From: Ferruccio Guidi Date: Fri, 1 Apr 2016 21:26:20 +0000 (+0000) Subject: - new component "s_transition" for the restored fqu and fquq X-Git-Tag: make_still_working~617 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=93bba1c94779e83184d111cd077d4167e42a74aa - new component "s_transition" for the restored fqu and fquq - minor additions - notation update - partial commit in the "static" component to park da and lsubd --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da.etc new file mode 100644 index 000000000..b8d3d4315 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da.etc @@ -0,0 +1,108 @@ +(**************************************************************************) +(* ___ *) +(* ||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/degree_6.ma". +include "basic_2/grammar/genv.ma". +include "basic_2/substitution/drop.ma". +include "basic_2/static/sd.ma". + +(* DEGREE ASSIGNMENT FOR TERMS **********************************************) + +(* activate genv *) +inductive da (h:sh) (o:sd h): relation4 genv lenv term nat ≝ +| da_sort: ∀G,L,s,d. deg h o s d → da h o G L (⋆s) d +| da_ldef: ∀G,L,K,V,i,d. ⬇[i] L ≡ K.ⓓV → da h o G K V d → da h o G L (#i) d +| da_ldec: ∀G,L,K,W,i,d. ⬇[i] L ≡ K.ⓛW → da h o G K W d → da h o G L (#i) (d+1) +| da_bind: ∀a,I,G,L,V,T,d. da h o G (L.ⓑ{I}V) T d → da h o G L (ⓑ{a,I}V.T) d +| da_flat: ∀I,G,L,V,T,d. da h o G L T d → da h o G L (ⓕ{I}V.T) d +. + +interpretation "degree assignment (term)" + 'Degree h o G L T d = (da h o G L T d). + +(* Basic inversion lemmas ***************************************************) + +fact da_inv_sort_aux: ∀h,o,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, o] d → + ∀s0. T = ⋆s0 → deg h o s0 d. +#h #o #G #L #T #d * -G -L -T -d +[ #G #L #s #d #Hkd #s0 #H destruct // +| #G #L #K #V #i #d #_ #_ #s0 #H destruct +| #G #L #K #W #i #d #_ #_ #s0 #H destruct +| #a #I #G #L #V #T #d #_ #s0 #H destruct +| #I #G #L #V #T #d #_ #s0 #H destruct +] +qed-. + +lemma da_inv_sort: ∀h,o,G,L,s,d. ⦃G, L⦄ ⊢ ⋆s ▪[h, o] d → deg h o s d. +/2 width=5 by da_inv_sort_aux/ qed-. + +fact da_inv_lref_aux: ∀h,o,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, o] d → ∀j. T = #j → + (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, o] d) ∨ + (∃∃K,W,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, o] d0 & + d = d0 + 1 + ). +#h #o #G #L #T #d * -G -L -T -d +[ #G #L #s #d #_ #j #H destruct +| #G #L #K #V #i #d #HLK #HV #j #H destruct /3 width=4 by ex2_2_intro, or_introl/ +| #G #L #K #W #i #d #HLK #HW #j #H destruct /3 width=6 by ex3_3_intro, or_intror/ +| #a #I #G #L #V #T #d #_ #j #H destruct +| #I #G #L #V #T #d #_ #j #H destruct +] +qed-. + +lemma da_inv_lref: ∀h,o,G,L,j,d. ⦃G, L⦄ ⊢ #j ▪[h, o] d → + (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, o] d) ∨ + (∃∃K,W,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, o] d0 & d = d0+1). +/2 width=3 by da_inv_lref_aux/ qed-. + +fact da_inv_gref_aux: ∀h,o,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, o] d → ∀p0. T = §p0 → ⊥. +#h #o #G #L #T #d * -G -L -T -d +[ #G #L #s #d #_ #p0 #H destruct +| #G #L #K #V #i #d #_ #_ #p0 #H destruct +| #G #L #K #W #i #d #_ #_ #p0 #H destruct +| #a #I #G #L #V #T #d #_ #p0 #H destruct +| #I #G #L #V #T #d #_ #p0 #H destruct +] +qed-. + +lemma da_inv_gref: ∀h,o,G,L,p,d. ⦃G, L⦄ ⊢ §p ▪[h, o] d → ⊥. +/2 width=9 by da_inv_gref_aux/ qed-. + +fact da_inv_bind_aux: ∀h,o,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, o] d → + ∀b,J,X,Y. T = ⓑ{b,J}Y.X → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, o] d. +#h #o #G #L #T #d * -G -L -T -d +[ #G #L #s #d #_ #b #J #X #Y #H destruct +| #G #L #K #V #i #d #_ #_ #b #J #X #Y #H destruct +| #G #L #K #W #i #d #_ #_ #b #J #X #Y #H destruct +| #a #I #G #L #V #T #d #HT #b #J #X #Y #H destruct // +| #I #G #L #V #T #d #_ #b #J #X #Y #H destruct +] +qed-. + +lemma da_inv_bind: ∀h,o,b,J,G,L,Y,X,d. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X ▪[h, o] d → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, o] d. +/2 width=4 by da_inv_bind_aux/ qed-. + +fact da_inv_flat_aux: ∀h,o,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, o] d → + ∀J,X,Y. T = ⓕ{J}Y.X → ⦃G, L⦄ ⊢ X ▪[h, o] d. +#h #o #G #L #T #d * -G -L -T -d +[ #G #L #s #d #_ #J #X #Y #H destruct +| #G #L #K #V #i #d #_ #_ #J #X #Y #H destruct +| #G #L #K #W #i #d #_ #_ #J #X #Y #H destruct +| #a #I #G #L #V #T #d #_ #J #X #Y #H destruct +| #I #G #L #V #T #d #HT #J #X #Y #H destruct // +] +qed-. + +lemma da_inv_flat: ∀h,o,J,G,L,Y,X,d. ⦃G, L⦄ ⊢ ⓕ{J}Y.X ▪[h, o] d → ⦃G, L⦄ ⊢ X ▪[h, o] d. +/2 width=5 by da_inv_flat_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_aaa.etc new file mode 100644 index 000000000..43d480c60 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_aaa.etc @@ -0,0 +1,31 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/aaa_lift.ma". +include "basic_2/static/da.ma". + +(* DEGREE ASSIGNMENT FOR TERMS **********************************************) + +(* Properties on atomic arity assignment for terms **************************) + +lemma aaa_da: ∀h,o,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃d. ⦃G, L⦄ ⊢ T ▪[h, o] d. +#h #o #G #L #T #A #H elim H -G -L -T -A +[ #G #L #s elim (deg_total h o s) /3 width=2 by da_sort, ex_intro/ +| * #G #L #K #V #B #i #HLK #_ * /3 width=5 by da_ldef, da_ldec, ex_intro/ +| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by da_bind, ex_intro/ +| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by da_bind, ex_intro/ +| #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by da_flat, ex_intro/ +| #G #L #W #T #A #_ #_ #_ * /3 width=2 by da_flat, ex_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_da.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_da.etc new file mode 100644 index 000000000..2bc863015 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_da.etc @@ -0,0 +1,38 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/da_lift.ma". + +(* DEGREE ASSIGNMENT FOR TERMS **********************************************) + +(* Main properties **********************************************************) + +theorem da_mono: ∀h,o,G,L,T,d1. ⦃G, L⦄ ⊢ T ▪[h, o] d1 → + ∀d2. ⦃G, L⦄ ⊢ T ▪[h, o] d2 → d1 = d2. +#h #o #G #L #T #d1 #H elim H -G -L -T -d1 +[ #G #L #s #d1 #Hkd1 #d2 #H + lapply (da_inv_sort … H) -G -L #Hkd2 + >(deg_mono … Hkd2 … Hkd1) -h -s -d2 // +| #G #L #K #V #i #d1 #HLK #_ #IHV #d2 #H + elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0 #HV0 [| #Hd0 ] + lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /2 width=1 by/ +| #G #L #K #W #i #d1 #HLK #_ #IHW #d2 #H + elim (da_inv_lref … H) -H * #K0 #W0 [| #d0 ] #HLK0 #HW0 [| #Hd0 ] + lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /3 width=1 by eq_f/ +| #a #I #G #L #V #T #d1 #_ #IHT #d2 #H + lapply (da_inv_bind … H) -H /2 width=1 by/ +| #I #G #L #V #T #d1 #_ #IHT #d2 #H + lapply (da_inv_flat … H) -H /2 width=1 by/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_lift.etc new file mode 100644 index 000000000..2f1a41554 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/da_lift.etc @@ -0,0 +1,78 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/drop_drop.ma". +include "basic_2/static/da.ma". + +(* DEGREE ASSIGNMENT FOR TERMS **********************************************) + +(* Properties on relocation *************************************************) + +lemma da_lift: ∀h,o,G,L1,T1,d. ⦃G, L1⦄ ⊢ T1 ▪[h, o] d → + ∀L2,c,l,k. ⬇[c, l, k] L2 ≡ L1 → ∀T2. ⬆[l, k] T1 ≡ T2 → + ⦃G, L2⦄ ⊢ T2 ▪[h, o] d. +#h #o #G #L1 #T1 #d #H elim H -G -L1 -T1 -d +[ #G #L1 #s #d #Hkd #L2 #c #l #k #_ #X #H + >(lift_inv_sort1 … H) -X /2 width=1 by da_sort/ +| #G #L1 #K1 #V1 #i #d #HLK1 #_ #IHV1 #L2 #c #l #k #HL21 #X #H + elim (lift_inv_lref1 … H) * #Hil #H destruct + [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H + elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct + /3 width=9 by da_ldef/ + | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 + /3 width=8 by da_ldef, drop_inv_gen/ + ] +| #G #L1 #K1 #W1 #i #d #HLK1 #_ #IHW1 #L2 #c #l #k #HL21 #X #H + elim (lift_inv_lref1 … H) * #Hil #H destruct + [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H + elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #W2 #HK21 #HW12 #H destruct + /3 width=8 by da_ldec/ + | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 + /3 width=8 by da_ldec, drop_inv_gen/ + ] +| #a #I #G #L1 #V1 #T1 #d #_ #IHT1 #L2 #c #l #k #HL21 #X #H + elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct + /4 width=5 by da_bind, drop_skip/ +| #I #G #L1 #V1 #T1 #d #_ #IHT1 #L2 #c #l #k #HL21 #X #H + elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct + /3 width=5 by da_flat/ +] +qed. + +(* Inversion lemmas on relocation *******************************************) + +lemma da_inv_lift: ∀h,o,G,L2,T2,d. ⦃G, L2⦄ ⊢ T2 ▪[h, o] d → + ∀L1,c,l,k. ⬇[c, l, k] L2 ≡ L1 → ∀T1. ⬆[l, k] T1 ≡ T2 → + ⦃G, L1⦄ ⊢ T1 ▪[h, o] d. +#h #o #G #L2 #T2 #d #H elim H -G -L2 -T2 -d +[ #G #L2 #s #d #Hkd #L1 #c #l #k #_ #X #H + >(lift_inv_sort2 … H) -X /2 width=1 by da_sort/ +| #G #L2 #K2 #V2 #i #d #HLK2 #HV2 #IHV2 #L1 #c #l #k #HL21 #X #H + elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HV2 | -IHV2 ] + [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldef/ + | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldef/ + ] +| #G #L2 #K2 #W2 #i #d #HLK2 #HW2 #IHW2 #L1 #c #l #k #HL21 #X #H + elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HW2 | -IHW2 ] + [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldec/ + | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldec/ + ] +| #a #I #G #L2 #V2 #T2 #d #_ #IHT2 #L1 #c #l #k #HL21 #X #H + elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /4 width=5 by da_bind, drop_skip/ +| #I #G #L2 #V2 #T2 #d #_ #IHT2 #L1 #c #l #k #HL21 #X #H + elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /3 width=5 by da_flat/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/degree_6.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/degree_6.etc new file mode 100644 index 000000000..bcc939ad1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/da/degree_6.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ▪ break [ term 46 o , break term 46 h ] break term 46 d )" + non associative with precedence 45 + for @{ 'Degree $h $o $G $L $T $d }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops.etc index 795f620cd..86ca4c76a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops.etc +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops.etc @@ -22,26 +22,6 @@ lemma drop_split: ∀L1,L2,l,m2,s. ⬇[s, l, m2] L1 ≡ L2 → ∀m1. m1 ≤ m2 ] qed-. -(* Basic_2A1: includes: drop_split *) -lemma drops_split_trans: ∀L1,L2,f,c. ⬇*[c, f] L1 ≡ L2 → ∀f1,f2. f1 ⊚ f2 ≡ f → - ∃∃L. ⬇*[c, f1] L1 ≡ L & ⬇*[c, f2] L ≡ L2. -#L1 #L2 #f #c #H elim H -L1 -L2 -f -[ #f #Hc #f1 #f2 #Hf @(ex2_intro … (⋆)) @drops_atom - #H lapply (Hc H) -c - #H elim (after_inv_isid3 … Hf H) -f // -| #I #L1 #L2 #V #f #HL12 #IHL12 #f1 #f2 #Hf elim (after_inv_xxS … Hf) -Hf * - [ #g1 #g2 #Hf #H1 #H2 destruct elim (IHL12 … Hf) -f - #L #HL1 #HL2 @(ex2_intro … (L.ⓑ{I}V)) /2 width=1 by drops_drop/ - @drops_skip // - | #g1 #Hf #H destruct elim (IHL12 … Hf) -f - /3 width=5 by ex2_intro, drops_drop/ - ] -| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 #f1 #f2 #Hf elim (after_inv_xxO … Hf) -Hf - #g1 #g2 #Hf #H1 #H2 destruct elim (lifts_split_trans … HV21 … Hf) -HV21 - elim (IHL12 … Hf) -f /3 width=5 by ex2_intro, drops_skip/ -] -qed-. - lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 0. /2 width=5 by drop_inv_length_eq/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/rdropstar_3.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/rdropstar_3.etc deleted file mode 100644 index 8a6f2f4c1..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/rdropstar_3.etc +++ /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 m ] break term 46 L1 ≡ break term 46 L2 )" - non associative with precedence 45 - for @{ 'RDropStar $m $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/fqu/fqu.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/fqu/fqu.etc new file mode 100644 index 000000000..b5066d00f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/fqu/fqu.etc @@ -0,0 +1,16 @@ +(* Removed theorems *********************************************************) + +include "basic_2/substitution/drop.ma". + +| fqu_drop : ∀G,L,K,T,U,k. + ⬇[⫯k] L ≡ K → ⬆[0, ⫯k] T ≡ U → fqu G L U G K T + +lemma fqu_drop_lt: ∀G,L,K,T,U,k. 0 < k → + ⬇[k] L ≡ K → ⬆[0, k] T ≡ U → ⦃G, L, U⦄ ⊐ ⦃G, K, T⦄. +#G #L #K #T #U #k #Hm lapply (ylt_inv_O1 … Hm) -Hm +#Hm (drop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T // +qed. + +inductive fquq: tri_relation genv lenv term ≝ +| fquq_lref_O : ∀I,G,L,V. fquq G (L.ⓑ{I}V) (#0) G L V +| fquq_pair_sn: ∀I,G,L,V,T. fquq G L (②{I}V.T) G L V +| fquq_bind_dx: ∀a,I,G,L,V,T. fquq G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T +| fquq_flat_dx: ∀I,G, L,V,T. fquq G L (ⓕ{I}V.T) G L T +| fquq_drop : ∀G,L,K,T,U,k. + ⬇[k] L ≡ K → ⬆[0, k] T ≡ U → fquq G L U G K T +. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lrsubeqd_5.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lrsubeqd_5.etc new file mode 100644 index 000000000..586c75e5c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lrsubeqd_5.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( G ⊢ break term 46 L1 ⫃ ▪ break [ term 46 o, break term 46 h ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'LRSubEqD $h $o $G $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd.etc new file mode 100644 index 000000000..518ad3bc8 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd.etc @@ -0,0 +1,151 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lrsubeqd_5.ma". +include "basic_2/static/lsubr.ma". +include "basic_2/static/da.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) + +inductive lsubd (h) (o) (G): relation lenv ≝ +| lsubd_atom: lsubd h o G (⋆) (⋆) +| lsubd_pair: ∀I,L1,L2,V. lsubd h o G L1 L2 → + lsubd h o G (L1.ⓑ{I}V) (L2.ⓑ{I}V) +| lsubd_beta: ∀L1,L2,W,V,d. ⦃G, L1⦄ ⊢ V ▪[h, o] d+1 → ⦃G, L2⦄ ⊢ W ▪[h, o] d → + lsubd h o G L1 L2 → lsubd h o G (L1.ⓓⓝW.V) (L2.ⓛW) +. + +interpretation + "local environment refinement (degree assignment)" + 'LRSubEqD h o G L1 L2 = (lsubd h o G L1 L2). + +(* Basic forward lemmas *****************************************************) + +lemma lsubd_fwd_lsubr: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → L1 ⫃ L2. +#h #o #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +fact lsubd_inv_atom1_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → L1 = ⋆ → L2 = ⋆. +#h #o #G #L1 #L2 * -L1 -L2 +[ // +| #I #L1 #L2 #V #_ #H destruct +| #L1 #L2 #W #V #d #_ #_ #_ #H destruct +] +qed-. + +lemma lsubd_inv_atom1: ∀h,o,G,L2. G ⊢ ⋆ ⫃▪[h, o] L2 → L2 = ⋆. +/2 width=6 by lsubd_inv_atom1_aux/ qed-. + +fact lsubd_inv_pair1_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → + ∀I,K1,X. L1 = K1.ⓑ{I}X → + (∃∃K2. G ⊢ K1 ⫃▪[h, o] K2 & L2 = K2.ⓑ{I}X) ∨ + ∃∃K2,W,V,d. ⦃G, K1⦄ ⊢ V ▪[h, o] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d & + G ⊢ K1 ⫃▪[h, o] K2 & + I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. +#h #o #G #L1 #L2 * -L1 -L2 +[ #J #K1 #X #H destruct +| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/ +| #L1 #L2 #W #V #d #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9 by ex6_4_intro, or_intror/ +] +qed-. + +lemma lsubd_inv_pair1: ∀h,o,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃▪[h, o] L2 → + (∃∃K2. G ⊢ K1 ⫃▪[h, o] K2 & L2 = K2.ⓑ{I}X) ∨ + ∃∃K2,W,V,d. ⦃G, K1⦄ ⊢ V ▪[h, o] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d & + G ⊢ K1 ⫃▪[h, o] K2 & + I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. +/2 width=3 by lsubd_inv_pair1_aux/ qed-. + +fact lsubd_inv_atom2_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → L2 = ⋆ → L1 = ⋆. +#h #o #G #L1 #L2 * -L1 -L2 +[ // +| #I #L1 #L2 #V #_ #H destruct +| #L1 #L2 #W #V #d #_ #_ #_ #H destruct +] +qed-. + +lemma lsubd_inv_atom2: ∀h,o,G,L1. G ⊢ L1 ⫃▪[h, o] ⋆ → L1 = ⋆. +/2 width=6 by lsubd_inv_atom2_aux/ qed-. + +fact lsubd_inv_pair2_aux: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → + ∀I,K2,W. L2 = K2.ⓑ{I}W → + (∃∃K1. G ⊢ K1 ⫃▪[h, o] K2 & L1 = K1.ⓑ{I}W) ∨ + ∃∃K1,V,d. ⦃G, K1⦄ ⊢ V ▪[h, o] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d & + G ⊢ K1 ⫃▪[h, o] K2 & I = Abst & L1 = K1. ⓓⓝW.V. +#h #o #G #L1 #L2 * -L1 -L2 +[ #J #K2 #U #H destruct +| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/ +| #L1 #L2 #W #V #d #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7 by ex5_3_intro, or_intror/ +] +qed-. + +lemma lsubd_inv_pair2: ∀h,o,I,G,L1,K2,W. G ⊢ L1 ⫃▪[h, o] K2.ⓑ{I}W → + (∃∃K1. G ⊢ K1 ⫃▪[h, o] K2 & L1 = K1.ⓑ{I}W) ∨ + ∃∃K1,V,d. ⦃G, K1⦄ ⊢ V ▪[h, o] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, o] d & + G ⊢ K1 ⫃▪[h, o] K2 & I = Abst & L1 = K1. ⓓⓝW.V. +/2 width=3 by lsubd_inv_pair2_aux/ qed-. + +(* Basic properties *********************************************************) + +lemma lsubd_refl: ∀h,o,G,L. G ⊢ L ⫃▪[h, o] L. +#h #o #G #L elim L -L /2 width=1 by lsubd_pair/ +qed. + +(* Note: the constant 0 cannot be generalized *) +lemma lsubd_drop_O1_conf: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → + ∀K1,c,k. ⬇[c, 0, k] L1 ≡ K1 → + ∃∃K2. G ⊢ K1 ⫃▪[h, o] K2 & ⬇[c, 0, k] L2 ≡ K2. +#h #o #G #L1 #L2 #H elim H -L1 -L2 +[ /2 width=3 by ex2_intro/ +| #I #L1 #L2 #V #_ #IHL12 #K1 #c #k #H + elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1 + [ destruct + elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H + <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, drop_pair, ex2_intro/ + | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/ + ] +| #L1 #L2 #W #V #d #HV #HW #_ #IHL12 #K1 #c #k #H + elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1 + [ destruct + elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H + <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_beta, drop_pair, ex2_intro/ + | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/ + ] +] +qed-. + +(* Note: the constant 0 cannot be generalized *) +lemma lsubd_drop_O1_trans: ∀h,o,G,L1,L2. G ⊢ L1 ⫃▪[h, o] L2 → + ∀K2,c,k. ⬇[c, 0, k] L2 ≡ K2 → + ∃∃K1. G ⊢ K1 ⫃▪[h, o] K2 & ⬇[c, 0, k] L1 ≡ K1. +#h #o #G #L1 #L2 #H elim H -L1 -L2 +[ /2 width=3 by ex2_intro/ +| #I #L1 #L2 #V #_ #IHL12 #K2 #c #k #H + elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2 + [ destruct + elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H + <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, drop_pair, ex2_intro/ + | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/ + ] +| #L1 #L2 #W #V #d #HV #HW #_ #IHL12 #K2 #c #k #H + elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2 + [ destruct + elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H + <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_beta, drop_pair, ex2_intro/ + | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/ + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd_da.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd_da.etc new file mode 100644 index 000000000..62b67e09c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd_da.etc @@ -0,0 +1,65 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/da_da.ma". +include "basic_2/static/lsubd.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) + +(* Properties on degree assignment ******************************************) + +lemma lsubd_da_trans: ∀h,o,G,L2,T,d. ⦃G, L2⦄ ⊢ T ▪[h, o] d → + ∀L1. G ⊢ L1 ⫃▪[h, o] L2 → ⦃G, L1⦄ ⊢ T ▪[h, o] d. +#h #o #G #L2 #T #d #H elim H -G -L2 -T -d +[ /2 width=1 by da_sort/ +| #G #L2 #K2 #V #i #d #HLK2 #_ #IHV #L1 #HL12 + elim (lsubd_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 + elim (lsubd_inv_pair2 … H) -H * #K1 [ | -IHV -HLK1 ] + [ #HK12 #H destruct /3 width=4 by da_ldef/ + | #W #d0 #_ #_ #_ #H destruct + ] +| #G #L2 #K2 #W #i #d #HLK2 #HW #IHW #L1 #HL12 + elim (lsubd_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 + elim (lsubd_inv_pair2 … H) -H * #K1 [ -HW | -IHW ] + [ #HK12 #H destruct /3 width=4 by da_ldec/ + | #V #d0 #HV #H0W #_ #_ #H destruct + lapply (da_mono … H0W … HW) -H0W -HW #H destruct /3 width=7 by da_ldef, da_flat/ + ] +| /4 width=1 by lsubd_pair, da_bind/ +| /3 width=1 by da_flat/ +] +qed-. + +lemma lsubd_da_conf: ∀h,o,G,L1,T,d. ⦃G, L1⦄ ⊢ T ▪[h, o] d → + ∀L2. G ⊢ L1 ⫃▪[h, o] L2 → ⦃G, L2⦄ ⊢ T ▪[h, o] d. +#h #o #G #L1 #T #d #H elim H -G -L1 -T -d +[ /2 width=1 by da_sort/ +| #G #L1 #K1 #V #i #d #HLK1 #HV #IHV #L2 #HL12 + elim (lsubd_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2 + elim (lsubd_inv_pair1 … H) -H * #K2 [ -HV | -IHV ] + [ #HK12 #H destruct /3 width=4 by da_ldef/ + | #W0 #V0 #d0 #HV0 #HW0 #_ #_ #H1 #H2 destruct + lapply (da_inv_flat … HV) -HV #H0V0 + lapply (da_mono … H0V0 … HV0) -H0V0 -HV0 #H destruct /2 width=4 by da_ldec/ + ] +| #G #L1 #K1 #W #i #d #HLK1 #HW #IHW #L2 #HL12 + elim (lsubd_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2 + elim (lsubd_inv_pair1 … H) -H * #K2 [ -HW | -IHW ] + [ #HK12 #H destruct /3 width=4 by da_ldec/ + | #W0 #V0 #d0 #HV0 #HW0 #_ #H destruct + ] +| /4 width=1 by lsubd_pair, da_bind/ +| /3 width=1 by da_flat/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd_lsubd.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd_lsubd.etc new file mode 100644 index 000000000..b198b6ab7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsubd/lsubd_lsubd.etc @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/static/lsubd_da.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) + +(* Main properties **********************************************************) + +theorem lsubd_trans: ∀h,o,G. Transitive … (lsubd h o G). +#h #o #G #L1 #L #H elim H -L1 -L +[ #X #H >(lsubd_inv_atom1 … H) -H // +| #I #L1 #L #Y #HL1 #IHL1 #X #H + elim (lsubd_inv_pair1 … H) -H * #L2 + [ #HL2 #H destruct /3 width=1 by lsubd_pair/ + | #W #V #d #HV #HW #HL2 #H1 #H2 #H3 destruct + /3 width=3 by lsubd_beta, lsubd_da_trans/ + ] +| #L1 #L #W #V #d #HV #HW #HL1 #IHL1 #X #H + elim (lsubd_inv_pair1 … H) -H * #L2 + [ #HL2 #H destruct /3 width=5 by lsubd_beta, lsubd_da_conf/ + | #W0 #V0 #d0 #_ #_ #_ #H destruct + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/append.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/append.ma new file mode 100644 index 000000000..cdbb45d48 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/append.ma @@ -0,0 +1,69 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "ground_2/notation/functions/append_2.ma". +include "basic_2/notation/functions/snbind2_3.ma". +include "basic_2/notation/functions/snabbr_2.ma". +include "basic_2/notation/functions/snabst_2.ma". +include "basic_2/grammar/lenv.ma". + +(* APPEND FOR LOCAL ENVIRONMENTS ********************************************) + +rec definition append L K on K ≝ match K with +[ LAtom ⇒ L +| LPair K I V ⇒ (append L K).ⓑ{I}V +]. + +interpretation "append (local environment)" 'Append L1 L2 = (append L1 L2). + +interpretation "local environment tail binding construction (binary)" + 'SnBind2 I T L = (append (LPair LAtom I T) L). + +interpretation "tail abbreviation (local environment)" + 'SnAbbr T L = (append (LPair LAtom Abbr T) L). + +interpretation "tail abstraction (local environment)" + 'SnAbst L T = (append (LPair LAtom Abst T) L). + +definition d_appendable_sn: predicate (lenv→relation term) ≝ λR. + ∀K,T1,T2. R K T1 T2 → ∀L. R (L@@K) T1 T2. + +(* Basic properties *********************************************************) + +lemma append_atom: ∀L. L @@ ⋆ = L. +// qed. + +lemma append_pair: ∀I,L,K,V. L@@(K.ⓑ{I}V) = (L@@K).ⓑ{I}V. +// qed. + +lemma append_atom_sn: ∀L. ⋆@@L = L. +#L elim L -L // +#L #I #V >append_pair // +qed. + +lemma append_assoc: associative … append. +#L1 #L2 #L3 elim L3 -L3 // +qed. + +lemma append_shift: ∀L,K,I,V. L@@(ⓑ{I}V.K) = (L.ⓑ{I}V)@@K. +#L #K #I #V append_pair #H +elim (destruct_lpair_lpair_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *) +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/append_length.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/append_length.ma new file mode 100644 index 000000000..21e5a868e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/append_length.ma @@ -0,0 +1,116 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/grammar/lenv_length.ma". +include "basic_2/grammar/append.ma". + +(* APPEND FOR LOCAL ENVIRONMENTS ********************************************) + +(* Properties with length for local environments ****************************) + +lemma append_length: ∀L1,L2. |L1 @@ L2| = |L1| + |L2|. +#L1 #L2 elim L2 -L2 // +#L2 #I #V2 >append_pair >length_pair >length_pair // +qed. + +lemma ltail_length: ∀I,L,V. |ⓑ{I}V.L| = ⫯|L|. +#I #L #V >append_length // +qed. + +(* Basic_1: was just: chead_ctail *) +lemma lpair_ltail: ∀L,I,V. ∃∃J,K,W. L.ⓑ{I}V = ⓑ{J}W.K & |L| = |K|. +#L elim L -L /2 width=5 by ex2_3_intro/ +#L #Z #X #IHL #I #V elim (IHL Z X) -IHL +#J #K #W #H #_ >H -H >ltail_length +@(ex2_3_intro … J (K.ⓑ{I}V) W) /2 width=1 by length_pair/ +qed-. + +(* Advanced inversion lemmas on length for local environments ***************) + +(* Basic_2A1: was: length_inv_pos_dx_ltail *) +lemma length_inv_succ_dx_ltail: ∀L,l. |L| = ⫯l → + ∃∃I,K,V. |K| = l & L = ⓑ{I}V.K. +#Y #l #H elim (length_inv_succ_dx … H) -H #I #L #V #Hl #HLK destruct +elim (lpair_ltail L I V) /2 width=5 by ex2_3_intro/ +qed-. + +(* Basic_2A1: was: length_inv_pos_sn_ltail *) +lemma length_inv_succ_sn_ltail: ∀L,l. ⫯l = |L| → + ∃∃I,K,V. l = |K| & L = ⓑ{I}V.K. +#Y #l #H elim (length_inv_succ_sn … H) -H #I #L #V #Hl #HLK destruct +elim (lpair_ltail L I V) /2 width=5 by ex2_3_intro/ +qed-. + +(* Inversion lemmas with length for local environments **********************) + +(* Basic_2A1: was: append_inj_sn *) +lemma append_inj_length_sn: ∀K1,K2,L1,L2. L1 @@ K1 = L2 @@ K2 → |K1| = |K2| → + L1 = L2 ∧ K1 = K2. +#K1 elim K1 -K1 +[ * /2 width=1 by conj/ + #K2 #I2 #V2 #L1 #L2 #_ >length_atom >length_pair + #H destruct +| #K1 #I1 #V1 #IH * + [ #L1 #L2 #_ >length_atom >length_pair + #H destruct + | #K2 #I2 #V2 #L1 #L2 #H1 #H2 + elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *) + elim (IH … H1) -IH -H1 /2 width=1 by conj/ + ] +] +qed-. + +(* Note: lemma 750 *) +(* Basic_2A1: was: append_inj_dx *) +lemma append_inj_length_dx: ∀K1,K2,L1,L2. L1 @@ K1 = L2 @@ K2 → |L1| = |L2| → + L1 = L2 ∧ K1 = K2. +#K1 elim K1 -K1 +[ * /2 width=1 by conj/ + #K2 #I2 #V2 #L1 #L2 >append_atom >append_pair #H destruct + >length_pair >append_length >plus_n_Sm + #H elim (plus_xSy_x_false … H) +| #K1 #I1 #V1 #IH * + [ #L1 #L2 >append_pair >append_atom #H destruct + >length_pair >append_length >plus_n_Sm #H + lapply (discr_plus_x_xy … H) -H #H destruct + | #K2 #I2 #V2 #L1 #L2 >append_pair >append_pair #H1 #H2 + elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *) + elim (IH … H1) -IH -H1 /2 width=1 by conj/ + ] +] +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma append_inj_dx: ∀L,K1,K2. L@@K1 = L@@K2 → K1 = K2. +#L #K1 #K2 #H elim (append_inj_length_dx … H) -H // +qed-. + +lemma append_inv_refl_dx: ∀L,K. L@@K = L → K = ⋆. +#L #K #H elim (append_inj_dx … (⋆) … H) // +qed-. + +lemma append_inv_pair_dx: ∀I,L,K,V. L@@K = L.ⓑ{I}V → K = ⋆.ⓑ{I}V. +#I #L #K #V #H elim (append_inj_dx … (⋆.ⓑ{I}V) … H) // +qed-. + +(* Basic eliminators ********************************************************) + +(* Basic_1: was: c_tail_ind *) +lemma lenv_ind_alt: ∀R:predicate lenv. + R (⋆) → (∀I,L,T. R L → R (ⓑ{I}T.L)) → + ∀L. R L. +#R #IH1 #IH2 #L @(f_ind … length … L) -L #x #IHx * // -IH1 +#L #I #V #H destruct elim (lpair_ltail L I V) /4 width=1 by/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma index ff183d177..80738ebbd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma @@ -66,3 +66,6 @@ lemma discr_lpair_x_xy: ∀I,V,L. L = L.ⓑ{I}V → ⊥. elim (destruct_lpair_lpair_aux … H) -H #H1 #H2 #H3 destruct /2 width=1 by/ (**) (* destruct lemma needed *) ] qed-. + +lemma discr_lpair_xy_x: ∀I,V,L. L.ⓑ{I}V = L→ ⊥. +/2 width=4 by discr_lpair_x_xy/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv_append.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv_append.ma deleted file mode 100644 index aa2e587d8..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv_append.ma +++ /dev/null @@ -1,141 +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 "ground_2/notation/functions/append_2.ma". -include "basic_2/notation/functions/snbind2_3.ma". -include "basic_2/notation/functions/snabbr_2.ma". -include "basic_2/notation/functions/snabst_2.ma". -include "basic_2/grammar/lenv_length.ma". - -(* LOCAL ENVIRONMENTS *******************************************************) - -let rec append L K on K ≝ match K with -[ LAtom ⇒ L -| LPair K I V ⇒ (append L K). ⓑ{I} V -]. - -interpretation "append (local environment)" 'Append L1 L2 = (append L1 L2). - -interpretation "local environment tail binding construction (binary)" - 'SnBind2 I T L = (append (LPair LAtom I T) L). - -interpretation "tail abbreviation (local environment)" - 'SnAbbr T L = (append (LPair LAtom Abbr T) L). - -interpretation "tail abstraction (local environment)" - 'SnAbst L T = (append (LPair LAtom Abst T) L). - -definition d_appendable_sn: predicate (lenv→relation term) ≝ λR. - ∀K,T1,T2. R K T1 T2 → ∀L. R (L @@ K) T1 T2. - -(* Basic properties *********************************************************) - -lemma append_atom: ∀L. L @@ ⋆ = L. -// qed. - -lemma append_pair: ∀I,L,K,V. L @@ (K.ⓑ{I}V) = (L @@ K).ⓑ{I}V. -// qed. - -lemma append_atom_sn: ∀L. ⋆ @@ L = L. -#L elim L -L // -#L #I #V >append_pair // -qed. - -lemma append_assoc: associative … append. -#L1 #L2 #L3 elim L3 -L3 // -qed. - -lemma append_length: ∀L1,L2. |L1 @@ L2| = |L1| + |L2|. -#L1 #L2 elim L2 -L2 // -#L2 #I #V2 >append_pair >length_pair >length_pair // -qed. - -lemma ltail_length: ∀I,L,V. |ⓑ{I}V.L| = ⫯|L|. -#I #L #V >append_length // -qed. - -(* Basic_1: was just: chead_ctail *) -lemma lpair_ltail: ∀L,I,V. ∃∃J,K,W. L.ⓑ{I}V = ⓑ{J}W.K & |L| = |K|. -#L elim L -L /2 width=5 by ex2_3_intro/ -#L #Z #X #IHL #I #V elim (IHL Z X) -IHL -#J #K #W #H #_ >H -H >ltail_length -@(ex2_3_intro … J (K.ⓑ{I}V) W) /2 width=1 by length_pair/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma append_inj_sn: ∀K1,K2,L1,L2. L1 @@ K1 = L2 @@ K2 → |K1| = |K2| → - L1 = L2 ∧ K1 = K2. -#K1 elim K1 -K1 -[ * /2 width=1 by conj/ - #K2 #I2 #V2 #L1 #L2 #_ >length_atom >length_pair - #H destruct -| #K1 #I1 #V1 #IH * - [ #L1 #L2 #_ >length_atom >length_pair - #H destruct - | #K2 #I2 #V2 #L1 #L2 #H1 #H2 - elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *) - elim (IH … H1) -IH -H1 /2 width=1 by conj/ - ] -] -qed-. - -(* Note: lemma 750 *) -lemma append_inj_dx: ∀K1,K2,L1,L2. L1 @@ K1 = L2 @@ K2 → |L1| = |L2| → - L1 = L2 ∧ K1 = K2. -#K1 elim K1 -K1 -[ * /2 width=1 by conj/ - #K2 #I2 #V2 #L1 #L2 >append_atom >append_pair #H destruct - >length_pair >append_length >plus_n_Sm - #H elim (plus_xSy_x_false … H) -| #K1 #I1 #V1 #IH * - [ #L1 #L2 >append_pair >append_atom #H destruct - >length_pair >append_length >plus_n_Sm #H - lapply (discr_plus_x_xy … H) -H #H destruct - | #K2 #I2 #V2 #L1 #L2 >append_pair >append_pair #H1 #H2 - elim (destruct_lpair_lpair_aux … H1) -H1 #H1 #H3 #H4 destruct (**) (* destruct lemma needed *) - elim (IH … H1) -IH -H1 /2 width=1 by conj/ - ] -] -qed-. - -lemma append_inv_refl_dx: ∀L,K. L @@ K = L → K = ⋆. -#L #K #H elim (append_inj_dx … (⋆) … H) // -qed-. - -lemma append_inv_pair_dx: ∀I,L,K,V. L @@ K = L.ⓑ{I}V → K = ⋆.ⓑ{I}V. -#I #L #K #V #H elim (append_inj_dx … (⋆.ⓑ{I}V) … H) // -qed-. - -lemma length_inv_pos_dx_ltail: ∀L,l. |L| = ⫯l → - ∃∃I,K,V. |K| = l & L = ⓑ{I}V.K. -#Y #l #H elim (length_inv_succ_dx … H) -H #I #L #V #Hl #HLK destruct -elim (lpair_ltail L I V) /2 width=5 by ex2_3_intro/ -qed-. - -lemma length_inv_pos_sn_ltail: ∀L,l. ⫯l = |L| → - ∃∃I,K,V. l = |K| & L = ⓑ{I}V.K. -#Y #l #H elim (length_inv_succ_sn … H) -H #I #L #V #Hl #HLK destruct -elim (lpair_ltail L I V) /2 width=5 by ex2_3_intro/ -qed-. - -(* Basic eliminators ********************************************************) - -(* Basic_1: was: c_tail_ind *) -lemma lenv_ind_alt: ∀R:predicate lenv. - R (⋆) → (∀I,L,T. R L → R (ⓑ{I}T.L)) → - ∀L. R L. -#R #IH1 #IH2 #L @(f_ind … length … L) -L #x #IHx * // -IH1 -#L #I #V #H destruct elim (lpair_ltail L I V) /4 width=1 by/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/names.txt b/matita/matita/contribs/lambdadelta/basic_2/names.txt index 1a26839d5..53d34c979 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/names.txt +++ b/matita/matita/contribs/lambdadelta/basic_2/names.txt @@ -20,12 +20,12 @@ c : local dropping kind parameter (true = restricted, false = general) d : term degree e : reserved: future use (\lambda\delta 3) f,g : local reference transforming map -h : sort degree parameter +h : sort hierarchy parameter i,j : local reference depth (de Bruijn's) k,l : global reference level m : n : type iterations -o : sort hierarchy parameter +o : sort degree parameter p,q : binder polarity r : reduction kind parameter (true = ordinary, false = extended) s : sort index diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma index 575359650..e60c75d42 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma @@ -20,11 +20,11 @@ notation "hvbox( ⦃G, L⦄ ⊢ break ⌘ ⦃ term 46 T ⦄ ≡ break term 46 k notation "hvbox( ⦃ term 46 h , break term 46 L ⦄ ⊢ break term 46 T ÷ break term 46 A )" non associative with precedence 45 - for @{ 'BinaryArity $o $L $T $A }. + for @{ 'BinaryArity $h $L $T $A }. notation "hvbox( h ⊢ break term 46 L1 ÷ ⫃ break term 46 L2 )" non associative with precedence 45 - for @{ 'LRSubEqB $o $L1 $L2 }. + for @{ 'LRSubEqB $h $L1 $L2 }. notation "hvbox( L1 ⊢ ⬌ break term 46 L2 )" non associative with precedence 45 @@ -36,12 +36,12 @@ notation "hvbox( L1 ⊢ ⬌* break term 46 L2 )" notation "hvbox( ⦃ term 46 h , break term 46 L ⦄ ⊢ break term 46 T1 : break term 46 T2 )" non associative with precedence 45 - for @{ 'NativeType $o $L $T1 $T2 }. + for @{ 'NativeType $h $L $T1 $T2 }. notation "hvbox( ⦃ term 46 h , break term 46 L ⦄ ⊢ break term 46 T1 : : break term 46 T2 )" non associative with precedence 45 - for @{ 'NativeTypeAlt $o $L $T1 $T2 }. + for @{ 'NativeTypeAlt $h $L $T1 $T2 }. notation "hvbox( ⦃ term 46 h , break term 46 L ⦄ ⊢ break term 46 T1 : * break term 46 T2 )" non associative with precedence 45 - for @{ 'NativeTypeStar $o $L $T1 $T2 }. + for @{ 'NativeTypeStar $h $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpred_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpred_8.ma index a555a5f60..68a41ffc0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpred_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpred_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≽ break [ term 46 o, break term 46 h ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'BTPRed $o $h $G1 $L1 $T1 $G2 $L2 $T2 }. + for @{ 'BTPRed $h $o $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredalt_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredalt_8.ma index 0b7e9ea0b..97ea5b59f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredalt_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredalt_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≽ ≽ break [ term 46 o, break term 46 h ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'BTPRedAlt $o $h $G1 $L1 $T1 $G2 $L2 $T2 }. + for @{ 'BTPRedAlt $h $o $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredproper_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredproper_8.ma index d155e84bd..2dce25d43 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredproper_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredproper_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≻ break [ term 46 o, break term 46 h ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'BTPRedProper $o $h $G1 $L1 $T1 $G2 $L2 $T2 }. + for @{ 'BTPRedProper $h $o $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstar_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstar_8.ma index c14e163d2..c48d221d4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstar_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstar_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≥ break [ term 46 o, break term 46 h ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'BTPRedStar $o $h $G1 $L1 $T1 $G2 $L2 $T2 }. + for @{ 'BTPRedStar $h $o $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstaralt_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstaralt_8.ma index 67609ec75..79b3fd9d2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstaralt_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btpredstaralt_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ ≥ ≥ break [ term 46 o, break term 46 h ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'BTPRedStarAlt $o $h $G1 $L1 $T1 $G2 $L2 $T2 }. + for @{ 'BTPRedStarAlt $h $o $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsn_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsn_5.ma index adacfe0a5..06d5b7b88 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsn_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsn_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦥ [ term 46 o, break term 46 h ] break ⦃ term 46 G, break term 46 L, break term 46 T ⦄ )" non associative with precedence 45 - for @{ 'BTSN $o $h $G $L $T }. + for @{ 'BTSN $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsnalt_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsnalt_5.ma index 5e8fd5d84..c303e309e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsnalt_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/btsnalt_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦥ ⦥ [ term 46 o, break term 46 h ] break ⦃ term 46 G, break term 46 L, break term 46 T ⦄ )" non associative with precedence 45 - for @{ 'BTSNAlt $o $h $G $L $T }. + for @{ 'BTSNAlt $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/cosn_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/cosn_5.ma index 5ac085cdc..e9c9e645d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/cosn_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/cosn_5.ma @@ -16,4 +16,4 @@ notation "hvbox( G ⊢ ~ ⬊ * break [ term 46 o , break term 46 h , break term 46 f ] break term 46 L )" non associative with precedence 45 - for @{ 'CoSN $o $h $f $G $L }. + for @{ 'CoSN $h $o $f $G $L }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/degree_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/degree_6.ma deleted file mode 100644 index 1de815516..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/degree_6.ma +++ /dev/null @@ -1,19 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ▪ break [ term 46 o , break term 46 h ] break term 46 d )" - non associative with precedence 45 - for @{ 'Degree $o $h $G $L $T $d }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpconvstar_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpconvstar_8.ma index 433923fea..8ae707641 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpconvstar_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpconvstar_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • * ⬌ * break [ term 46 o , break term 46 h , break term 46 n1 , break term 46 n2 ] break term 46 T2 )" non associative with precedence 45 - for @{ 'DPConvStar $o $h $n1 $n2 $G $L $T1 $T2 }. + for @{ 'DPConvStar $h $o $n1 $n2 $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpredstar_7.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpredstar_7.ma index ceb7aaf65..a85c0e4f8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpredstar_7.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/dpredstar_7.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • * ➡ * break [ term 46 o , break term 46 h , break term 46 n ] break term 46 T2 )" non associative with precedence 45 - for @{ 'DPRedStar $o $h $n $G $L $T1 $T2 }. + for @{ 'DPRedStar $h $o $n $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazybtpredstarproper_8.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazybtpredstarproper_8.ma index a2a86fddd..c99a3af6d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazybtpredstarproper_8.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazybtpredstarproper_8.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G1, break term 46 L1, break term 46 T1 ⦄ >≡ break [ term 46 o, break term 46 h ] break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'LazyBTPRedStarProper $o $h $G1 $L1 $T1 $G2 $L2 $T2 }. + for @{ 'LazyBTPRedStarProper $h $o $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqd_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqd_5.ma deleted file mode 100644 index b43c8bf17..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqd_5.ma +++ /dev/null @@ -1,19 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( G ⊢ break term 46 L1 ⫃ ▪ break [ term 46 o, break term 46 h ] break term 46 L2 )" - non associative with precedence 45 - for @{ 'LRSubEqD $o $h $G $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqv_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqv_5.ma index 55067b57e..f39c6bcc7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqv_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lrsubeqv_5.ma @@ -16,4 +16,4 @@ notation "hvbox( G ⊢ break term 46 L1 ⫃ ¡ break [ term 46 o, break term 46 h ] break term 46 L2 )" non associative with precedence 45 - for @{ 'LRSubEqV $o $h $G $L1 $L2 }. + for @{ 'LRSubEqV $h $o $G $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_5.ma index 631034964..50e2452c2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ¡ break [ term 46 o, break term 46 h ] )" non associative with precedence 45 - for @{ 'NativeValid $o $h $G $L $T }. + for @{ 'NativeValid $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_6.ma index f2a9b70ed..2cd233332 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/nativevalid_6.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T ¡ break [ term 46 o , break term 46 h , break term 46 d ] )" non associative with precedence 45 - for @{ 'NativeValid $o $h $d $G $L $T }. + for @{ 'NativeValid $h $o $d $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_6.ma index 536568f7c..4fa78f516 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_6.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ break [ term 46 o , break term 46 h ] break term 46 T2 )" non associative with precedence 45 - for @{ 'PRed $o $h $G $L $T1 $T2 }. + for @{ 'PRed $h $o $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predeval_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predeval_6.ma index cd474d498..80c61f26f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predeval_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predeval_6.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ * break [ term 46 o , break term 46 h ] break 𝐍 ⦃ term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'PRedEval $o $h $G $L $T1 $T2 }. + for @{ 'PRedEval $h $o $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednormal_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednormal_5.ma index 8eda91490..9df81e650 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednormal_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednormal_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐍 break ⦃ term 46 T ⦄ )" non associative with precedence 45 - for @{ 'PRedNormal $o $h $G $L $T }. + for @{ 'PRedNormal $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednotreducible_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednotreducible_5.ma index 788e2b428..c6f7e645e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednotreducible_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/prednotreducible_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐈 break ⦃ term 46 T ⦄ )" non associative with precedence 45 - for @{ 'PRedNotReducible $o $h $G $L $T }. + for @{ 'PRedNotReducible $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predreducible_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predreducible_5.ma index af010f4fb..ca0d12b72 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predreducible_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predreducible_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] 𝐑 break ⦃ term 46 T ⦄ )" non associative with precedence 45 - for @{ 'PRedReducible $o $h $G $L $T }. + for @{ 'PRedReducible $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsn_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsn_5.ma index 86fd19efb..8c49c6b19 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsn_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsn_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L1 ⦄ ⊢ ➡ break [ term 46 o , break term 46 h ] break term 46 L2 )" non associative with precedence 45 - for @{ 'PRedSn $o $h $G $L1 $L2 }. + for @{ 'PRedSn $h $o $G $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsnstar_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsnstar_5.ma index f7b69c931..3b12bc0e9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsnstar_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predsnstar_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L1 ⦄ ⊢ ➡ * break [ term 46 o , break term 46 h ] break term 46 L2 )" non associative with precedence 45 - for @{ 'PRedSnStar $o $h $G $L1 $L2 }. + for @{ 'PRedSnStar $h $o $G $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_6.ma index 7991322a3..b336dd9d7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_6.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ➡ * break [ term 46 o , break term 46 h ] break term 46 T2 )" non associative with precedence 45 - for @{ 'PRedStar $o $h $G $L $T1 $T2 }. + for @{ 'PRedStar $h $o $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/rdropstar_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/rdropstar_3.ma new file mode 100644 index 000000000..32b9c8445 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/rdropstar_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 i ] break term 46 L1 ≡ break term 46 L2 )" + non associative with precedence 45 + for @{ 'RDropStar $i $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_5.ma index ac83d2797..050140b69 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ⬊ * break [ term 46 o, break term 46 h ] break term 46 T )" non associative with precedence 45 - for @{ 'SN $o $h $G $L $T }. + for @{ 'SN $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_6.ma index 00db2bfe0..ea018c41b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/sn_6.ma @@ -16,4 +16,4 @@ notation "hvbox( G ⊢ ⬊ * break [ term 46 o , break term 46 h , break term 46 T , break term 46 f ] break term 46 L )" non associative with precedence 45 - for @{ 'SN $o $h $T $f $G $L }. + for @{ 'SN $h $o $T $f $G $L }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_5.ma index c64d194e3..0c624b913 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_5.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_5.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ ⬊ ⬊ * break [ term 46 o , break term 46 h ] break term 46 T )" non associative with precedence 45 - for @{ 'SNAlt $o $h $G $L $T }. + for @{ 'SNAlt $h $o $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_6.ma index 251ae4158..09e6918e0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/snalt_6.ma @@ -16,4 +16,4 @@ notation "hvbox( G ⊢ ⬊ ⬊ * break [ term 46 o , break term 46 h , break term 46 T , break term 46 f ] break term 46 L )" non associative with precedence 45 - for @{ 'SNAlt $o $h $T $f $G $L }. + for @{ 'SNAlt $h $o $T $f $G $L }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/statictypestar_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/statictypestar_6.ma index 369838563..e50f58275 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/statictypestar_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/statictypestar_6.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 •* break [ term 46 o , break term 46 n ] break term 46 T2 )" non associative with precedence 45 - for @{ 'StaticTypeStar $o $G $L $n $T1 $T2 }. + for @{ 'StaticTypeStar $h $G $L $n $T1 $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 deleted file mode 100644 index 7748c1836..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_6.ma +++ /dev/null @@ -1,19 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 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/reduction/cir.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma index ba5ba7e7e..412297546 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma @@ -58,7 +58,7 @@ qed-. (* Basic properties *********************************************************) -lemma cir_sort: ∀G,L,k. ⦃G, L⦄ ⊢ ➡ 𝐈⦃⋆k⦄. +lemma cir_sort: ∀G,L,s. ⦃G, L⦄ ⊢ ➡ 𝐈⦃⋆s⦄. /2 width=4 by crr_inv_sort/ qed. lemma cir_gref: ∀G,L,p. ⦃G, L⦄ ⊢ ➡ 𝐈⦃§p⦄. 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 68578e5f0..147d4b43c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_lift.ma @@ -19,10 +19,10 @@ include "basic_2/reduction/cir.ma". (* Properties on relocation *************************************************) -lemma cir_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄ → ∀L,s,l,m. ⬇[s, l, m] L ≡ K → - ∀U. ⬆[l, m] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄. +lemma cir_lift: ∀G,K,T. ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄ → ∀L,c,l,k. ⬇[c, l, k] L ≡ K → + ∀U. ⬆[l, k] T ≡ U → ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄. /3 width=8 by crr_inv_lift/ qed. -lemma cir_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄ → ∀K,s,l,m. ⬇[s, l, m] L ≡ K → - ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄. +lemma cir_inv_lift: ∀G,L,U. ⦃G, L⦄ ⊢ ➡ 𝐈⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K → + ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡ 𝐈⦃T⦄. /3 width=8 by crr_lift/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma index 9532052b8..5a7215ddf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma @@ -19,43 +19,43 @@ include "basic_2/reduction/crx.ma". (* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ***************) definition cix: ∀h. sd h → relation3 genv lenv term ≝ - λh,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ⊥. + λh,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ⊥. interpretation "irreducibility for context-sensitive extended reduction (term)" - 'PRedNotReducible h g G L T = (cix h g G L T). + 'PRedNotReducible h o G L T = (cix h o G L T). (* Basic inversion lemmas ***************************************************) -lemma cix_inv_sort: ∀h,g,G,L,k,d. deg h g k (d+1) → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃⋆k⦄ → ⊥. +lemma cix_inv_sort: ∀h,o,G,L,s,d. deg h o s (d+1) → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃⋆s⦄ → ⊥. /3 width=2 by crx_sort/ qed-. -lemma cix_inv_delta: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃#i⦄ → ⊥. +lemma cix_inv_delta: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃#i⦄ → ⊥. /3 width=4 by crx_delta/ qed-. -lemma cix_inv_ri2: ∀h,g,I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃②{I}V.T⦄ → ⊥. +lemma cix_inv_ri2: ∀h,o,I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃②{I}V.T⦄ → ⊥. /3 width=1 by crx_ri2/ qed-. -lemma cix_inv_ib2: ∀h,g,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃ⓑ{a,I}V.T⦄ → - ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, g] 𝐈⦃T⦄. +lemma cix_inv_ib2: ∀h,o,a,I,G,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄ → + ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ ∧ ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄. /4 width=1 by crx_ib2_sn, crx_ib2_dx, conj/ qed-. -lemma cix_inv_bind: ∀h,g,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃ⓑ{a,I}V.T⦄ → - ∧∧ ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ & ib2 a I. -#h #g #a * [ elim a -a ] +lemma cix_inv_bind: ∀h,o,a,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓑ{a,I}V.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L.ⓑ{I}V⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & ib2 a I. +#h #o #a * [ elim a -a ] #G #L #V #T #H [ elim H -H /3 width=1 by crx_ri2, or_introl/ ] elim (cix_inv_ib2 … H) -H /3 width=1 by and3_intro, or_introl/ qed-. -lemma cix_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃ⓐV.T⦄ → - ∧∧ ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ & 𝐒⦃T⦄. -#h #g #G #L #V #T #HVT @and3_intro /3 width=1 by crx_appl_sn, crx_appl_dx/ +lemma cix_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓐV.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & 𝐒⦃T⦄. +#h #o #G #L #V #T #HVT @and3_intro /3 width=1 by crx_appl_sn, crx_appl_dx/ generalize in match HVT; -HVT elim T -T // * // #a * #U #T #_ #_ #H elim H -H /2 width=1 by crx_beta, crx_theta/ qed-. -lemma cix_inv_flat: ∀h,g,I,G,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 * #G #L #V #T #H +lemma cix_inv_flat: ∀h,o,I,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃ⓕ{I}V.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl. +#h #o * #G #L #V #T #H [ elim (cix_inv_appl … H) -H /2 width=1 by and4_intro/ | elim (cix_inv_ri2 … H) -H // ] @@ -63,31 +63,31 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma cix_inv_cir: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄. +lemma cix_inv_cir: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄. /3 width=1 by crr_crx/ qed-. (* Basic properties *********************************************************) -lemma cix_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃⋆k⦄. -#h #g #G #L #k #Hk #H elim (crx_inv_sort … H) -L #d #Hkd -lapply (deg_mono … Hk Hkd) -h -k (cpr_inv_sort1 … H) // qed. 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 d280b1e83..eef9aefc8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_lift.ma @@ -30,20 +30,20 @@ qed. (* Relocation properties ****************************************************) (* Basic_1: was: nf2_lift *) -lemma cnr_lift: ∀G,L0,L,T,T0,s,l,m. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → - ⬇[s, l, m] L0 ≡ L → ⬆[l, m] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄. -#G #L0 #L #T #T0 #s #l #m #HLT #HL0 #HT0 #X #H +lemma cnr_lift: ∀G,L0,L,T,T0,c,l,k. ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → + ⬇[c, l, k] L0 ≡ L → ⬆[l, k] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄. +#G #L0 #L #T #T0 #c #l #k #HLT #HL0 #HT0 #X #H elim (cpr_inv_lift1 … H … HL0 … HT0) -L0 #T1 #HT10 #HT1 <(HLT … HT1) in HT0; -L #HT0 >(lift_mono … HT10 … HT0) -T1 -X // qed. (* Note: this was missing in basic_1 *) -lemma cnr_inv_lift: ∀G,L0,L,T,T0,s,l,m. ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄ → - ⬇[s, l, m] L0 ≡ L → ⬆[l, m] T ≡ T0 → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄. -#G #L0 #L #T #T0 #s #l #m #HLT0 #HL0 #HT0 #X #H -elim (lift_total X l m) #X0 #HX0 +lemma cnr_inv_lift: ∀G,L0,L,T,T0,c,l,k. ⦃G, L0⦄ ⊢ ➡ 𝐍⦃T0⦄ → + ⬇[c, l, k] L0 ≡ L → ⬆[l, k] T ≡ T0 → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄. +#G #L0 #L #T #T0 #c #l #k #HLT0 #HL0 #HT0 #X #H +elim (lift_total X l k) #X0 #HX0 lapply (cpr_lift … H … HL0 … HT0 … HX0) -L #HTX0 >(HLT0 … HTX0) in HX0; -L0 -X0 #H ->(lift_inj … H … HT0) -T0 -X -s -l -m // +>(lift_inj … H … HT0) -T0 -X -c -l -k // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma index 942ccc285..0259a08aa 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma @@ -19,48 +19,48 @@ include "basic_2/reduction/cpx.ma". (* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************) definition cnx: ∀h. sd h → relation3 genv lenv term ≝ - λh,g,G,L. NF … (cpx h g G L) (eq …). + λh,o,G,L. NF … (cpx h o G L) (eq …). interpretation "normality for context-sensitive extended reduction (term)" - 'PRedNormal h g L T = (cnx h g L T). + 'PRedNormal h o L T = (cnx h o L T). (* Basic inversion lemmas ***************************************************) -lemma cnx_inv_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆k⦄ → deg h g k 0. -#h #g #G #L #k #H elim (deg_total h g k) +lemma cnx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆s⦄ → deg h o s 0. +#h #o #G #L #s #H elim (deg_total h o s) #d @(nat_ind_plus … d) -d // #d #_ #Hkd -lapply (H (⋆(next h k)) ?) -H /2 width=2 by cpx_st/ -L -d #H +lapply (H (⋆(next h s)) ?) -H /2 width=2 by cpx_st/ -L -d #H lapply (destruct_tatom_tatom_aux … H) -H #H (**) (* destruct lemma needed *) lapply (destruct_sort_sort_aux … H) -H #H (**) (* destruct lemma needed *) -lapply (next_lt h k) >H -H #H elim (lt_refl_false … H) +lapply (next_lt h s) >H -H #H elim (lt_refl_false … H) qed-. -lemma cnx_inv_delta: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄ → ⊥. -#h #g #I #G #L #K #V #i #HLK #H +lemma cnx_inv_delta: ∀h,o,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄ → ⊥. +#h #o #I #G #L #K #V #i #HLK #H elim (lift_total V 0 (i+1)) #W #HVW lapply (H W ?) -H [ /3 width=7 by cpx_delta/ ] -HLK #H destruct elim (lift_inv_lref2_be … HVW) -HVW /2 width=1 by ylt_inj/ qed-. -lemma cnx_inv_abst: ∀h,g,a,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓛ{a}V.T⦄ → - ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ➡[h, g] 𝐍⦃T⦄. -#h #g #a #G #L #V1 #T1 #HVT1 @conj +lemma cnx_inv_abst: ∀h,o,a,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓛ{a}V.T⦄ → + ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ➡[h, o] 𝐍⦃T⦄. +#h #o #a #G #L #V1 #T1 #HVT1 @conj [ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct // | #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct // ] qed-. -lemma cnx_inv_abbr: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃-ⓓV.T⦄ → - ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ➡[h, g] 𝐍⦃T⦄. -#h #g #G #L #V1 #T1 #HVT1 @conj +lemma cnx_inv_abbr: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃-ⓓV.T⦄ → + ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ➡[h, o] 𝐍⦃T⦄. +#h #o #G #L #V1 #T1 #HVT1 @conj [ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 #H destruct // | #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 #H destruct // ] qed-. -lemma cnx_inv_zeta: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃+ⓓV.T⦄ → ⊥. -#h #g #G #L #V #T #H elim (is_lift_dec T 0 1) +lemma cnx_inv_zeta: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃+ⓓV.T⦄ → ⊥. +#h #o #G #L #V #T #H elim (is_lift_dec T 0 1) [ * #U #HTU lapply (H U ?) -H /2 width=3 by cpx_zeta/ #H destruct elim (lift_inv_pair_xy_y … HTU) @@ -71,9 +71,9 @@ lemma cnx_inv_zeta: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃+ⓓV.T⦄ ] qed-. -lemma cnx_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓐV.T⦄ → - ∧∧ ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ & 𝐒⦃T⦄. -#h #g #G #L #V1 #T1 #HVT1 @and3_intro +lemma cnx_inv_appl: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓐV.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ & ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ & 𝐒⦃T⦄. +#h #o #G #L #V1 #T1 #HVT1 @and3_intro [ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1 by cpx_pair_sn/ -HV2 #H destruct // | #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1 by cpx_flat/ -HT2 #H destruct // | generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H @@ -84,53 +84,53 @@ lemma cnx_inv_appl: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓐV.T⦄ ] qed-. -lemma cnx_inv_eps: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓝV.T⦄ → ⊥. -#h #g #G #L #V #T #H lapply (H T ?) -H +lemma cnx_inv_eps: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓝV.T⦄ → ⊥. +#h #o #G #L #V #T #H lapply (H T ?) -H /2 width=4 by cpx_eps, discr_tpair_xy_y/ qed-. (* Basic forward lemmas *****************************************************) -lemma cnx_fwd_cnr: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄. -#h #g #G #L #T #H #U #HTU +lemma cnx_fwd_cnr: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄. +#h #o #G #L #T #H #U #HTU @H /2 width=1 by cpr_cpx/ (**) (* auto fails because a δ-expansion gets in the way *) qed-. (* Basic properties *********************************************************) -lemma cnx_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆k⦄. -#h #g #G #L #k #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #d #Hkd #_ +lemma cnx_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃⋆s⦄. +#h #o #G #L #s #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #d #Hkd #_ lapply (deg_mono … Hkd Hk) -h -L (drop_fwd_length … HL) -HL // +lemma cnx_lref_atom: ∀h,o,G,L,i. ⬇[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃#i⦄. +#h #o #G #L #i #HL @cnx_lref_free >(drop_fwd_length … HL) -HL // qed. -lemma cnx_abst: ∀h,g,a,G,L,W,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → - ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓛ{a}W.T⦄. -#h #g #a #G #L #W #T #HW #HT #X #H +lemma cnx_abst: ∀h,o,a,G,L,W,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → + ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓛ{a}W.T⦄. +#h #o #a #G #L #W #T #HW #HT #X #H elim (cpx_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct >(HW … HW0) -W0 >(HT … HT0) -T0 // qed. -lemma cnx_appl_simple: ∀h,g,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → 𝐒⦃T⦄ → - ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃ⓐV.T⦄. -#h #g #G #L #V #T #HV #HT #HS #X #H +lemma cnx_appl_simple: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → 𝐒⦃T⦄ → + ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃ⓐV.T⦄. +#h #o #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,G,L,T1. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T1⦄ ∨ - ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & (T1 = T2 → ⊥). +axiom cnx_dec: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T1⦄ ∨ + ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] 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 a5ad258f0..32b9a76bf 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 irreducibility ****************************************) -theorem cix_cnx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄. +theorem cix_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄. /2 width=6 by cpx_fwd_cix/ qed. (* Main inversion lemmas on irreducibility **********************************) -theorem cnx_inv_cix: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃T⦄. +theorem cnx_inv_cix: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐈⦃T⦄. /2 width=7 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 ca66d9271..bb251965e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_crx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_crx.ma @@ -20,11 +20,11 @@ include "basic_2/reduction/cnx.ma". (* Advanced inversion lemmas on reducibility ********************************) (* Note: this property is unusual *) -lemma cnx_inv_crx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⊥. -#h #g #G #L #T #H elim H -L -T -[ #L #k #d #Hkd #H +lemma cnx_inv_crx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ⊥. +#h #o #G #L #T #H elim H -L -T +[ #L #s #d #Hkd #H lapply (cnx_inv_sort … H) -H #H - lapply (deg_mono … H Hkd) -h -L -k (lift_mono … HT10 … HT0) -T1 -X // qed. -lemma cnx_inv_lift: ∀h,g,G,L0,L,T,T0,s,l,m. ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄ → ⬇[s, l, m] L0 ≡ L → - ⬆[l, m] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄. -#h #g #G #L0 #L #T #T0 #s #l #m #HLT0 #HL0 #HT0 #X #H -elim (lift_total X l m) #X0 #HX0 +lemma cnx_inv_lift: ∀h,o,G,L0,L,T,T0,c,l,k. ⦃G, L0⦄ ⊢ ➡[h, o] 𝐍⦃T0⦄ → ⬇[c, l, k] L0 ≡ L → + ⬆[l, k] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄. +#h #o #G #L0 #L #T #T0 #c #l #k #HLT0 #HL0 #HT0 #X #H +elim (lift_total X l k) #X0 #HX0 lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0 >(HLT0 … HTX0) in HX0; -L0 -X0 #H ->(lift_inj … H … HT0) -T0 -X -l -m // +>(lift_inj … H … HT0) -T0 -X -l -k // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma index 27682aa74..d5374513a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -131,8 +131,8 @@ lemma cpr_inv_atom1: ∀I,G,L,T2. ⦃G, 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: ∀G,L,T2,k. ⦃G, L⦄ ⊢ ⋆k ➡ T2 → T2 = ⋆k. -#G #L #T2 #k #H +lemma cpr_inv_sort1: ∀G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡ T2 → T2 = ⋆s. +#G #L #T2 #s #H elim (cpr_inv_atom1 … H) -H // * #K #V #V2 #i #_ #_ #_ #H destruct qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma index 62418cd5d..18dcdc3bd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma @@ -23,9 +23,9 @@ include "basic_2/reduction/cpr.ma". (* Basic_1: includes: pr0_lift pr2_lift *) lemma cpr_lift: ∀G. d_liftable (cpr G). #G #K #T1 #T2 #H elim H -G -K -T1 -T2 -[ #I #G #K #L #s #l #m #_ #U1 #H1 #U2 #H2 +[ #I #G #K #L #c #l #k #_ #U1 #H1 #U2 #H2 >(lift_mono … H1 … H2) -H1 -H2 // -| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #s #l #m #HLK #U1 #H #U2 #HWU2 +| #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #c #l #k #HLK #U1 #H #U2 #HWU2 elim (lift_inv_lref1 … H) * #Hil #H destruct [ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2 elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H @@ -34,23 +34,23 @@ lemma cpr_lift: ∀G. d_liftable (cpr G). | lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2 lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil /3 width=6 by cpr_delta, drop_inv_gen/ ] -| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #HLK #U1 #H1 #U2 #H2 +| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2 elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpr_bind, drop_skip/ -| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #HLK #U1 #H1 #U2 #H2 +| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2 elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6 by cpr_flat/ -| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #H #U2 #HTU2 +| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #c #l #k #HLK #U1 #H #U2 #HTU2 elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpr_zeta, drop_skip/ -| #G #K #V #T1 #T2 #_ #IHT12 #L #s #l #m #HLK #U1 #H #U2 #HTU2 +| #G #K #V #T1 #T2 #_ #IHT12 #L #c #l #k #HLK #U1 #H #U2 #HTU2 elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpr_eps/ -| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #s #l #m #HLK #X1 #HX1 #X2 #HX2 +| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2 elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct elim (lift_inv_bind1 … HX2) -HX2 #X #T3 #HX #HT23 #HX2 destruct elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpr_beta, drop_skip/ -| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #s #l #m #HLK #X1 #HX1 #X2 #HX2 +| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2 elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct @@ -62,12 +62,12 @@ qed. (* Basic_1: includes: pr0_gen_lift pr2_gen_lift *) lemma cpr_inv_lift1: ∀G. d_deliftable_sn (cpr G). #G #L #U1 #U2 #H elim H -G -L -U1 -U2 -[ * #i #G #L #K #s #l #m #_ #T1 #H +[ * #i #G #L #K #c #l #k #_ #T1 #H [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ ] -| #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #s #l #m #HLK #T1 #H +| #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #c #l #k #HLK #T1 #H elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ elim (drop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2 @@ -75,36 +75,36 @@ lemma cpr_inv_lift1: ∀G. d_deliftable_sn (cpr G). | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi lapply (yle_inv_inj … Hmi) -Hmi #Hmi lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV - elim (lift_split … HVW2 l (i - m + 1)) -HVW2 [2,3,4: /2 width=1 by yle_succ_dx, le_S_S/ ] -Hil -Hlim + elim (lift_split … HVW2 l (i - k + 1)) -HVW2 [2,3,4: /2 width=1 by yle_succ_dx, le_S_S/ ] -Hil -Hlim #V1 #HV1 >plus_minus // (lift_mono … H1 … H2) -H1 -H2 // -| #G #K #k #d #Hkd #L #s #l #m #_ #U1 #H1 #U2 #H2 +| #G #K #s #d #Hkd #L #c #l #k #_ #U1 #H1 #U2 #H2 >(lift_inv_sort1 … H1) -U1 >(lift_inv_sort1 … H2) -U2 /2 width=2 by cpx_st/ -| #I #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #s #l #m #HLK #U1 #H #U2 #HWU2 +| #I #G #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #c #l #k #HLK #U1 #H #U2 #HWU2 elim (lift_inv_lref1 … H) * #Hil #H destruct [ elim (lift_trans_ge … HVW2 … HWU2) -W2 /2 width=1 by ylt_fwd_le_succ1/ #W2 #HVW2 #HWU2 elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H @@ -64,25 +64,25 @@ lemma cpx_lift: ∀h,g,G. d_liftable (cpx h g G). | lapply (lift_trans_be … HVW2 … HWU2 ? ?) -W2 /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2 lapply (drop_trans_ge_comm … HLK … HKV ?) -K /3 width=7 by cpx_delta, drop_inv_gen/ ] -| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #HLK #U1 #H1 #U2 #H2 +| #a #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2 elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=6 by cpx_bind, drop_skip/ -| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #s #l #m #HLK #U1 #H1 #U2 #H2 +| #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #c #l #k #HLK #U1 #H1 #U2 #H2 elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6 by cpx_flat/ -| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #H #U2 #HTU2 +| #G #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #c #l #k #HLK #U1 #H #U2 #HTU2 elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=6 by cpx_zeta, drop_skip/ -| #G #K #V #T1 #T2 #_ #IHT12 #L #s #l #m #HLK #U1 #H #U2 #HTU2 +| #G #K #V #T1 #T2 #_ #IHT12 #L #c #l #k #HLK #U1 #H #U2 #HTU2 elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpx_eps/ -| #G #K #V1 #V2 #T #_ #IHV12 #L #s #l #m #HLK #U1 #H #U2 #HVU2 +| #G #K #V1 #V2 #T #_ #IHV12 #L #c #l #k #HLK #U1 #H #U2 #HVU2 elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=6 by cpx_ct/ -| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #s #l #m #HLK #X1 #HX1 #X2 #HX2 +| #a #G #K #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #IHV12 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2 elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct elim (lift_inv_bind1 … HX2) -HX2 #X #T3 #HX #HT23 #HX2 destruct elim (lift_inv_flat1 … HX) -HX #W3 #V3 #HW23 #HV23 #HX destruct /4 width=6 by cpx_beta, drop_skip/ -| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #s #l #m #HLK #X1 #HX1 #X2 #HX2 +| #a #G #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #c #l #k #HLK #X1 #HX1 #X2 #HX2 elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct @@ -91,16 +91,16 @@ lemma cpx_lift: ∀h,g,G. d_liftable (cpx h g G). ] qed. -lemma cpx_inv_lift1: ∀h,g,G. d_deliftable_sn (cpx h g G). -#h #g #G #L #U1 #U2 #H elim H -G -L -U1 -U2 -[ * #i #G #L #K #s #l #m #_ #T1 #H +lemma cpx_inv_lift1: ∀h,o,G. d_deliftable_sn (cpx h o G). +#h #o #G #L #U1 #U2 #H elim H -G -L -U1 -U2 +[ * #i #G #L #K #c #l #k #_ #T1 #H [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by cpx_atom, lift_sort, ex2_intro/ | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by cpx_atom, lift_lref_ge_minus, lift_lref_lt, ex2_intro/ | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by cpx_atom, lift_gref, ex2_intro/ ] -| #G #L #k #d #Hkd #K #s #l #m #_ #T1 #H +| #G #L #s #d #Hkd #K #c #l #k #_ #T1 #H lapply (lift_inv_sort2 … H) -H #H destruct /3 width=3 by cpx_st, lift_sort, ex2_intro/ -| #I #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #s #l #m #HLK #T1 #H +| #I #G #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #c #l #k #HLK #T1 #H elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ elim (drop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2 @@ -108,39 +108,39 @@ lemma cpx_inv_lift1: ∀h,g,G. d_deliftable_sn (cpx h g G). | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi lapply (yle_inv_inj … Hmi) -Hmi #Hmi lapply (drop_conf_ge … HLK … HLV ?) -L // #HKLV - elim (lift_split … HVW2 l (i - m + 1)) -HVW2 /3 width=1 by yle_succ, yle_pred_sn, le_S_S/ -Hil -Hlim + elim (lift_split … HVW2 l (i - k + 1)) -HVW2 /3 width=1 by yle_succ, yle_pred_sn, le_S_S/ -Hil -Hlim #V1 #HV1 >plus_minus // (lift_inv_sort1 … H) -X /2 width=2 by crx_sort/ -| #I #K #K0 #V #i #HK0 #L #s #l #m #HLK #X #H +| #I #K #K0 #V #i #HK0 #L #c #l #k #HLK #X #H elim (lift_inv_lref1 … H) -H * #Hil #H destruct [ elim (drop_trans_lt … HLK … HK0) -K /2 width=4 by crx_delta/ | lapply (drop_trans_ge … HLK … HK0 ?) -K /3 width=5 by crx_delta, drop_inv_gen/ ] -| #K #V #T #_ #IHV #L #s #l #m #HLK #X #H +| #K #V #T #_ #IHV #L #c #l #k #HLK #X #H elim (lift_inv_flat1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crx_appl_sn/ -| #K #V #T #_ #IHT #L #s #l #m #HLK #X #H +| #K #V #T #_ #IHT #L #c #l #k #HLK #X #H elim (lift_inv_flat1 … H) -H #W #U #_ #HTU #H destruct /3 width=5 by crx_appl_dx/ -| #I #K #V #T #HI #L #s #l #m #_ #X #H +| #I #K #V #T #HI #L #c #l #k #_ #X #H elim (lift_fwd_pair1 … H) -H #W #U #_ #H destruct /2 width=1 by crx_ri2/ -| #a #I #K #V #T #HI #_ #IHV #L #s #l #m #HLK #X #H +| #a #I #K #V #T #HI #_ #IHV #L #c #l #k #HLK #X #H elim (lift_inv_bind1 … H) -H #W #U #HVW #_ #H destruct /3 width=5 by crx_ib2_sn/ -| #a #I #K #V #T #HI #_ #IHT #L #s #l #m #HLK #X #H +| #a #I #K #V #T #HI #_ #IHT #L #c #l #k #HLK #X #H elim (lift_inv_bind1 … H) -H #W #U #HVW #HTU #H destruct /4 width=5 by crx_ib2_dx, drop_skip/ -| #a #K #V #V0 #T #L #s #l #m #_ #X #H +| #a #K #V #V0 #T #L #c #l #k #_ #X #H elim (lift_inv_flat1 … H) -H #W #X0 #_ #H0 #H destruct elim (lift_inv_bind1 … H0) -H0 #W0 #U #_ #_ #H0 destruct /2 width=1 by crx_beta/ -| #a #K #V #V0 #T #L #s #l #m #_ #X #H +| #a #K #V #V0 #T #L #c #l #k #_ #X #H elim (lift_inv_flat1 … H) -H #W #X0 #_ #H0 #H destruct elim (lift_inv_bind1 … H0) -H0 #W0 #U #_ #_ #H0 destruct /2 width=1 by crx_theta/ ] qed. -lemma crx_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃U⦄ → ∀K,s,l,m. ⬇[s, l, m] L ≡ K → - ∀T. ⬆[l, m] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, g] 𝐑⦃T⦄. -#h #g #G #L #U #H elim H -L -U -[ #L #k #d #Hkd #K #s #l #m #_ #X #H +lemma crx_inv_lift: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ➡[h, o] 𝐑⦃U⦄ → ∀K,c,l,k. ⬇[c, l, k] L ≡ K → + ∀T. ⬆[l, k] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, o] 𝐑⦃T⦄. +#h #o #G #L #U #H elim H -L -U +[ #L #s #d #Hkd #K #c #l #k #_ #X #H >(lift_inv_sort2 … H) -X /2 width=2 by crx_sort/ -| #I #L #L0 #W #i #HK0 #K #s #l #m #HLK #X #H +| #I #L #L0 #W #i #HK0 #K #c #l #k #HLK #X #H elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ elim (drop_conf_lt … HLK … HK0) -L /2 width=4 by crx_delta/ | lapply (drop_conf_ge … HLK … HK0 ?) -L /2 width=4 by crx_delta/ ] -| #L #W #U #_ #IHW #K #s #l #m #HLK #X #H +| #L #W #U #_ #IHW #K #c #l #k #HLK #X #H elim (lift_inv_flat2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crx_appl_sn/ -| #L #W #U #_ #IHU #K #s #l #m #HLK #X #H +| #L #W #U #_ #IHU #K #c #l #k #HLK #X #H elim (lift_inv_flat2 … H) -H #V #T #_ #HTU #H destruct /3 width=5 by crx_appl_dx/ -| #I #L #W #U #HI #K #s #l #m #_ #X #H +| #I #L #W #U #HI #K #c #l #k #_ #X #H elim (lift_fwd_pair2 … H) -H #V #T #_ #H destruct /2 width=1 by crx_ri2/ -| #a #I #L #W #U #HI #_ #IHW #K #s #l #m #HLK #X #H +| #a #I #L #W #U #HI #_ #IHW #K #c #l #k #HLK #X #H elim (lift_inv_bind2 … H) -H #V #T #HVW #_ #H destruct /3 width=5 by crx_ib2_sn/ -| #a #I #L #W #U #HI #_ #IHU #K #s #l #m #HLK #X #H +| #a #I #L #W #U #HI #_ #IHU #K #c #l #k #HLK #X #H elim (lift_inv_bind2 … H) -H #V #T #HVW #HTU #H destruct /4 width=5 by crx_ib2_dx, drop_skip/ -| #a #L #W #W0 #U #K #s #l #m #_ #X #H +| #a #L #W #W0 #U #K #c #l #k #_ #X #H elim (lift_inv_flat2 … H) -H #V #X0 #_ #H0 #H destruct elim (lift_inv_bind2 … H0) -H0 #V0 #T #_ #_ #H0 destruct /2 width=1 by crx_beta/ -| #a #L #W #W0 #U #K #s #l #m #_ #X #H +| #a #L #W #W0 #U #K #c #l #k #_ #X #H elim (lift_inv_flat2 … H) -H #V #X0 #_ #H0 #H destruct elim (lift_inv_bind2 … H0) -H0 #V0 #T #_ #_ #H0 destruct /2 width=1 by crx_theta/ ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma index 8872199fd..cce7d39ff 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma @@ -19,22 +19,22 @@ include "basic_2/reduction/lpx.ma". (* "RST" PROPER PARALLEL COMPUTATION FOR CLOSURES ***************************) -inductive fpb (h) (g) (G1) (L1) (T1): relation3 genv lenv term ≝ -| fpb_fqu: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → fpb h g G1 L1 T1 G2 L2 T2 -| fpb_cpx: ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → fpb h g G1 L1 T1 G1 L1 T2 -| fpb_lpx: ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, g] L2 → (L1 ≡[T1, 0] L2 → ⊥) → fpb h g G1 L1 T1 G1 L2 T1 +inductive fpb (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝ +| fpb_fqu: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → fpb h o G1 L1 T1 G2 L2 T2 +| fpb_cpx: ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → (T1 = T2 → ⊥) → fpb h o G1 L1 T1 G1 L1 T2 +| fpb_lpx: ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → (L1 ≡[T1, 0] L2 → ⊥) → fpb h o G1 L1 T1 G1 L2 T1 . interpretation "'rst' proper parallel reduction (closure)" - 'BTPRedProper h g G1 L1 T1 G2 L2 T2 = (fpb h g G1 L1 T1 G2 L2 T2). + 'BTPRedProper h o G1 L1 T1 G2 L2 T2 = (fpb h o G1 L1 T1 G2 L2 T2). (* Basic properties *********************************************************) -lemma cpr_fpb: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) → - ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄. +lemma cpr_fpb: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) → + ⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄. /3 width=1 by fpb_cpx, cpr_cpx/ qed. -lemma lpr_fpb: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → (L1 ≡[T, 0] L2 → ⊥) → - ⦃G, L1, T⦄ ≻[h, g] ⦃G, L2, T⦄. +lemma lpr_fpb: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → (L1 ≡[T, 0] L2 → ⊥) → + ⦃G, L1, T⦄ ≻[h, o] ⦃G, L2, T⦄. /3 width=1 by fpb_lpx, lpr_lpx/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_fleq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_fleq.ma index 788a59364..6463870f1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_fleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_fleq.ma @@ -19,19 +19,19 @@ include "basic_2/reduction/fpb_lleq.ma". (* Properties on lazy equivalence for closures ******************************) -lemma fleq_fpb_trans: ∀h,g,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≡[0] ⦃F2, K2, T2⦄ → - ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, g] ⦃G2, L2, U2⦄ → - ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, g] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≡[0] ⦃G2, L2, U2⦄. -#h #g #F1 #F2 #K1 #K2 #T1 #T2 * -F2 -K2 -T2 +lemma fleq_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≡[0] ⦃F2, K2, T2⦄ → + ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ → + ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≡[0] ⦃G2, L2, U2⦄. +#h #o #F1 #F2 #K1 #K2 #T1 #T2 * -F2 -K2 -T2 #K2 #HK12 #G2 #L2 #U2 #H12 elim (lleq_fpb_trans … HK12 … H12) -K2 /3 width=5 by fleq_intro, ex2_3_intro/ qed-. (* Inversion lemmas on lazy equivalence for closures ************************) -lemma fpb_inv_fleq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → +lemma fpb_inv_fleq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 [ #G2 #L2 #T2 #H12 #H elim (fleq_inv_gen … H) -H /3 width=8 by lleq_fwd_length, fqu_inv_eq/ | #T2 #_ #HnT #H elim (fleq_inv_gen … H) -H /2 width=1 by/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lift.ma index 8efea75f0..6fae399d9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lift.ma @@ -20,9 +20,9 @@ include "basic_2/reduction/fpb.ma". (* Advanced properties ******************************************************) -lemma sta_fpb: ∀h,g,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 → - ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ ≻[h, g] ⦃G, L, T2⦄. -#h #g #G #L #T1 #T2 #d #HT1 #HT12 elim (eq_term_dec T1 T2) +lemma sta_fpb: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → + ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → ⦃G, L, T1⦄ ≻[h, o] ⦃G, L, T2⦄. +#h #o #G #L #T1 #T2 #d #HT1 #HT12 elim (eq_term_dec T1 T2) /3 width=2 by fpb_cpx, sta_cpx/ #H destruct elim (lstas_inv_refl_pos h G L T2 0) // qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lleq.ma index 519880ff8..19e8d009a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb_lleq.ma @@ -21,10 +21,10 @@ include "basic_2/reduction/fpb.ma". (* Properties on lazy equivalence for local environments ********************) -lemma lleq_fpb_trans: ∀h,g,F,K1,K2,T. K1 ≡[T, 0] K2 → - ∀G,L2,U. ⦃F, K2, T⦄ ≻[h, g] ⦃G, L2, U⦄ → - ∃∃L1. ⦃F, K1, T⦄ ≻[h, g] ⦃G, L1, U⦄ & L1 ≡[U, 0] L2. -#h #g #F #K1 #K2 #T #HT #G #L2 #U * -G -L2 -U +lemma lleq_fpb_trans: ∀h,o,F,K1,K2,T. K1 ≡[T, 0] K2 → + ∀G,L2,U. ⦃F, K2, T⦄ ≻[h, o] ⦃G, L2, U⦄ → + ∃∃L1. ⦃F, K1, T⦄ ≻[h, o] ⦃G, L1, U⦄ & L1 ≡[U, 0] L2. +#h #o #F #K1 #K2 #T #HT #G #L2 #U * -G -L2 -U [ #G #L2 #U #H2 elim (lleq_fqu_trans … H2 … HT) -K2 /3 width=3 by fpb_fqu, ex2_intro/ | /4 width=10 by fpb_cpx, cpx_lleq_conf_sn, lleq_cpx_trans, ex2_intro/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq.ma index 3e417b723..c8be98384 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq.ma @@ -19,24 +19,24 @@ include "basic_2/reduction/lpx.ma". (* "QRST" PARALLEL REDUCTION FOR CLOSURES ***********************************) -inductive fpbq (h) (g) (G1) (L1) (T1): relation3 genv lenv term ≝ -| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h g G1 L1 T1 G2 L2 T2 -| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T2 → fpbq h g G1 L1 T1 G1 L1 T2 -| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, g] L2 → fpbq h g G1 L1 T1 G1 L2 T1 -| fpbq_lleq: ∀L2. L1 ≡[T1, 0] L2 → fpbq h g G1 L1 T1 G1 L2 T1 +inductive fpbq (h) (o) (G1) (L1) (T1): relation3 genv lenv term ≝ +| fpbq_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → fpbq h o G1 L1 T1 G2 L2 T2 +| fpbq_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T2 → fpbq h o G1 L1 T1 G1 L1 T2 +| fpbq_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, o] L2 → fpbq h o G1 L1 T1 G1 L2 T1 +| fpbq_lleq: ∀L2. L1 ≡[T1, 0] L2 → fpbq h o G1 L1 T1 G1 L2 T1 . interpretation "'qrst' parallel reduction (closure)" - 'BTPRed h g G1 L1 T1 G2 L2 T2 = (fpbq h g G1 L1 T1 G2 L2 T2). + 'BTPRed h o G1 L1 T1 G2 L2 T2 = (fpbq h o G1 L1 T1 G2 L2 T2). (* Basic properties *********************************************************) -lemma fpbq_refl: ∀h,g. tri_reflexive … (fpbq h g). +lemma fpbq_refl: ∀h,o. tri_reflexive … (fpbq h o). /2 width=1 by fpbq_cpx/ qed. -lemma cpr_fpbq: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L, T1⦄ ≽[h, g] ⦃G, L, T2⦄. +lemma cpr_fpbq: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄. /3 width=1 by fpbq_cpx, cpr_cpx/ qed. -lemma lpr_fpbq: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1, T⦄ ≽[h, g] ⦃G, L2, T⦄. +lemma lpr_fpbq: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1, T⦄ ≽[h, o] ⦃G, L2, T⦄. /3 width=1 by fpbq_lpx, lpr_lpx/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_aaa.ma index 0f3d7fbc6..503216c57 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_aaa.ma @@ -21,8 +21,8 @@ include "basic_2/reduction/fpbq.ma". (* Properties on atomic arity assignment for terms **************************) -lemma fpbq_aaa_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → +lemma fpbq_aaa_conf: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → ∀A1. ⦃G1, L1⦄ ⊢ T1 ⁝ A1 → ∃A2. ⦃G2, L2⦄ ⊢ T2 ⁝ A2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /3 width=6 by aaa_lleq_conf, lpx_aaa_conf, cpx_aaa_conf, aaa_fquq_conf, ex_intro/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_alt.ma index 8bc5c210d..2a6cf5087 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_alt.ma @@ -20,32 +20,32 @@ include "basic_2/reduction/fpbq.ma". (* alternative definition of fpbq *) definition fpbqa: ∀h. sd h → tri_relation genv lenv term ≝ - λh,g,G1,L1,T1,G2,L2,T2. - ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄. + λh,o,G1,L1,T1,G2,L2,T2. + ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄. interpretation "'qrst' parallel reduction (closure) alternative" - 'BTPRedAlt h g G1 L1 T1 G2 L2 T2 = (fpbqa h g G1 L1 T1 G2 L2 T2). + 'BTPRedAlt h o G1 L1 T1 G2 L2 T2 = (fpbqa h o G1 L1 T1 G2 L2 T2). (* Basic properties *********************************************************) -lemma fleq_fpbq: ∀h,g,G1,G2,L1,L2,T1,T2. - ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/ +lemma fleq_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. + ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fpbq_lleq/ qed. -lemma fpb_fpbq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +lemma fpb_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /3 width=1 by fpbq_fquq, fpbq_cpx, fpbq_lpx, fqu_fquq/ qed. (* Main properties **********************************************************) -theorem fpbq_fpbqa: ∀h,g,G1,G2,L1,L2,T1,T2. - ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ≽≽[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 +theorem fpbq_fpbqa: ∀h,o,G1,G2,L1,L2,T1,T2. + ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 [ #G2 #L2 #T2 #H elim (fquq_inv_gen … H) -H [ /3 width=1 by fpb_fqu, or_intror/ | * #H1 #H2 #H3 destruct /2 width=1 by or_introl/ @@ -58,29 +58,29 @@ theorem fpbq_fpbqa: ∀h,g,G1,G2,L1,L2,T1,T2. ] qed. -theorem fpbqa_inv_fpbq: ∀h,g,G1,G2,L1,L2,T1,T2. - ⦃G1, L1, T1⦄ ≽≽[h, g] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fleq_fpbq, fpb_fpbq/ +theorem fpbqa_inv_fpbq: ∀h,o,G1,G2,L1,L2,T1,T2. + ⦃G1, L1, T1⦄ ≽≽[h, o] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=1 by fleq_fpbq, fpb_fpbq/ qed-. (* Advanced eliminators *****************************************************) -lemma fpbq_ind_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ∀R:Prop. +lemma fpbq_ind_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ∀R:Prop. (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → R) → - (⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → R) → - ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → R. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/ + (⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → R) → + ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → R. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #R #HI1 #HI2 #H elim (fpbq_fpbqa … H) /2 width=1 by/ qed-. (* aternative definition of fpb *********************************************) -lemma fpb_fpbq_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥). +lemma fpb_fpbq_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ ∧ (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥). /3 width=10 by fpb_fpbq, fpb_inv_fleq, conj/ qed. -lemma fpbq_inv_fpb_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → - (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H // +lemma fpbq_inv_fpb_alt: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≽[h, o] ⦃G2, L2, T2⦄ → + (⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ → ⊥) → ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #H0 @(fpbq_ind_alt … H) -H // #H elim H0 -H0 // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_lift.ma index 9d2846331..4029d9417 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpbq_lift.ma @@ -19,7 +19,7 @@ include "basic_2/reduction/fpbq.ma". (* Advanced properties ******************************************************) -lemma sta_fpbq: ∀h,g,G,L,T1,T2,d. - ⦃G, L⦄ ⊢ T1 ▪[h, g] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → - ⦃G, L, T1⦄ ≽[h, g] ⦃G, L, T2⦄. +lemma sta_fpbq: ∀h,o,G,L,T1,T2,d. + ⦃G, L⦄ ⊢ T1 ▪[h, o] d+1 → ⦃G, L⦄ ⊢ T1 •*[h, 1] T2 → + ⦃G, L, T1⦄ ≽[h, o] ⦃G, L, T2⦄. /3 width=4 by fpbq_cpx, sta_cpx/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fqu.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fqu.ma deleted file mode 100644 index babf330d2..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fqu.ma +++ /dev/null @@ -1,96 +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/supterm_6.ma". -include "basic_2/grammar/cl_weight.ma". -include "basic_2/substitution/drop.ma". - -(* SUPCLOSURE ***************************************************************) - -(* activate genv *) -inductive fqu: tri_relation genv lenv term ≝ -| fqu_lref_O : ∀I,G,L,V. fqu G (L.ⓑ{I}V) (#0) G L V -| fqu_pair_sn: ∀I,G,L,V,T. fqu G L (②{I}V.T) G L V -| fqu_bind_dx: ∀a,I,G,L,V,T. fqu G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T -| fqu_flat_dx: ∀I,G,L,V,T. fqu G L (ⓕ{I}V.T) G L T -| fqu_drop : ∀G,L,K,T,U,m. - ⬇[⫯m] L ≡ K → ⬆[0, ⫯m] T ≡ U → fqu G L U G K T -. - -interpretation - "structural successor (closure)" - 'SupTerm G1 L1 T1 G2 L2 T2 = (fqu G1 L1 T1 G2 L2 T2). - -(* Basic properties *********************************************************) - -lemma fqu_drop_lt: ∀G,L,K,T,U,m. 0 < m → - ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → ⦃G, L, U⦄ ⊐ ⦃G, K, T⦄. -#G #L #K #T #U #m #Hm lapply (ylt_inv_O1 … Hm) -Hm -#Hm (drop_fwd_length … HLK) in H; -L - #H elim (discr_yplus_xy_x … H) -H /2 width=2 by ysucc_inv_O_sn/ - #H elim (ylt_yle_false (|K|) (∞)) // -] -/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x/ -qed-. - -lemma fqu_inv_eq: ∀G,L1,L2,T. ⦃G, L1, T⦄ ⊐ ⦃G, L2, T⦄ → |L1| = |L2| → ⊥. -#G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *) -qed-. - -(* Advanced eliminators *****************************************************) - -lemma fqu_wf_ind: ∀R:relation3 …. ( - ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) → - R G1 L1 T1 - ) → ∀G1,L1,T1. R G1 L1 T1. -#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqu_fwd_fw/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq.ma deleted file mode 100644 index 914af3702..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq.ma +++ /dev/null @@ -1,64 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/suptermopt_6.ma". -include "basic_2/substitution/fqu.ma". - -(* OPTIONAL SUPCLOSURE ******************************************************) - -(* activate genv *) -inductive fquq: tri_relation genv lenv term ≝ -| fquq_lref_O : ∀I,G,L,V. fquq G (L.ⓑ{I}V) (#0) G L V -| fquq_pair_sn: ∀I,G,L,V,T. fquq G L (②{I}V.T) G L V -| fquq_bind_dx: ∀a,I,G,L,V,T. fquq G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T -| fquq_flat_dx: ∀I,G, L,V,T. fquq G L (ⓕ{I}V.T) G L T -| fquq_drop : ∀G,L,K,T,U,m. - ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → fquq G L U G K T -. - -interpretation - "optional structural successor (closure)" - 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fquq G1 L1 T1 G2 L2 T2). - -(* Basic properties *********************************************************) - -lemma fquq_refl: tri_reflexive … fquq. -/2 width=3 by fquq_drop/ qed. - -lemma fqu_fquq: ∀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=3 by fquq_drop/ -qed. - -(* Basic forward lemmas *****************************************************) - -lemma fquq_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 /2 width=1 by lt_to_le/ -#G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 -lapply (drop_fwd_lw … HLK1) -HLK1 -lapply (lift_fwd_tw … HTU1) -HTU1 -/2 width=1 by le_plus, le_n/ -qed-. - -fact fquq_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀i. T1 = #i → |L2| ≤ |L1|. -#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 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #i #H destruct - /2 width=3 by drop_fwd_length_le4/ -] -qed-. - -lemma fquq_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐⸮ ⦃G2, L2, T2⦄ → |L2| ≤ |L1|. -/2 width=7 by fquq_fwd_length_lref1_aux/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq_alt.ma deleted file mode 100644 index 7d7f05bd9..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq_alt.ma +++ /dev/null @@ -1,58 +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/suptermoptalt_6.ma". -include "basic_2/substitution/fquq.ma". - -(* OPTIONAL SUPCLOSURE ******************************************************) - -(* alternative definition of fquq *) -definition fquqa: tri_relation genv lenv term ≝ tri_RC … fqu. - -interpretation - "optional structural successor (closure) alternative" - 'SupTermOptAlt G1 L1 T1 G2 L2 T2 = (fquqa G1 L1 T1 G2 L2 T2). - -(* Basic properties *********************************************************) - -lemma fquqa_refl: tri_reflexive … fquqa. -// qed. - -lemma fquqa_drop: ∀G,L,K,T,U,m. - ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → ⦃G, L, U⦄ ⊐⊐⸮ ⦃G, K, T⦄. -#G #L #K #T #U #m @(ynat_ind … m) -m /3 width=3 by fqu_drop, or_introl/ -#HLK #HTU >(drop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T // -qed. - -(* Main properties **********************************************************) - -theorem fquq_fquqa: ∀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=3 by fquqa_drop, fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, or_introl/ -qed. - -(* Main inversion properties ************************************************) - -theorem fquqa_inv_fquq: ∀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 by fqu_fquq/ -* #H1 #H2 #H3 destruct // -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma fquq_inv_gen: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2). -#G1 #G2 #L1 #L2 #T1 #T2 #H elim (fquq_fquqa … H) -H [| * ] -/2 width=1 by or_introl/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_drop.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_drop.ma index 8254c0bde..5ec500e5f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_drop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_drop.ma @@ -39,8 +39,8 @@ lemma fqu_cpr_trans_dx: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2 ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄. #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/ -#G #L #K #U #T #m #HLK #HUT #U2 #HU2 -elim (lift_total U2 0 (m+1)) #T2 #HUT2 +#G #L #K #U #T #k #HLK #HUT #U2 #HU2 +elim (lift_total U2 0 (k+1)) #T2 #HUT2 lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/ qed-. @@ -58,8 +58,8 @@ lemma fqu_cpr_trans_sn: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2 ∃∃L,U1. ⦃G1, L1⦄ ⊢ ➡ L & ⦃G1, L1⦄ ⊢ T1 ➡ U1 & ⦃G1, L, U1⦄ ⊐ ⦃G2, L2, U2⦄. #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpr_pair, cpr_pair_sn, cpr_atom, cpr_bind, cpr_flat, ex3_2_intro/ -#G #L #K #U #T #m #HLK #HUT #U2 #HU2 -elim (lift_total U2 0 (m+1)) #T2 #HUT2 +#G #L #K #U #T #k #HLK #HUT #U2 #HU2 +elim (lift_total U2 0 (k+1)) #T2 #HUT2 lapply (cpr_lift … HU2 … HLK … HUT … HUT2) -HU2 -HUT /3 width=9 by fqu_drop, ex3_2_intro/ qed-. @@ -80,7 +80,7 @@ lemma fqu_lpr_trans: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ [ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpr_inv_pair1 … H) -H #K2 #W2 #HLK2 #HVW2 #H destruct /3 width=5 by fqu_fquq, cpr_pair_sn, fqu_bind_dx, ex3_2_intro/ -| #G #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #K2 #HK12 +| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #K2 #HK12 elim (drop_lpr_trans … HLK1 … HK12) -HK12 /3 width=7 by fqu_drop, ex3_2_intro/ ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma index b10b13a28..bd4e16df9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma @@ -19,47 +19,47 @@ include "basic_2/reduction/cpx.ma". (* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************) definition lpx: ∀h. sd h → relation3 genv lenv lenv ≝ - λh,g,G. lpx_sn (cpx h g G). + λh,o,G. lpx_sn (cpx h o G). interpretation "extended parallel reduction (local environment, sn variant)" - 'PRedSn h g G L1 L2 = (lpx h g G L1 L2). + 'PRedSn h o G L1 L2 = (lpx h o G L1 L2). (* Basic inversion lemmas ***************************************************) -lemma lpx_inv_atom1: ∀h,g,G,L2. ⦃G, ⋆⦄ ⊢ ➡[h, g] L2 → L2 = ⋆. +lemma lpx_inv_atom1: ∀h,o,G,L2. ⦃G, ⋆⦄ ⊢ ➡[h, o] L2 → L2 = ⋆. /2 width=4 by lpx_sn_inv_atom1_aux/ qed-. -lemma lpx_inv_pair1: ∀h,g,I,G,K1,V1,L2. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, g] L2 → - ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡[h, g] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, g] V2 & +lemma lpx_inv_pair1: ∀h,o,I,G,K1,V1,L2. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, o] L2 → + ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 & L2 = K2. ⓑ{I} V2. /2 width=3 by lpx_sn_inv_pair1_aux/ qed-. -lemma lpx_inv_atom2: ∀h,g,G,L1. ⦃G, L1⦄ ⊢ ➡[h, g] ⋆ → L1 = ⋆. +lemma lpx_inv_atom2: ∀h,o,G,L1. ⦃G, L1⦄ ⊢ ➡[h, o] ⋆ → L1 = ⋆. /2 width=4 by lpx_sn_inv_atom2_aux/ qed-. -lemma lpx_inv_pair2: ∀h,g,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡[h, g] K2.ⓑ{I}V2 → - ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡[h, g] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, g] V2 & +lemma lpx_inv_pair2: ∀h,o,I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡[h, o] K2.ⓑ{I}V2 → + ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡[h, o] K2 & ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 & L1 = K1. ⓑ{I} V1. /2 width=3 by lpx_sn_inv_pair2_aux/ qed-. -lemma lpx_inv_pair: ∀h,g,I1,I2,G,L1,L2,V1,V2. ⦃G, L1.ⓑ{I1}V1⦄ ⊢ ➡[h, g] L2.ⓑ{I2}V2 → - ∧∧ ⦃G, L1⦄ ⊢ ➡[h, g] L2 & ⦃G, L1⦄ ⊢ V1 ➡[h, g] V2 & I1 = I2. +lemma lpx_inv_pair: ∀h,o,I1,I2,G,L1,L2,V1,V2. ⦃G, L1.ⓑ{I1}V1⦄ ⊢ ➡[h, o] L2.ⓑ{I2}V2 → + ∧∧ ⦃G, L1⦄ ⊢ ➡[h, o] L2 & ⦃G, L1⦄ ⊢ V1 ➡[h, o] V2 & I1 = I2. /2 width=1 by lpx_sn_inv_pair/ qed-. (* Basic properties *********************************************************) -lemma lpx_refl: ∀h,g,G,L. ⦃G, L⦄ ⊢ ➡[h, g] L. +lemma lpx_refl: ∀h,o,G,L. ⦃G, L⦄ ⊢ ➡[h, o] L. /2 width=1 by lpx_sn_refl/ qed. -lemma lpx_pair: ∀h,g,I,G,K1,K2,V1,V2. ⦃G, K1⦄ ⊢ ➡[h, g] K2 → ⦃G, K1⦄ ⊢ V1 ➡[h, g] V2 → - ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, g] K2.ⓑ{I}V2. +lemma lpx_pair: ∀h,o,I,G,K1,K2,V1,V2. ⦃G, K1⦄ ⊢ ➡[h, o] K2 → ⦃G, K1⦄ ⊢ V1 ➡[h, o] V2 → + ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡[h, o] K2.ⓑ{I}V2. /2 width=1 by lpx_sn_pair/ qed. -lemma lpr_lpx: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡[h, g] L2. -#h #g #G #L1 #L2 #H elim H -L1 -L2 /3 width=1 by lpx_pair, cpr_cpx/ +lemma lpr_lpx: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1⦄ ⊢ ➡[h, o] L2. +#h #o #G #L1 #L2 #H elim H -L1 -L2 /3 width=1 by lpx_pair, cpr_cpx/ qed. (* Basic forward lemmas *****************************************************) -lemma lpx_fwd_length: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → |L1| = |L2|. +lemma lpx_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → |L1| = |L2|. /2 width=2 by lpx_sn_fwd_length/ 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 6c1fbdbfa..6f4b3b2a0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_aaa.ma @@ -21,11 +21,11 @@ include "basic_2/reduction/lpx_drop.ma". (* Properties on atomic arity assignment for terms **************************) (* Note: lemma 500 *) -lemma cpx_lpx_aaa_conf: ∀h,g,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → - ∀T2. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2 → - ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A. -#h #g #G #L1 #T1 #A #H elim H -G -L1 -T1 -A -[ #g #L1 #k #X #H +lemma cpx_lpx_aaa_conf: ∀h,o,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → + ∀T2. ⦃G, L1⦄ ⊢ T1 ➡[h, o] T2 → + ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A. +#h #o #G #L1 #T1 #A #H elim H -G -L1 -T1 -A +[ #o #L1 #s #X #H elim (cpx_inv_sort1 … H) -H // * // | #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #X #H #L2 #HL12 elim (cpx_inv_lref1 … H) -H @@ -70,10 +70,10 @@ lemma cpx_lpx_aaa_conf: ∀h,g,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ] qed-. -lemma cpx_aaa_conf: ∀h,g,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. +lemma cpx_aaa_conf: ∀h,o,G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, o] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. /2 width=7 by cpx_lpx_aaa_conf/ qed-. -lemma lpx_aaa_conf: ∀h,g,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → ⦃G, L2⦄ ⊢ T ⁝ A. +lemma lpx_aaa_conf: ∀h,o,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A → ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ⦃G, L2⦄ ⊢ T ⁝ A. /2 width=7 by cpx_lpx_aaa_conf/ qed-. lemma cpr_aaa_conf: ∀G,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_drop.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_drop.ma index 1944e9dd2..6473e73c0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_drop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_drop.ma @@ -20,58 +20,58 @@ include "basic_2/reduction/lpx.ma". (* Properties on local environment slicing ***********************************) -lemma lpx_drop_conf: ∀h,g,G. dropable_sn (lpx h g G). +lemma lpx_drop_conf: ∀h,o,G. dropable_sn (lpx h o G). /3 width=6 by lpx_sn_deliftable_dropable, cpx_inv_lift1/ qed-. -lemma drop_lpx_trans: ∀h,g,G. dedropable_sn (lpx h g G). +lemma drop_lpx_trans: ∀h,o,G. dedropable_sn (lpx h o G). /3 width=10 by lpx_sn_liftable_dedropable, cpx_lift/ qed-. -lemma lpx_drop_trans_O1: ∀h,g,G. dropable_dx (lpx h g G). +lemma lpx_drop_trans_O1: ∀h,o,G. dropable_dx (lpx h o G). /2 width=3 by lpx_sn_dropable/ qed-. (* Properties on supclosure *************************************************) -lemma fqu_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 → - ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 +lemma fqu_lpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → + ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, o] K2 → + ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, o] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpx_pair, ex3_2_intro/ [ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H #K2 #W2 #HLK2 #HVW2 #H destruct /3 width=5 by cpx_pair_sn, fqu_bind_dx, ex3_2_intro/ -| #G #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #K2 #HK12 +| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #K2 #HK12 elim (drop_lpx_trans … HLK1 … HK12) -HK12 /3 width=7 by fqu_drop, ex3_2_intro/ ] qed-. -lemma fquq_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 → - ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H +lemma fquq_lpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → + ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, o] K2 → + ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, o] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H [ #HT12 elim (fqu_lpx_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/ | * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/ ] qed-. -lemma lpx_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 → - ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 +lemma lpx_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → + ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → + ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /3 width=7 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpx_pair, ex3_2_intro/ [ #I #G1 #L1 #V1 #X #H elim (lpx_inv_pair2 … H) -H #K1 #W1 #HKL1 #HWV1 #H destruct elim (lift_total V1 0 1) /4 width=7 by cpx_delta, fqu_drop, drop_drop, ex3_2_intro/ -| #G #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #L0 #HL01 +| #G #L1 #K1 #T1 #U1 #k #HLK1 #HTU1 #L0 #HL01 elim (lpx_drop_trans_O1 … HL01 … HLK1) -L1 /3 width=5 by fqu_drop, ex3_2_intro/ ] qed-. -lemma lpx_fquq_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 → - ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H +lemma lpx_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → + ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → + ∃∃K2,T. ⦃G1, K1⦄ ⊢ T1 ➡[h, o] T & ⦃G1, K1, T⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H [ #HT12 elim (lpx_fqu_trans … HT12 … HKL1) /3 width=5 by fqu_fquq, ex3_2_intro/ | * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/ ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_frees.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_frees.ma index 960e41f3d..fa314630d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_frees.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_frees.ma @@ -20,12 +20,12 @@ include "basic_2/reduction/lpx_drop.ma". (* Properties on context-sensitive free variables ***************************) -lemma lpx_cpx_frees_trans: ∀h,g,G,L1,U1,U2. ⦃G, L1⦄ ⊢ U1 ➡[h, g] U2 → - ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → +lemma lpx_cpx_frees_trans: ∀h,o,G,L1,U1,U2. ⦃G, L1⦄ ⊢ U1 ➡[h, o] U2 → + ∀L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ∀i. L2 ⊢ i ϵ 𝐅*[0]⦃U2⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U1⦄. -#h #g #G #L1 #U1 @(fqup_wf_ind_eq … G L1 U1) -G -L1 -U1 +#h #o #G #L1 #U1 @(fqup_wf_ind_eq … G L1 U1) -G -L1 -U1 #G0 #L0 #U0 #IH #G #L1 * * -[ -IH #k #HG #HL #HU #U2 #H1 #L2 #_ #i #H2 elim (cpx_inv_sort1 … H1) -H1 +[ -IH #s #HG #HL #HU #U2 #H1 #L2 #_ #i #H2 elim (cpx_inv_sort1 … H1) -H1 [| * #d #_ ] #H destruct elim (frees_inv_sort … H2) | #j #HG #HL #HU #U2 #H1 #L2 #HL12 #i #H2 elim (cpx_inv_lref1 … H1) -H1 [ #H destruct elim (frees_inv_lref … H2) -H2 // @@ -80,9 +80,9 @@ lemma lpx_cpx_frees_trans: ∀h,g,G,L1,U1,U2. ⦃G, L1⦄ ⊢ U1 ➡[h, g] U2 ] qed-. -lemma cpx_frees_trans: ∀h,g,G. frees_trans (cpx h g G). +lemma cpx_frees_trans: ∀h,o,G. frees_trans (cpx h o G). /2 width=8 by lpx_cpx_frees_trans/ qed-. -lemma lpx_frees_trans: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → +lemma lpx_frees_trans: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡[h, o] L2 → ∀U,i. L2 ⊢ i ϵ 𝐅*[0]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[0]⦃U⦄. /2 width=8 by lpx_cpx_frees_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma index 9e731c5a9..8f2f7497b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma @@ -26,10 +26,10 @@ include "basic_2/reduction/lpx_frees.ma". (* Properties on lazy equivalence for local environments ********************) (* Note: contains a proof of llpx_cpx_conf *) -lemma lleq_lpx_trans: ∀h,g,G,L2,K2. ⦃G, L2⦄ ⊢ ➡[h, g] K2 → +lemma lleq_lpx_trans: ∀h,o,G,L2,K2. ⦃G, L2⦄ ⊢ ➡[h, o] K2 → ∀L1,T,l. L1 ≡[T, l] L2 → - ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, g] K1 & K1 ≡[T, l] K2. -#h #g #G #L2 #K2 #HLK2 #L1 #T #l #HL12 + ∃∃K1. ⦃G, L1⦄ ⊢ ➡[h, o] K1 & K1 ≡[T, l] K2. +#h #o #G #L2 #K2 #HLK2 #L1 #T #l #HL12 lapply (lpx_fwd_length … HLK2) #H1 lapply (lleq_fwd_length … HL12) #H2 lapply (lpx_sn_llpx_sn … T … l HLK2) // -HLK2 #H @@ -42,10 +42,10 @@ lapply (llpx_sn_llor_dx_sym … H … HLK1) ] qed-. -lemma lpx_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 → K1 ≡[T1, 0] L1 → - ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2 & K2 ≡[T2, 0] L2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 +lemma lpx_lleq_fqu_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → + ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 → + ∃∃K2. ⦃G1, K1, T1⦄ ⊐ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 [ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1 #K0 #V0 #H1KL1 #_ #H destruct elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 // @@ -59,22 +59,22 @@ lemma lpx_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, /3 width=4 by lpx_pair, fqu_bind_dx, ex3_intro/ | #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H /2 width=4 by fqu_flat_dx, ex3_intro/ -| #G1 #L1 #L #T1 #U1 #m #HL1 #HTU1 #K1 #H1KL1 #H2KL1 - elim (drop_O1_le (Ⓕ) (m+1) K1) +| #G1 #L1 #L #T1 #U1 #k #HL1 #HTU1 #K1 #H1KL1 #H2KL1 + elim (drop_O1_le (Ⓕ) (k+1) K1) [ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 // #H2KL elim (lpx_drop_trans_O1 … H1KL1 … HL1) -L1 #K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct /3 width=4 by fqu_drop, ex3_intro/ - | lapply (drop_fwd_length_le2 … HL1) -L -T1 -g + | lapply (drop_fwd_length_le2 … HL1) -L -T1 -o lapply (lleq_fwd_length … H2KL1) // ] ] qed-. -lemma lpx_lleq_fquq_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 → K1 ≡[T1, 0] L1 → - ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2 & K2 ≡[T2, 0] L2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1 +lemma lpx_lleq_fquq_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → + ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 → + ∃∃K2. ⦃G1, K1, T1⦄ ⊐⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1 elim (fquq_inv_gen … H) -H [ #H elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1 /3 width=4 by fqu_fquq, ex3_intro/ @@ -82,10 +82,10 @@ elim (fquq_inv_gen … H) -H ] qed-. -lemma lpx_lleq_fqup_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → - ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 → K1 ≡[T1, 0] L1 → - ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2 & K2 ≡[T2, 0] L2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2 +lemma lpx_lleq_fqup_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → + ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 → + ∃∃K2. ⦃G1, K1, T1⦄ ⊐+ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2 [ #G2 #L2 #T2 #H #K1 #H1KL1 #H2KL1 elim (lpx_lleq_fqu_trans … H … H1KL1 H2KL1) -L1 /3 width=4 by fqu_fqup, ex3_intro/ | #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #K1 #H1KL1 #H2KL1 elim (IHT1 … H2KL1) // -L1 @@ -94,10 +94,10 @@ lemma lpx_lleq_fqup_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2 ] qed-. -lemma lpx_lleq_fqus_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → - ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 → K1 ≡[T1, 0] L1 → - ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2 & K2 ≡[T2, 0] L2. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1 +lemma lpx_lleq_fqus_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → + ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, o] L1 → K1 ≡[T1, 0] L1 → + ∃∃K2. ⦃G1, K1, T1⦄ ⊐* ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, o] L2 & K2 ≡[T2, 0] L2. +#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #H1KL1 #H2KL1 elim (fqus_inv_gen … H) -H [ #H elim (lpx_lleq_fqup_trans … H … H1KL1 H2KL1) -L1 /3 width=4 by fqup_fqus, ex3_intro/ @@ -105,22 +105,22 @@ elim (fqus_inv_gen … H) -H ] qed-. -fact lreq_lpx_trans_lleq_aux: ∀h,g,G,L1,L0,l,m. L1 ⩬[l, m] L0 → m = ∞ → - ∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 → - ∃∃L. L ⩬[l, m] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L & +fact lreq_lpx_trans_lleq_aux: ∀h,o,G,L1,L0,l,k. L1 ⩬[l, k] L0 → k = ∞ → + ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 → + ∃∃L. L ⩬[l, k] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L & (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L). -#h #g #G #L1 #L0 #l #m #H elim H -L1 -L0 -l -m -[ #l #m #_ #L2 #H >(lpx_inv_atom1 … H) -H +#h #o #G #L1 #L0 #l #k #H elim H -L1 -L0 -l -k +[ #l #k #_ #L2 #H >(lpx_inv_atom1 … H) -H /3 width=5 by ex3_intro, conj/ | #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct -| #I #L1 #L0 #V1 #m #HL10 #IHL10 #Hm #Y #H +| #I #L1 #L0 #V1 #k #HL10 #IHL10 #Hm #Y #H elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpx_pair, lreq_cpx_trans, lreq_pair/ #T elim (IH T) #HL0dx #HL0sn @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/ -| #I1 #I0 #L1 #L0 #V1 #V0 #l #m #HL10 #IHL10 #Hm #Y #H +| #I1 #I0 #L1 #L0 #V1 #V0 #l #k #HL10 #IHL10 #Hm #Y #H elim (lpx_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpx_pair, lreq_succ/ @@ -129,8 +129,8 @@ fact lreq_lpx_trans_lleq_aux: ∀h,g,G,L1,L0,l,m. L1 ⩬[l, m] L0 → m = ∞ ] qed-. -lemma lreq_lpx_trans_lleq: ∀h,g,G,L1,L0,l. L1 ⩬[l, ∞] L0 → - ∀L2. ⦃G, L0⦄ ⊢ ➡[h, g] L2 → - ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, g] L & +lemma lreq_lpx_trans_lleq: ∀h,o,G,L1,L0,l. L1 ⩬[l, ∞] L0 → + ∀L2. ⦃G, L0⦄ ⊢ ➡[h, o] L2 → + ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡[h, o] L & (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L). /2 width=1 by lreq_lpx_trans_lleq_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma index 1fc24f51e..2e21e55ef 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma @@ -13,6 +13,7 @@ (**************************************************************************) include "ground_2/relocation/rtmap_isfin.ma". +include "basic_2/notation/relations/rdropstar_3.ma". include "basic_2/notation/relations/rdropstar_4.ma". include "basic_2/relocation/lreq.ma". include "basic_2/relocation/lifts.ma". @@ -31,6 +32,9 @@ inductive drops (c:bool): rtmap → relation lenv ≝ drops c (↑f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2) . +interpretation "uniform slicing (local environment)" + 'RDropStar i L1 L2 = (drops true (uni i) L1 L2). + interpretation "general slicing (local environment)" 'RDropStar c f L1 L2 = (drops c f L1 L2). @@ -64,69 +68,6 @@ definition dedropable_sn: predicate (rtmap → relation lenv) ≝ ∀f2. f ⊚ f1 ≡ f2 → ∃∃L2. R f2 L1 L2 & ⬇*[c, f] L2 ≡ K2 & L1 ≡[f] L2. -(* Basic inversion lemmas ***************************************************) - -fact drops_inv_atom1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → X = ⋆ → - Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄). -#X #Y #c #f * -X -Y -f -[ /3 width=1 by conj/ -| #I #L1 #L2 #V #f #_ #H destruct -| #I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct -] -qed-. - -(* Basic_1: includes: drop_gen_sort *) -(* Basic_2A1: includes: drop_inv_atom1 *) -lemma drops_inv_atom1: ∀Y,c,f. ⬇*[c, f] ⋆ ≡ Y → Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄). -/2 width=3 by drops_inv_atom1_aux/ qed-. - -fact drops_inv_drop1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K,V,g. X = K.ⓑ{I}V → f = ⫯g → - ⬇*[c, g] K ≡ Y. -#X #Y #c #f * -X -Y -f -[ #f #Ht #J #K #W #g #H destruct -| #I #L1 #L2 #V #f #HL #J #K #W #g #H1 #H2 <(injective_next … H2) -g destruct // -| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K #W #g #_ #H2 elim (discr_push_next … H2) -] -qed-. - -(* Basic_1: includes: drop_gen_drop *) -(* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *) -lemma drops_inv_drop1: ∀I,K,Y,V,c,f. ⬇*[c, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[c, f] K ≡ Y. -/2 width=7 by drops_inv_drop1_aux/ qed-. - - -fact drops_inv_skip1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K1,V1,g. X = K1.ⓑ{I}V1 → f = ↑g → - ∃∃K2,V2. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2. -#X #Y #c #f * -X -Y -f -[ #f #Ht #J #K1 #W1 #g #H destruct -| #I #L1 #L2 #V #f #_ #J #K1 #W1 #g #_ #H2 elim (discr_next_push … H2) -| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct - /2 width=5 by ex3_2_intro/ -] -qed-. - -(* Basic_1: includes: drop_gen_skip_l *) -(* Basic_2A1: includes: drop_inv_skip1 *) -lemma drops_inv_skip1: ∀I,K1,V1,Y,c,f. ⬇*[c, ↑f] K1.ⓑ{I}V1 ≡ Y → - ∃∃K2,V2. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2. -/2 width=5 by drops_inv_skip1_aux/ qed-. - -fact drops_inv_skip2_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K2,V2,g. Y = K2.ⓑ{I}V2 → f = ↑g → - ∃∃K1,V1. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1. -#X #Y #c #f * -X -Y -f -[ #f #Ht #J #K2 #W2 #g #H destruct -| #I #L1 #L2 #V #f #_ #J #K2 #W2 #g #_ #H2 elim (discr_next_push … H2) -| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct - /2 width=5 by ex3_2_intro/ -] -qed-. - -(* Basic_1: includes: drop_gen_skip_r *) -(* Basic_2A1: includes: drop_inv_skip2 *) -lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 → - ∃∃K1,V1. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1. -/2 width=5 by drops_inv_skip2_aux/ qed-. - (* Basic properties *********************************************************) lemma drops_eq_repl_back: ∀L1,L2,c. eq_repl_back … (λf. ⬇*[c, f] L1 ≡ L2). @@ -150,6 +91,27 @@ lemma drops_refl: ∀c,L,f. 𝐈⦃f⦄ → ⬇*[c, f] L ≡ L. /3 width=1 by drops_skip, lifts_refl/ qed. +(* Basic_2A1: includes: drop_split *) +lemma drops_split_trans: ∀L1,L2,f,c. ⬇*[c, f] L1 ≡ L2 → ∀f1,f2. f1 ⊚ f2 ≡ f → 𝐔⦃f1⦄ → + ∃∃L. ⬇*[c, f1] L1 ≡ L & ⬇*[c, f2] L ≡ L2. +#L1 #L2 #f #c #H elim H -L1 -L2 -f +[ #f #Hc #f1 #f2 #Hf #Hf1 @(ex2_intro … (⋆)) @drops_atom + #H lapply (Hc H) -c + #H elim (after_inv_isid3 … Hf H) -f // +| #I #L1 #L2 #V #f #HL12 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxn … Hf) -Hf [1,3: * |*: // ] + [ #g1 #g2 #Hf #H1 #H2 destruct + lapply (isuni_inv_push … Hf1 ??) -Hf1 [1,2: // ] #Hg1 + elim (IHL12 … Hf) -f + /4 width=5 by drops_drop, drops_skip, lifts_refl, isuni_isid, ex2_intro/ + | #g1 #Hf #H destruct elim (IHL12 … Hf) -f + /3 width=5 by ex2_intro, drops_drop, isuni_inv_next/ + ] +| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 #f1 #f2 #Hf #Hf1 elim (after_inv_xxp … Hf) -Hf [2,3: // ] + #g1 #g2 #Hf #H1 #H2 destruct elim (lifts_split_trans … HV21 … Hf) -HV21 + elim (IHL12 … Hf) -f /3 width=5 by ex2_intro, drops_skip, isuni_fwd_push/ +] +qed-. + (* Basic_2A1: includes: drop_FT *) lemma drops_TF: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2. #L1 #L2 #f #H elim H -L1 -L2 -f @@ -204,17 +166,114 @@ qed-. lemma drops_isuni_fwd_drop2: ∀I,X,K,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] X ≡ K.ⓑ{I}V → ⬇*[c, ⫯f] X ≡ K. /3 width=7 by drops_after_fwd_drop2, after_isid_isuni/ qed-. -(* forward lemmas with test for finite colength *****************************) +(* Forward lemmas with test for finite colength *****************************) lemma drops_fwd_isfin: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → 𝐅⦃f⦄. #L1 #L2 #f #H elim H -L1 -L2 -f /3 width=1 by isfin_next, isfin_push, isfin_isid/ qed-. -(* Basic_2A1: removed theorems 14: +(* Basic inversion lemmas ***************************************************) + +fact drops_inv_atom1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → X = ⋆ → + Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄). +#X #Y #c #f * -X -Y -f +[ /3 width=1 by conj/ +| #I #L1 #L2 #V #f #_ #H destruct +| #I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct +] +qed-. + +(* Basic_1: includes: drop_gen_sort *) +(* Basic_2A1: includes: drop_inv_atom1 *) +lemma drops_inv_atom1: ∀Y,c,f. ⬇*[c, f] ⋆ ≡ Y → Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄). +/2 width=3 by drops_inv_atom1_aux/ qed-. + +fact drops_inv_drop1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K,V,g. X = K.ⓑ{I}V → f = ⫯g → + ⬇*[c, g] K ≡ Y. +#X #Y #c #f * -X -Y -f +[ #f #Ht #J #K #W #g #H destruct +| #I #L1 #L2 #V #f #HL #J #K #W #g #H1 #H2 <(injective_next … H2) -g destruct // +| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K #W #g #_ #H2 elim (discr_push_next … H2) +] +qed-. + +(* Basic_1: includes: drop_gen_drop *) +(* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *) +lemma drops_inv_drop1: ∀I,K,Y,V,c,f. ⬇*[c, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[c, f] K ≡ Y. +/2 width=7 by drops_inv_drop1_aux/ qed-. + + +fact drops_inv_skip1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K1,V1,g. X = K1.ⓑ{I}V1 → f = ↑g → + ∃∃K2,V2. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2. +#X #Y #c #f * -X -Y -f +[ #f #Ht #J #K1 #W1 #g #H destruct +| #I #L1 #L2 #V #f #_ #J #K1 #W1 #g #_ #H2 elim (discr_next_push … H2) +| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct + /2 width=5 by ex3_2_intro/ +] +qed-. + +(* Basic_1: includes: drop_gen_skip_l *) +(* Basic_2A1: includes: drop_inv_skip1 *) +lemma drops_inv_skip1: ∀I,K1,V1,Y,c,f. ⬇*[c, ↑f] K1.ⓑ{I}V1 ≡ Y → + ∃∃K2,V2. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2. +/2 width=5 by drops_inv_skip1_aux/ qed-. + +fact drops_inv_skip2_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K2,V2,g. Y = K2.ⓑ{I}V2 → f = ↑g → + ∃∃K1,V1. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1. +#X #Y #c #f * -X -Y -f +[ #f #Ht #J #K2 #W2 #g #H destruct +| #I #L1 #L2 #V #f #_ #J #K2 #W2 #g #_ #H2 elim (discr_next_push … H2) +| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct + /2 width=5 by ex3_2_intro/ +] +qed-. + +(* Basic_1: includes: drop_gen_skip_r *) +(* Basic_2A1: includes: drop_inv_skip2 *) +lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 → + ∃∃K1,V1. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1. +/2 width=5 by drops_inv_skip2_aux/ qed-. + +(* Inversion lemmas with test for uniformity ********************************) + +(* Basic_2A1: was: drop_inv_O1_pair1 *) +lemma drops_inv_pair1_isuni: ∀I,K,L2,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] K.ⓑ{I}V ≡ L2 → + (𝐈⦃f⦄ ∧ L2 = K.ⓑ{I}V) ∨ + ∃∃g. ⬇*[c, g] K ≡ L2 & f = ⫯g. +#I #K #L2 #V #c #f #Hf #H elim (isuni_split … Hf) -Hf * #g #Hg #H0 destruct +[ lapply (drops_inv_skip1 … H) -H * #Y #X #HY #HX #H destruct + <(drops_fwd_isid … HY Hg) -Y >(lifts_fwd_isid … HX Hg) -X + /4 width=3 by isid_push, or_introl, conj/ +| lapply (drops_inv_drop1 … H) -H /3 width=3 by ex2_intro, or_intror/ +] +qed-. + +(* Basic_2A1: was: drop_inv_O1_pair2 *) +lemma drops_inv_pair2_isuni: ∀I,K,V,c,f,L1. 𝐔⦃f⦄ → ⬇*[c, f] L1 ≡ K.ⓑ{I}V → + (𝐈⦃f⦄ ∧ L1 = K.ⓑ{I}V) ∨ + ∃∃I1,K1,V1,g. ⬇*[c, g] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1 & f = ⫯g. +#I #K #V #c #f * +[ #Hf #H elim (drops_inv_atom1 … H) -H #H destruct +| #L1 #I1 #V1 #Hf #H elim (drops_inv_pair1_isuni … Hf H) -Hf -H * + [ #Hf #H destruct /3 width=1 by or_introl, conj/ + | /3 width=7 by ex3_4_intro, or_intror/ + ] +] +qed-. + +lemma drops_inv_pair2_isuni_next: ∀I,K,V,c,f,L1. 𝐔⦃f⦄ → ⬇*[c, ⫯f] L1 ≡ K.ⓑ{I}V → + ∃∃I1,K1,V1. ⬇*[c, f] K1 ≡ K.ⓑ{I}V & L1 = K1.ⓑ{I1}V1. +#I #K #V #c #f #L1 #Hf #H elim (drops_inv_pair2_isuni … H) -H /2 width=3 by isuni_next/ -Hf * +[ #H elim (isid_inv_next … H) -H // +| /2 width=5 by ex2_3_intro/ +] +qed-. + +(* Basic_2A1: removed theorems 12: drops_inv_nil drops_inv_cons d1_liftable_liftables - drop_refl_atom_O2 - drop_inv_O1_pair1 drop_inv_pair1 drop_inv_O1_pair2 + drop_refl_atom_O2 drop_inv_pair1 drop_inv_Y1 drop_Y1 drop_O_Y drop_fwd_Y2 drop_fwd_length_minus2 drop_fwd_length_minus4 *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma index cae6be8f9..a72f58604 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma @@ -34,8 +34,12 @@ inductive lifts: rtmap → relation term ≝ lifts f (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) . +interpretation "uniform relocation (term)" + 'RLiftStar i T1 T2 = (lifts (uni i) T1 T2). + interpretation "generic relocation (term)" - 'RLiftStar cs T1 T2 = (lifts cs T1 T2). + 'RLiftStar f T1 T2 = (lifts f T1 T2). + (* Basic inversion lemmas ***************************************************) @@ -101,7 +105,7 @@ lemma lifts_inv_bind1: ∀p,I,V1,T1,Y,f. ⬆*[f] ⓑ{p,I}V1.T1 ≡ Y → Y = ⓑ{p,I}V2.T2. /2 width=3 by lifts_inv_bind1_aux/ qed-. -fact lifts_inv_flat1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → +fact lifts_inv_flat1_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y → ∀I,V1,T1. X = ⓕ{I}V1.T1 → ∃∃V2,T2. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 & Y = ⓕ{I}V2.T2. @@ -116,7 +120,7 @@ qed-. (* Basic_1: was: lift1_flat *) (* Basic_2A1: includes: lift_inv_flat1 *) -lemma lifts_inv_flat1: ∀I,V1,T1,Y,f. ⬆*[f] ⓕ{I}V1.T1 ≡ Y → +lemma lifts_inv_flat1: ∀I,V1,T1,Y. ∀f:rtmap. ⬆*[f] ⓕ{I}V1.T1 ≡ Y → ∃∃V2,T2. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 & Y = ⓕ{I}V2.T2. /2 width=3 by lifts_inv_flat1_aux/ qed-. @@ -183,7 +187,7 @@ lemma lifts_inv_bind2: ∀p,I,V2,T2,X,f. ⬆*[f] X ≡ ⓑ{p,I}V2.T2 → X = ⓑ{p,I}V1.T1. /2 width=3 by lifts_inv_bind2_aux/ qed-. -fact lifts_inv_flat2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → +fact lifts_inv_flat2_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y → ∀I,V2,T2. Y = ⓕ{I}V2.T2 → ∃∃V1,T1. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 & X = ⓕ{I}V1.T1. @@ -198,7 +202,7 @@ qed-. (* Basic_1: includes: lift_gen_flat *) (* Basic_2A1: includes: lift_inv_flat2 *) -lemma lifts_inv_flat2: ∀I,V2,T2,X,f. ⬆*[f] X ≡ ⓕ{I}V2.T2 → +lemma lifts_inv_flat2: ∀I,V2,T2,X. ∀f:rtmap. ⬆*[f] X ≡ ⓕ{I}V2.T2 → ∃∃V1,T1. ⬆*[f] V1 ≡ V2 & ⬆*[f] T1 ≡ T2 & X = ⓕ{I}V1.T1. /2 width=3 by lifts_inv_flat2_aux/ qed-. @@ -245,7 +249,7 @@ lemma lifts_fwd_isid: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 → 𝐈⦃f⦄ → T1 = T2. qed-. (* Basic_2A1: includes: lift_fwd_pair1 *) -lemma lifts_fwd_pair1: ∀I,V1,T1,Y,f. ⬆*[f] ②{I}V1.T1 ≡ Y → +lemma lifts_fwd_pair1: ∀I,V1,T1,Y. ∀f:rtmap. ⬆*[f] ②{I}V1.T1 ≡ Y → ∃∃V2,T2. ⬆*[f] V1 ≡ V2 & Y = ②{I}V2.T2. * [ #p ] #I #V1 #T1 #Y #f #H [ elim (lifts_inv_bind1 … H) -H /2 width=4 by ex2_2_intro/ @@ -254,7 +258,7 @@ lemma lifts_fwd_pair1: ∀I,V1,T1,Y,f. ⬆*[f] ②{I}V1.T1 ≡ Y → qed-. (* Basic_2A1: includes: lift_fwd_pair2 *) -lemma lifts_fwd_pair2: ∀I,V2,T2,X,f. ⬆*[f] X ≡ ②{I}V2.T2 → +lemma lifts_fwd_pair2: ∀I,V2,T2,X. ∀f:rtmap. ⬆*[f] X ≡ ②{I}V2.T2 → ∃∃V1,T1. ⬆*[f] V1 ≡ V2 & X = ②{I}V1.T1. * [ #p ] #I #V2 #T2 #X #f #H [ elim (lifts_inv_bind2 … H) -H /2 width=4 by ex2_2_intro/ @@ -291,6 +295,9 @@ elim (IHV1 f) -IHV1 #V2 #HV12 ] qed-. +lemma lift_SO: ∀i. ⬆*[1] #i ≡ #(⫯i). +/2 width=1 by lifts_lref/ qed. + (* Basic_1: includes: lift_free (right to left) *) (* Basic_2A1: includes: lift_split *) lemma lifts_split_trans: ∀T1,T2,f. ⬆*[f] T1 ≡ T2 → diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts_vector.ma index f741b3e5a..a83fa3590 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts_vector.ma @@ -20,34 +20,34 @@ include "basic_2/relocation/lifts_vector.ma". (* Main properties **********************************************************) (* Basic_1: includes: lifts_inj *) -theorem liftsv_inj: ∀T1c,Us,f. ⬆*[f] T1c ≡ Us → - ∀T2c. ⬆*[f] T2c ≡ Us → T1c = T2c. -#T1c #Us #f #H elim H -T1c -Us -[ #T2c #H >(liftsv_inv_nil2 … H) -H // -| #T1c #Us #T1 #U #HT1U #_ #IHT1Us #X #H destruct - elim (liftsv_inv_cons2 … H) -H #T2 #T2c #HT2U #HT2Us #H destruct +theorem liftsv_inj: ∀T1s,Us,f. ⬆*[f] T1s ≡ Us → + ∀T2s. ⬆*[f] T2s ≡ Us → T1s = T2s. +#T1s #Us #f #H elim H -T1s -Us +[ #T2s #H >(liftsv_inv_nil2 … H) -H // +| #T1s #Us #T1 #U #HT1U #_ #IHT1Us #X #H destruct + elim (liftsv_inv_cons2 … H) -H #T2 #T2s #HT2U #HT2Us #H destruct >(lifts_inj … HT1U … HT2U) -U /3 width=1 by eq_f/ ] qed-. (* Basic_2A1: includes: liftv_mono *) -theorem liftsv_mono: ∀Ts,U1c,f. ⬆*[f] Ts ≡ U1c → - ∀U2c. ⬆*[f] Ts ≡ U2c → U1c = U2c. -#Ts #U1c #f #H elim H -Ts -U1c -[ #U2c #H >(liftsv_inv_nil1 … H) -H // -| #Ts #U1c #T #U1 #HTU1 #_ #IHTU1c #X #H destruct - elim (liftsv_inv_cons1 … H) -H #U2 #U2c #HTU2 #HTU2c #H destruct +theorem liftsv_mono: ∀Ts,U1s,f. ⬆*[f] Ts ≡ U1s → + ∀U2s. ⬆*[f] Ts ≡ U2s → U1s = U2s. +#Ts #U1s #f #H elim H -Ts -U1s +[ #U2s #H >(liftsv_inv_nil1 … H) -H // +| #Ts #U1s #T #U1 #HTU1 #_ #IHTU1s #X #H destruct + elim (liftsv_inv_cons1 … H) -H #U2 #U2s #HTU2 #HTU2s #H destruct >(lifts_mono … HTU1 … HTU2) -T /3 width=1 by eq_f/ ] qed-. (* Basic_1: includes: lifts1_xhg (right to left) *) (* Basic_2A1: includes: liftsv_liftv_trans_le *) -theorem liftsv_trans: ∀T1c,Ts,f1. ⬆*[f1] T1c ≡ Ts → ∀T2c,f2. ⬆*[f2] Ts ≡ T2c → - ∀f. f2 ⊚ f1 ≡ f → ⬆*[f] T1c ≡ T2c. -#T1c #Ts #f1 #H elim H -T1c -Ts -[ #T2c #f2 #H >(liftsv_inv_nil1 … H) -T2c /2 width=3 by liftsv_nil/ -| #T1c #Ts #T1 #T #HT1 #_ #IHT1c #X #f2 #H elim (liftsv_inv_cons1 … H) -H - #T2 #T2c #HT2 #HT2c #H destruct /3 width=6 by lifts_trans, liftsv_cons/ +theorem liftsv_trans: ∀T1s,Ts,f1. ⬆*[f1] T1s ≡ Ts → ∀T2s,f2. ⬆*[f2] Ts ≡ T2s → + ∀f. f2 ⊚ f1 ≡ f → ⬆*[f] T1s ≡ T2s. +#T1s #Ts #f1 #H elim H -T1s -Ts +[ #T2s #f2 #H >(liftsv_inv_nil1 … H) -T2s /2 width=3 by liftsv_nil/ +| #T1s #Ts #T1 #T #HT1 #_ #IHT1s #X #f2 #H elim (liftsv_inv_cons1 … H) -H + #T2 #T2s #HT2 #HT2s #H destruct /3 width=6 by lifts_trans, liftsv_cons/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_vector.ma index 9c6128335..e7003d4e4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_vector.ma @@ -18,105 +18,108 @@ include "basic_2/relocation/lifts.ma". (* GENERIC RELOCATION FOR TERM VECTORS *************************************) (* Basic_2A1: includes: liftv_nil liftv_cons *) -inductive liftsv (f): relation (list term) ≝ +inductive liftsv (f:rtmap): relation (list term) ≝ | liftsv_nil : liftsv f (◊) (◊) -| liftsv_cons: ∀T1c,T2c,T1,T2. - ⬆*[f] T1 ≡ T2 → liftsv f T1c T2c → - liftsv f (T1 @ T1c) (T2 @ T2c) +| liftsv_cons: ∀T1s,T2s,T1,T2. + ⬆*[f] T1 ≡ T2 → liftsv f T1s T2s → + liftsv f (T1 @ T1s) (T2 @ T2s) . +interpretation "uniform relocation (vector)" + 'RLiftStar i T1s T2s = (liftsv (uni i) T1s T2s). + interpretation "generic relocation (vector)" - 'RLiftStar f T1c T2c = (liftsv f T1c T2c). + 'RLiftStar f T1s T2s = (liftsv f T1s T2s). (* Basic inversion lemmas ***************************************************) fact liftsv_inv_nil1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → X = ◊ → Y = ◊. #X #Y #f * -X -Y // -#T1c #T2c #T1 #T2 #_ #_ #H destruct +#T1s #T2s #T1 #T2 #_ #_ #H destruct qed-. (* Basic_2A1: includes: liftv_inv_nil1 *) lemma liftsv_inv_nil1: ∀Y,f. ⬆*[f] ◊ ≡ Y → Y = ◊. /2 width=5 by liftsv_inv_nil1_aux/ qed-. -fact liftsv_inv_cons1_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → - ∀T1,T1c. X = T1 @ T1c → - ∃∃T2,T2c. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1c ≡ T2c & - Y = T2 @ T2c. +fact liftsv_inv_cons1_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y → + ∀T1,T1s. X = T1 @ T1s → + ∃∃T2,T2s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s & + Y = T2 @ T2s. #X #Y #f * -X -Y -[ #U1 #U1c #H destruct -| #T1c #T2c #T1 #T2 #HT12 #HT12c #U1 #U1c #H destruct /2 width=5 by ex3_2_intro/ +[ #U1 #U1s #H destruct +| #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5 by ex3_2_intro/ ] qed-. (* Basic_2A1: includes: liftv_inv_cons1 *) -lemma liftsv_inv_cons1: ∀T1,T1c,Y,f. ⬆*[f] T1 @ T1c ≡ Y → - ∃∃T2,T2c. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1c ≡ T2c & - Y = T2 @ T2c. +lemma liftsv_inv_cons1: ∀T1,T1s,Y. ∀f:rtmap. ⬆*[f] T1 @ T1s ≡ Y → + ∃∃T2,T2s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s & + Y = T2 @ T2s. /2 width=3 by liftsv_inv_cons1_aux/ qed-. fact liftsv_inv_nil2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → Y = ◊ → X = ◊. #X #Y #f * -X -Y // -#T1c #T2c #T1 #T2 #_ #_ #H destruct +#T1s #T2s #T1 #T2 #_ #_ #H destruct qed-. lemma liftsv_inv_nil2: ∀X,f. ⬆*[f] X ≡ ◊ → X = ◊. /2 width=5 by liftsv_inv_nil2_aux/ qed-. -fact liftsv_inv_cons2_aux: ∀X,Y,f. ⬆*[f] X ≡ Y → - ∀T2,T2c. Y = T2 @ T2c → - ∃∃T1,T1c. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1c ≡ T2c & - X = T1 @ T1c. +fact liftsv_inv_cons2_aux: ∀X,Y. ∀f:rtmap. ⬆*[f] X ≡ Y → + ∀T2,T2s. Y = T2 @ T2s → + ∃∃T1,T1s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s & + X = T1 @ T1s. #X #Y #f * -X -Y -[ #U2 #U2c #H destruct -| #T1c #T2c #T1 #T2 #HT12 #HT12c #U2 #U2c #H destruct /2 width=5 by ex3_2_intro/ +[ #U2 #U2s #H destruct +| #T1s #T2s #T1 #T2 #HT12 #HT12s #U2 #U2s #H destruct /2 width=5 by ex3_2_intro/ ] qed-. -lemma liftsv_inv_cons2: ∀X,T2,T2c,f. ⬆*[f] X ≡ T2 @ T2c → - ∃∃T1,T1c. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1c ≡ T2c & - X = T1 @ T1c. +lemma liftsv_inv_cons2: ∀X,T2,T2s. ∀f:rtmap. ⬆*[f] X ≡ T2 @ T2s → + ∃∃T1,T1s. ⬆*[f] T1 ≡ T2 & ⬆*[f] T1s ≡ T2s & + X = T1 @ T1s. /2 width=3 by liftsv_inv_cons2_aux/ qed-. (* Basic_1: was: lifts1_flat (left to right) *) -lemma lifts_inv_applv1: ∀V1c,U1,T2,f. ⬆*[f] Ⓐ V1c.U1 ≡ T2 → - ∃∃V2c,U2. ⬆*[f] V1c ≡ V2c & ⬆*[f] U1 ≡ U2 & - T2 = Ⓐ V2c.U2. -#V1c elim V1c -V1c +lemma lifts_inv_applv1: ∀V1s,U1,T2. ∀f:rtmap. ⬆*[f] Ⓐ V1s.U1 ≡ T2 → + ∃∃V2s,U2. ⬆*[f] V1s ≡ V2s & ⬆*[f] U1 ≡ U2 & + T2 = Ⓐ V2s.U2. +#V1s elim V1s -V1s [ /3 width=5 by ex3_2_intro, liftsv_nil/ -| #V1 #V1c #IHV1c #T1 #X #f #H elim (lifts_inv_flat1 … H) -H - #V2 #Y #HV12 #HY #H destruct elim (IHV1c … HY) -IHV1c -HY - #V2c #T2 #HV12c #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/ +| #V1 #V1s #IHV1s #T1 #X #f #H elim (lifts_inv_flat1 … H) -H + #V2 #Y #HV12 #HY #H destruct elim (IHV1s … HY) -IHV1s -HY + #V2c #T2 #HV12s #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/ ] qed-. -lemma lifts_inv_applv2: ∀V2c,U2,T1,f. ⬆*[f] T1 ≡ Ⓐ V2c.U2 → - ∃∃V1c,U1. ⬆*[f] V1c ≡ V2c & ⬆*[f] U1 ≡ U2 & - T1 = Ⓐ V1c.U1. -#V2c elim V2c -V2c +lemma lifts_inv_applv2: ∀V2s,U2,T1. ∀f:rtmap. ⬆*[f] T1 ≡ Ⓐ V2s.U2 → + ∃∃V1s,U1. ⬆*[f] V1s ≡ V2s & ⬆*[f] U1 ≡ U2 & + T1 = Ⓐ V1s.U1. +#V2s elim V2s -V2s [ /3 width=5 by ex3_2_intro, liftsv_nil/ -| #V2 #V2c #IHV2c #T2 #X #f #H elim (lifts_inv_flat2 … H) -H - #V1 #Y #HV12 #HY #H destruct elim (IHV2c … HY) -IHV2c -HY - #V1c #T1 #HV12c #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/ +| #V2 #V2s #IHV2s #T2 #X #f #H elim (lifts_inv_flat2 … H) -H + #V1 #Y #HV12 #HY #H destruct elim (IHV2s … HY) -IHV2s -HY + #V1s #T1 #HV12s #HT12 #H destruct /3 width=5 by ex3_2_intro, liftsv_cons/ ] qed-. (* Basic properties *********************************************************) (* Basic_2A1: includes: liftv_total *) -lemma liftsv_total: ∀f. ∀T1c:list term. ∃T2c. ⬆*[f] T1c ≡ T2c. -#f #T1c elim T1c -T1c +lemma liftsv_total: ∀f. ∀T1s:list term. ∃T2s. ⬆*[f] T1s ≡ T2s. +#f #T1s elim T1s -T1s [ /2 width=2 by liftsv_nil, ex_intro/ -| #T1 #T1c * #T2c #HT12c +| #T1 #T1s * #T2s #HT12s elim (lifts_total T1 f) /3 width=2 by liftsv_cons, ex_intro/ ] qed-. (* Basic_1: was: lifts1_flat (right to left) *) -lemma lifts_applv: ∀V1c,V2c,f. ⬆*[f] V1c ≡ V2c → +lemma lifts_applv: ∀V1s,V2s. ∀f:rtmap. ⬆*[f] V1s ≡ V2s → ∀T1,T2. ⬆*[f] T1 ≡ T2 → - ⬆*[f] Ⓐ V1c. T1 ≡ Ⓐ V2c. T2. -#V1c #V2c #f #H elim H -V1c -V2c /3 width=1 by lifts_flat/ + ⬆*[f] Ⓐ V1s.T1 ≡ Ⓐ V2s.T2. +#V1s #V2s #f #H elim H -V1s -V2s /3 width=1 by lifts_flat/ qed. (* Basic_1: removed theorems 2: lifts1_nil lifts1_cons *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu.ma b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu.ma new file mode 100644 index 000000000..cfc15cab9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu.ma @@ -0,0 +1,42 @@ +(**************************************************************************) +(* ___ *) +(* ||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/supterm_6.ma". +include "basic_2/grammar/lenv.ma". +include "basic_2/grammar/genv.ma". +include "basic_2/relocation/lifts.ma". + +(* SUPCLOSURE ***************************************************************) + +(* activate genv *) +inductive fqu: tri_relation genv lenv term ≝ +| fqu_lref_O : ∀I,G,L,V. fqu G (L.ⓑ{I}V) (#0) G L V +| fqu_pair_sn: ∀I,G,L,V,T. fqu G L (②{I}V.T) G L V +| fqu_bind_dx: ∀p,I,G,L,V,T. fqu G L (ⓑ{p,I}V.T) G (L.ⓑ{I}V) T +| fqu_flat_dx: ∀I,G,L,V,T. fqu G L (ⓕ{I}V.T) G L T +| fqu_drop : ∀I,G,L,V,T,U. ⬆*[1] T ≡ U → fqu G (L.ⓑ{I}V) U G L T +. + +interpretation + "structural successor (closure)" + 'SupTerm G1 L1 T1 G2 L2 T2 = (fqu G1 L1 T1 G2 L2 T2). + +(* Basic properties *********************************************************) + +lemma fqu_lref_S: ∀I,G,L,V,i. ⦃G, L.ⓑ{I}V, #(⫯i)⦄ ⊐ ⦃G, L, #(i)⦄. +/2 width=1 by fqu_drop/ qed. + +(* Basic_2A1: removed theorems 3: + fqu_drop fqu_drop_lt fqu_lref_S_lt +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_length.ma b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_length.ma new file mode 100644 index 000000000..9719388a5 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_length.ma @@ -0,0 +1,46 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/grammar/lenv_length.ma". +include "basic_2/s_transition/fqu.ma". + +(* SUPCLOSURE ***************************************************************) + +(* Forward lemmas with length for local environments ************************) + +fact fqu_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → + ∀i. T1 = #i → |L2| < |L1|. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // [2: #p] +#I #G #L #V #T #j #H destruct +qed-. + +lemma fqu_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐ ⦃G2, L2, T2⦄ → + |L2| < |L1|. +/2 width=7 by fqu_fwd_length_lref1_aux/ +qed-. + +(* Inversion lemmas with length for local environments **********************) + +fact fqu_inv_eq_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → + G1 = G2 → |L1| = |L2| → T1 = T2 → ⊥. +#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2 +[1: #I #G #L #V #_ #H elim (succ_inv_refl_sn … H) +|5: #I #G #L #V #T #U #_ #_ #H elim (succ_inv_refl_sn … H) +] +/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x/ +qed-. + +lemma fqu_inv_eq: ∀G,L1,L2,T. ⦃G, L1, T⦄ ⊐ ⦃G, L2, T⦄ → |L1| = |L2| → ⊥. +#G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *) +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_weight.ma b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_weight.ma new file mode 100644 index 000000000..6a564e0c9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_weight.ma @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/grammar/cl_weight.ma". +include "basic_2/relocation/lifts_weight.ma". +include "basic_2/s_transition/fqu.ma". + +(* SUPCLOSURE ***************************************************************) + +(* Forward lemmas with weight for closures **********************************) + +lemma fqu_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 // +#I #G #L #V #T #U #HTU normalize in ⊢ (?%%); -I +<(lifts_fwd_tw … HTU) -U /3 width=1 by monotonic_lt_plus_r, monotonic_lt_plus_l/ +qed-. + +(* Advanced eliminators *****************************************************) + +lemma fqu_wf_ind: ∀R:relation3 …. ( + ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) → + R G1 L1 T1 + ) → ∀G1,L1,T1. R G1 L1 T1. +#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqu_fwd_fw/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq.ma b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq.ma new file mode 100644 index 000000000..5efd7b490 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq.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/notation/relations/suptermopt_6.ma". +include "basic_2/s_transition/fqu.ma". + +(* OPTIONAL SUPCLOSURE ******************************************************) + +(* Basic_2A1: was: fquqa *) +(* Basic_2A1: includes: fquq_inv_gen *) +definition fquq: tri_relation genv lenv term ≝ tri_RC … fqu. + +interpretation + "optional structural successor (closure)" + 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fquq G1 L1 T1 G2 L2 T2). + +(* Basic properties *********************************************************) + +(* Basic_2A1: includes: fquqa_refl *) +lemma fquq_refl: tri_reflexive … fquq. +// qed. + +lemma fqu_fquq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄. +/2 width=1 by or_introl/ qed. + +(* Basic_2A1: removed theorems 8: + fquq_lref_O fquq_pair_sn fquq_bind_dx fquq_flat_dx fquq_drop + fquqa_drop fquq_fquqa fquqa_inv_fquq +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq_length.ma b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq_length.ma new file mode 100644 index 000000000..9e9027694 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq_length.ma @@ -0,0 +1,25 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/s_transition/fqu_length.ma". +include "basic_2/s_transition/fquq.ma". + +(* OPTIONAL SUPCLOSURE ******************************************************) + +(* Forward lemmas with length for local environments ************************) + +lemma fquq_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐⸮ ⦃G2, L2, T2⦄ → |L2| ≤ |L1|. +#G1 #G2 #L1 #L2 #T2 #i #H elim H -H [2: * ] +/3 width=5 by fqu_fwd_length_lref1, lt_to_le/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq_weight.ma b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq_weight.ma new file mode 100644 index 000000000..29b474f22 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/s_transition/fquq_weight.ma @@ -0,0 +1,25 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/s_transition/fqu_weight.ma". +include "basic_2/s_transition/fquq.ma". + +(* OPTIONAL SUPCLOSURE ******************************************************) + +(* Forward lemmas with weight for closures **********************************) + +lemma fquq_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 -H [2: * ] +/3 width=1 by fqu_fwd_fw, lt_to_le/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma index 9c1f53736..a57a60b24 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma @@ -14,19 +14,20 @@ include "basic_2/notation/relations/atomicarity_4.ma". include "basic_2/grammar/aarity.ma". +include "basic_2/grammar/lenv.ma". include "basic_2/grammar/genv.ma". -include "basic_2/substitution/drop.ma". -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) +(* ATONIC ARITY ASSIGNMENT FOR TERMS ****************************************) (* activate genv *) inductive aaa: relation4 genv lenv term aarity ≝ -| aaa_sort: ∀G,L,k. aaa G L (⋆k) (⓪) -| aaa_lref: ∀I,G,L,K,V,B,i. ⬇[i] L ≡ K. ⓑ{I}V → aaa G K V B → aaa G L (#i) B -| aaa_abbr: ∀a,G,L,V,T,B,A. - aaa G L V B → aaa G (L.ⓓV) T A → aaa G L (ⓓ{a}V.T) A -| aaa_abst: ∀a,G,L,V,T,B,A. - aaa G L V B → aaa G (L.ⓛV) T A → aaa G L (ⓛ{a}V.T) (②B.A) +| aaa_sort: ∀G,L,s. aaa G L (⋆s) (⓪) +| aaa_zero: ∀I,G,L,V,B. aaa G L V B → aaa G (L.ⓑ{I}V) (#0) B +| aaa_lref: ∀I,G,L,V,A,i. aaa G L (#i) A → aaa G (L.ⓑ{I}V) (#⫯i) A +| aaa_abbr: ∀p,G,L,V,T,B,A. + aaa G L V B → aaa G (L.ⓓV) T A → aaa G L (ⓓ{p}V.T) A +| aaa_abst: ∀p,G,L,V,T,B,A. + aaa G L V B → aaa G (L.ⓛV) T A → aaa G L (ⓛ{p}V.T) (②B.A) | aaa_appl: ∀G,L,V,T,B,A. aaa G L V B → aaa G L T (②B.A) → aaa G L (ⓐV.T) A | aaa_cast: ∀G,L,V,T,A. aaa G L V A → aaa G L T A → aaa G L (ⓝV.T) A . @@ -36,89 +37,110 @@ interpretation "atomic arity assignment (term)" (* Basic inversion lemmas ***************************************************) -fact aaa_inv_sort_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀k. T = ⋆k → A = ⓪. -#G #L #T #A * -G -L -T -A -[ // -| #I #G #L #K #V #B #i #_ #_ #k #H destruct -| #a #G #L #V #T #B #A #_ #_ #k #H destruct -| #a #G #L #V #T #B #A #_ #_ #k #H destruct -| #G #L #V #T #B #A #_ #_ #k #H destruct -| #G #L #V #T #A #_ #_ #k #H destruct +fact aaa_inv_sort_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀s. T = ⋆s → A = ⓪. +#G #L #T #A * -G -L -T -A // +[ #I #G #L #V #B #_ #s #H destruct +| #I #G #L #V #A #i #_ #s #H destruct +| #p #G #L #V #T #B #A #_ #_ #s #H destruct +| #p #G #L #V #T #B #A #_ #_ #s #H destruct +| #G #L #V #T #B #A #_ #_ #s #H destruct +| #G #L #V #T #A #_ #_ #s #H destruct ] qed-. -lemma aaa_inv_sort: ∀G,L,A,k. ⦃G, L⦄ ⊢ ⋆k ⁝ A → A = ⓪. +lemma aaa_inv_sort: ∀G,L,A,s. ⦃G, L⦄ ⊢ ⋆s ⁝ A → A = ⓪. /2 width=6 by aaa_inv_sort_aux/ qed-. -fact aaa_inv_lref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀i. T = #i → - ∃∃I,K,V. ⬇[i] L ≡ K.ⓑ{I} V & ⦃G, K⦄ ⊢ V ⁝ A. +fact aaa_inv_zero_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → T = #0 → + ∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A. +#G #L #T #A * -G -L -T -A /2 width=5 by ex2_3_intro/ +[ #G #L #s #H destruct +| #I #G #L #V #A #i #_ #H destruct +| #p #G #L #V #T #B #A #_ #_ #H destruct +| #p #G #L #V #T #B #A #_ #_ #H destruct +| #G #L #V #T #B #A #_ #_ #H destruct +| #G #L #V #T #A #_ #_ #H destruct +] +qed-. + +lemma aaa_inv_zero: ∀G,L,A. ⦃G, L⦄ ⊢ #0 ⁝ A → + ∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A. +/2 width=3 by aaa_inv_zero_aux/ qed-. + +fact aaa_inv_lref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀i. T = #(⫯i) → + ∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ #i ⁝ A. #G #L #T #A * -G -L -T -A -[ #G #L #k #i #H destruct -| #I #G #L #K #V #B #j #HLK #HB #i #H destruct /2 width=5 by ex2_3_intro/ -| #a #G #L #V #T #B #A #_ #_ #i #H destruct -| #a #G #L #V #T #B #A #_ #_ #i #H destruct -| #G #L #V #T #B #A #_ #_ #i #H destruct -| #G #L #V #T #A #_ #_ #i #H destruct +[ #G #L #s #j #H destruct +| #I #G #L #V #B #_ #j #H destruct +| #I #G #L #V #A #i #HA #j #H destruct /2 width=5 by ex2_3_intro/ +| #p #G #L #V #T #B #A #_ #_ #j #H destruct +| #p #G #L #V #T #B #A #_ #_ #j #H destruct +| #G #L #V #T #B #A #_ #_ #j #H destruct +| #G #L #V #T #A #_ #_ #j #H destruct ] qed-. -lemma aaa_inv_lref: ∀G,L,A,i. ⦃G, L⦄ ⊢ #i ⁝ A → - ∃∃I,K,V. ⬇[i] L ≡ K. ⓑ{I} V & ⦃G, K⦄ ⊢ V ⁝ A. +lemma aaa_inv_lref: ∀G,L,A,i. ⦃G, L⦄ ⊢ #⫯i ⁝ A → + ∃∃I,K,V. L = K.ⓑ{I}V & ⦃G, K⦄ ⊢ #i ⁝ A. /2 width=3 by aaa_inv_lref_aux/ qed-. -fact aaa_inv_gref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀p. T = §p → ⊥. +fact aaa_inv_gref_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀l. T = §l → ⊥. #G #L #T #A * -G -L -T -A -[ #G #L #k #q #H destruct -| #I #G #L #K #V #B #i #HLK #HB #q #H destruct -| #a #G #L #V #T #B #A #_ #_ #q #H destruct -| #a #G #L #V #T #B #A #_ #_ #q #H destruct -| #G #L #V #T #B #A #_ #_ #q #H destruct -| #G #L #V #T #A #_ #_ #q #H destruct +[ #G #L #s #k #H destruct +| #I #G #L #V #B #_ #k #H destruct +| #I #G #L #V #A #i #_ #k #H destruct +| #p #G #L #V #T #B #A #_ #_ #k #H destruct +| #p #G #L #V #T #B #A #_ #_ #k #H destruct +| #G #L #V #T #B #A #_ #_ #k #H destruct +| #G #L #V #T #A #_ #_ #k #H destruct ] qed-. -lemma aaa_inv_gref: ∀G,L,A,p. ⦃G, L⦄ ⊢ §p ⁝ A → ⊥. +lemma aaa_inv_gref: ∀G,L,A,l. ⦃G, L⦄ ⊢ §l ⁝ A → ⊥. /2 width=7 by aaa_inv_gref_aux/ qed-. -fact aaa_inv_abbr_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀a,W,U. T = ⓓ{a}W. U → +fact aaa_inv_abbr_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀p,W,U. T = ⓓ{p}W.U → ∃∃B. ⦃G, L⦄ ⊢ W ⁝ B & ⦃G, L.ⓓW⦄ ⊢ U ⁝ A. #G #L #T #A * -G -L -T -A -[ #G #L #k #a #W #U #H destruct -| #I #G #L #K #V #B #i #_ #_ #a #W #U #H destruct -| #b #G #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=2 by ex2_intro/ -| #b #G #L #V #T #B #A #_ #_ #a #W #U #H destruct -| #G #L #V #T #B #A #_ #_ #a #W #U #H destruct -| #G #L #V #T #A #_ #_ #a #W #U #H destruct +[ #G #L #s #q #W #U #H destruct +| #I #G #L #V #B #_ #q #W #U #H destruct +| #I #G #L #V #A #i #_ #q #W #U #H destruct +| #p #G #L #V #T #B #A #HV #HT #q #W #U #H destruct /2 width=2 by ex2_intro/ +| #p #G #L #V #T #B #A #_ #_ #q #W #U #H destruct +| #G #L #V #T #B #A #_ #_ #q #W #U #H destruct +| #G #L #V #T #A #_ #_ #q #W #U #H destruct ] qed-. -lemma aaa_inv_abbr: ∀a,G,L,V,T,A. ⦃G, L⦄ ⊢ ⓓ{a}V. T ⁝ A → +lemma aaa_inv_abbr: ∀p,G,L,V,T,A. ⦃G, L⦄ ⊢ ⓓ{p}V.T ⁝ A → ∃∃B. ⦃G, L⦄ ⊢ V ⁝ B & ⦃G, L.ⓓV⦄ ⊢ T ⁝ A. /2 width=4 by aaa_inv_abbr_aux/ qed-. -fact aaa_inv_abst_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀a,W,U. T = ⓛ{a}W. U → +fact aaa_inv_abst_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀p,W,U. T = ⓛ{p}W.U → ∃∃B1,B2. ⦃G, L⦄ ⊢ W ⁝ B1 & ⦃G, L.ⓛW⦄ ⊢ U ⁝ B2 & A = ②B1.B2. #G #L #T #A * -G -L -T -A -[ #G #L #k #a #W #U #H destruct -| #I #G #L #K #V #B #i #_ #_ #a #W #U #H destruct -| #b #G #L #V #T #B #A #_ #_ #a #W #U #H destruct -| #b #G #L #V #T #B #A #HV #HT #a #W #U #H destruct /2 width=5 by ex3_2_intro/ -| #G #L #V #T #B #A #_ #_ #a #W #U #H destruct -| #G #L #V #T #A #_ #_ #a #W #U #H destruct +[ #G #L #s #q #W #U #H destruct +| #I #G #L #V #B #_ #q #W #U #H destruct +| #I #G #L #V #A #i #_ #q #W #U #H destruct +| #p #G #L #V #T #B #A #_ #_ #q #W #U #H destruct +| #p #G #L #V #T #B #A #HV #HT #q #W #U #H destruct /2 width=5 by ex3_2_intro/ +| #G #L #V #T #B #A #_ #_ #q #W #U #H destruct +| #G #L #V #T #A #_ #_ #q #W #U #H destruct ] qed-. -lemma aaa_inv_abst: ∀a,G,L,W,T,A. ⦃G, L⦄ ⊢ ⓛ{a}W. T ⁝ A → +lemma aaa_inv_abst: ∀p,G,L,W,T,A. ⦃G, L⦄ ⊢ ⓛ{p}W.T ⁝ A → ∃∃B1,B2. ⦃G, L⦄ ⊢ W ⁝ B1 & ⦃G, L.ⓛW⦄ ⊢ T ⁝ B2 & A = ②B1.B2. /2 width=4 by aaa_inv_abst_aux/ qed-. fact aaa_inv_appl_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓐW.U → ∃∃B. ⦃G, L⦄ ⊢ W ⁝ B & ⦃G, L⦄ ⊢ U ⁝ ②B.A. #G #L #T #A * -G -L -T -A -[ #G #L #k #W #U #H destruct -| #I #G #L #K #V #B #i #_ #_ #W #U #H destruct -| #a #G #L #V #T #B #A #_ #_ #W #U #H destruct -| #a #G #L #V #T #B #A #_ #_ #W #U #H destruct +[ #G #L #s #W #U #H destruct +| #I #G #L #V #B #_ #W #U #H destruct +| #I #G #L #V #A #i #_ #W #U #H destruct +| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct +| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct | #G #L #V #T #B #A #HV #HT #W #U #H destruct /2 width=3 by ex2_intro/ | #G #L #V #T #A #_ #_ #W #U #H destruct ] @@ -131,10 +153,11 @@ lemma aaa_inv_appl: ∀G,L,V,T,A. ⦃G, L⦄ ⊢ ⓐV.T ⁝ A → fact aaa_inv_cast_aux: ∀G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓝW. U → ⦃G, L⦄ ⊢ W ⁝ A ∧ ⦃G, L⦄ ⊢ U ⁝ A. #G #L #T #A * -G -L -T -A -[ #G #L #k #W #U #H destruct -| #I #G #L #K #V #B #i #_ #_ #W #U #H destruct -| #a #G #L #V #T #B #A #_ #_ #W #U #H destruct -| #a #G #L #V #T #B #A #_ #_ #W #U #H destruct +[ #G #L #s #W #U #H destruct +| #I #G #L #V #B #_ #W #U #H destruct +| #I #G #L #V #A #i #_ #W #U #H destruct +| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct +| #p #G #L #V #T #B #A #_ #_ #W #U #H destruct | #G #L #V #T #B #A #_ #_ #W #U #H destruct | #G #L #V #T #A #HV #HT #W #U #H destruct /2 width=1 by conj/ ] 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 76bf7e529..66d0f47d9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_aaa.ma @@ -21,7 +21,7 @@ include "basic_2/static/aaa.ma". theorem aaa_mono: ∀G,L,T,A1. ⦃G, L⦄ ⊢ T ⁝ A1 → ∀A2. ⦃G, L⦄ ⊢ T ⁝ A2 → A1 = A2. #G #L #T #A1 #H elim H -G -L -T -A1 -[ #G #L #k #A2 #H +[ #G #L #s #A2 #H >(aaa_inv_sort … H) -H // | #I1 #G #L #K1 #V1 #B #i #HLK1 #_ #IHA1 #A2 #H elim (aaa_inv_lref … H) -H #I2 #K2 #V2 #HLK2 #HA2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma new file mode 100644 index 000000000..9fc9a9ebd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma @@ -0,0 +1,106 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/drops.ma". +include "basic_2/static/aaa.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties with generic slicing for local environments *******************) + +(* Basic_2A1: was: aaa_lref *) +lemma aaa_lref_gen: ∀I,G,K,V,B,i,L. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ⁝ B → ⦃G, L⦄ ⊢ #i ⁝ B. +#I #G #K #V #B #i elim i -i +[ #L #H lapply (drops_fwd_isid … H ?) -H // + #H destruct /2 width=1 by aaa_zero/ +| #i #IH #L (lifts_inv_sort1 … H) -H // +| #I #G #L1 #V1 #B #_ #IHB #L2 #c #f #HL21 #T2 #H + elim (lifts_inv_lref1 … H) -H #i2 #Hi #H destruct + @aaa_lref_gen [5: @IHB ] +| #I #G #L1 #V1 #B #i1 #_ #IHB #L2 #c #f #HL21 #T2 #H + elim (lifts_inv_lref1 … H) -H #i2 #Hi #H destruct + lapply (at_inv_nxx … H) + + + + +| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #c #l #k #HL21 #T2 #H + elim (lift_inv_lref1 … H) -H * #Hil #H destruct + [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H + elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct + /3 width=9 by aaa_lref/ + | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 + /3 width=9 by aaa_lref, drop_inv_gen/ + ] +| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #c #l #k #HL21 #X #H + elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /4 width=5 by aaa_abbr, drop_skip/ +| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #c #l #k #HL21 #X #H + elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /4 width=5 by aaa_abst, drop_skip/ +| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #c #l #k #HL21 #X #H + elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /3 width=5 by aaa_appl/ +| #G #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #c #l #k #HL21 #X #H + elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /3 width=5 by aaa_cast/ +] +qed. + +(* Inversion lemmas with generic slicing for local environments *************) + +(* Basic_2A1: was: aaa_inv_lref *) +lemma aaa_inv_lref_gen: ∀G,A,i,L. ⦃G, L⦄ ⊢ #i ⁝ A → + ∃∃I,K,V. ⬇*[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A. +#G #A #i elim i -i +[ #L #H elim (aaa_inv_zero … H) -H /3 width=5 by drops_refl, ex2_3_intro/ +| #i #IH #L #H elim (aaa_inv_lref … H) -H + #I #K #V #H #HA destruct elim (IH … HA) -IH -HA /3 width=5 by drops_drop, ex2_3_intro/ +] +qed-. + +lemma aaa_inv_lift: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,c,l,k. ⬇[c, l, k] L2 ≡ L1 → + ∀T1. ⬆[l, k] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A. +#G #L2 #T2 #A #H elim H -G -L2 -T2 -A +[ #G #L2 #s #L1 #c #l #k #_ #T1 #H + >(lift_inv_sort2 … H) -H // +| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #c #l #k #HL21 #T1 #H + elim (lift_inv_lref2 … H) -H * #Hil #H destruct + [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=9 by aaa_lref/ + | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /3 width=9 by aaa_lref/ + ] +| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #c #l #k #HL21 #X #H + elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /4 width=5 by aaa_abbr, drop_skip/ +| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #c #l #k #HL21 #X #H + elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /4 width=5 by aaa_abst, drop_skip/ +| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #c #l #k #HL21 #X #H + elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /3 width=5 by aaa_appl/ +| #G #L2 #V2 #T2 #A #_ #_ #IH1 #IH2 #L1 #c #l #k #HL21 #X #H + elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + /3 width=5 by aaa_cast/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma deleted file mode 100644 index e83255951..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lift.ma +++ /dev/null @@ -1,73 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/substitution/drop_drop.ma". -include "basic_2/static/aaa.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties on basic relocation *******************************************) - -lemma aaa_lift: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀L2,s,l,m. ⬇[s, l, m] L2 ≡ L1 → - ∀T2. ⬆[l, m] T1 ≡ T2 → ⦃G, L2⦄ ⊢ T2 ⁝ A. -#G #L1 #T1 #A #H elim H -G -L1 -T1 -A -[ #G #L1 #k #L2 #s #l #m #_ #T2 #H - >(lift_inv_sort1 … H) -H // -| #I #G #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #s #l #m #HL21 #T2 #H - elim (lift_inv_lref1 … H) -H * #Hil #H destruct - [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H - elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct - /3 width=9 by aaa_lref/ - | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 - /3 width=9 by aaa_lref, drop_inv_gen/ - ] -| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #l #m #HL21 #X #H - elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /4 width=5 by aaa_abbr, drop_skip/ -| #a #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #l #m #HL21 #X #H - elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /4 width=5 by aaa_abst, drop_skip/ -| #G #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #s #l #m #HL21 #X #H - elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /3 width=5 by aaa_appl/ -| #G #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #s #l #m #HL21 #X #H - elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /3 width=5 by aaa_cast/ -] -qed. - -lemma aaa_inv_lift: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀L1,s,l,m. ⬇[s, l, m] L2 ≡ L1 → - ∀T1. ⬆[l, m] T1 ≡ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A. -#G #L2 #T2 #A #H elim H -G -L2 -T2 -A -[ #G #L2 #k #L1 #s #l #m #_ #T1 #H - >(lift_inv_sort2 … H) -H // -| #I #G #L2 #K2 #V2 #B #i #HLK2 #_ #IHB #L1 #s #l #m #HL21 #T1 #H - elim (lift_inv_lref2 … H) -H * #Hil #H destruct - [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=9 by aaa_lref/ - | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /3 width=9 by aaa_lref/ - ] -| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #l #m #HL21 #X #H - elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /4 width=5 by aaa_abbr, drop_skip/ -| #a #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #l #m #HL21 #X #H - elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /4 width=5 by aaa_abst, drop_skip/ -| #G #L2 #V2 #T2 #B #A #_ #_ #IHB #IHA #L1 #s #l #m #HL21 #X #H - elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /3 width=5 by aaa_appl/ -| #G #L2 #V2 #T2 #A #_ #_ #IH1 #IH2 #L1 #s #l #m #HL21 #X #H - elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /3 width=5 by aaa_cast/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma deleted file mode 100644 index cf972d088..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_lifts.ma +++ /dev/null @@ -1,30 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/multiple/drops.ma". -include "basic_2/static/aaa_lift.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties concerning generic relocation *********************************) - -lemma aaa_lifts: ∀G,L1,L2,T2,A,s,cs. ⬇*[s, cs] L2 ≡ L1 → ∀T1. ⬆*[cs] T1 ≡ T2 → - ⦃G, L1⦄ ⊢ T1 ⁝ A → ⦃G, L2⦄ ⊢ T2 ⁝ A. -#G #L1 #L2 #T2 #A #s #cs #H elim H -L1 -L2 -cs -[ #L #T1 #H #HT1 - <(lifts_inv_nil … H) -H // -| #L1 #L #L2 #cs #l #m #_ #HL2 #IHL1 #T1 #H #HT1 - elim (lifts_inv_cons … H) -H /3 width=10 by aaa_lift/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/da.ma b/matita/matita/contribs/lambdadelta/basic_2/static/da.ma deleted file mode 100644 index 9b1331820..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/da.ma +++ /dev/null @@ -1,108 +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/degree_6.ma". -include "basic_2/grammar/genv.ma". -include "basic_2/substitution/drop.ma". -include "basic_2/static/sd.ma". - -(* DEGREE ASSIGNMENT FOR TERMS **********************************************) - -(* activate genv *) -inductive da (h:sh) (g:sd h): relation4 genv lenv term nat ≝ -| da_sort: ∀G,L,k,d. deg h g k d → da h g G L (⋆k) d -| da_ldef: ∀G,L,K,V,i,d. ⬇[i] L ≡ K.ⓓV → da h g G K V d → da h g G L (#i) d -| da_ldec: ∀G,L,K,W,i,d. ⬇[i] L ≡ K.ⓛW → da h g G K W d → da h g G L (#i) (d+1) -| da_bind: ∀a,I,G,L,V,T,d. da h g G (L.ⓑ{I}V) T d → da h g G L (ⓑ{a,I}V.T) d -| da_flat: ∀I,G,L,V,T,d. da h g G L T d → da h g G L (ⓕ{I}V.T) d -. - -interpretation "degree assignment (term)" - 'Degree h g G L T d = (da h g G L T d). - -(* Basic inversion lemmas ***************************************************) - -fact da_inv_sort_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → - ∀k0. T = ⋆k0 → deg h g k0 d. -#h #g #G #L #T #d * -G -L -T -d -[ #G #L #k #d #Hkd #k0 #H destruct // -| #G #L #K #V #i #d #_ #_ #k0 #H destruct -| #G #L #K #W #i #d #_ #_ #k0 #H destruct -| #a #I #G #L #V #T #d #_ #k0 #H destruct -| #I #G #L #V #T #d #_ #k0 #H destruct -] -qed-. - -lemma da_inv_sort: ∀h,g,G,L,k,d. ⦃G, L⦄ ⊢ ⋆k ▪[h, g] d → deg h g k d. -/2 width=5 by da_inv_sort_aux/ qed-. - -fact da_inv_lref_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → ∀j. T = #j → - (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] d) ∨ - (∃∃K,W,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] d0 & - d = d0 + 1 - ). -#h #g #G #L #T #d * -G -L -T -d -[ #G #L #k #d #_ #j #H destruct -| #G #L #K #V #i #d #HLK #HV #j #H destruct /3 width=4 by ex2_2_intro, or_introl/ -| #G #L #K #W #i #d #HLK #HW #j #H destruct /3 width=6 by ex3_3_intro, or_intror/ -| #a #I #G #L #V #T #d #_ #j #H destruct -| #I #G #L #V #T #d #_ #j #H destruct -] -qed-. - -lemma da_inv_lref: ∀h,g,G,L,j,d. ⦃G, L⦄ ⊢ #j ▪[h, g] d → - (∃∃K,V. ⬇[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ▪[h, g] d) ∨ - (∃∃K,W,d0. ⬇[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W ▪[h, g] d0 & d = d0+1). -/2 width=3 by da_inv_lref_aux/ qed-. - -fact da_inv_gref_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → ∀p0. T = §p0 → ⊥. -#h #g #G #L #T #d * -G -L -T -d -[ #G #L #k #d #_ #p0 #H destruct -| #G #L #K #V #i #d #_ #_ #p0 #H destruct -| #G #L #K #W #i #d #_ #_ #p0 #H destruct -| #a #I #G #L #V #T #d #_ #p0 #H destruct -| #I #G #L #V #T #d #_ #p0 #H destruct -] -qed-. - -lemma da_inv_gref: ∀h,g,G,L,p,d. ⦃G, L⦄ ⊢ §p ▪[h, g] d → ⊥. -/2 width=9 by da_inv_gref_aux/ qed-. - -fact da_inv_bind_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → - ∀b,J,X,Y. T = ⓑ{b,J}Y.X → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] d. -#h #g #G #L #T #d * -G -L -T -d -[ #G #L #k #d #_ #b #J #X #Y #H destruct -| #G #L #K #V #i #d #_ #_ #b #J #X #Y #H destruct -| #G #L #K #W #i #d #_ #_ #b #J #X #Y #H destruct -| #a #I #G #L #V #T #d #HT #b #J #X #Y #H destruct // -| #I #G #L #V #T #d #_ #b #J #X #Y #H destruct -] -qed-. - -lemma da_inv_bind: ∀h,g,b,J,G,L,Y,X,d. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X ▪[h, g] d → ⦃G, L.ⓑ{J}Y⦄ ⊢ X ▪[h, g] d. -/2 width=4 by da_inv_bind_aux/ qed-. - -fact da_inv_flat_aux: ∀h,g,G,L,T,d. ⦃G, L⦄ ⊢ T ▪[h, g] d → - ∀J,X,Y. T = ⓕ{J}Y.X → ⦃G, L⦄ ⊢ X ▪[h, g] d. -#h #g #G #L #T #d * -G -L -T -d -[ #G #L #k #d #_ #J #X #Y #H destruct -| #G #L #K #V #i #d #_ #_ #J #X #Y #H destruct -| #G #L #K #W #i #d #_ #_ #J #X #Y #H destruct -| #a #I #G #L #V #T #d #_ #J #X #Y #H destruct -| #I #G #L #V #T #d #HT #J #X #Y #H destruct // -] -qed-. - -lemma da_inv_flat: ∀h,g,J,G,L,Y,X,d. ⦃G, L⦄ ⊢ ⓕ{J}Y.X ▪[h, g] d → ⦃G, L⦄ ⊢ X ▪[h, g] d. -/2 width=5 by da_inv_flat_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/da_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/static/da_aaa.ma deleted file mode 100644 index c7903c24e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/da_aaa.ma +++ /dev/null @@ -1,31 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/aaa_lift.ma". -include "basic_2/static/da.ma". - -(* DEGREE ASSIGNMENT FOR TERMS **********************************************) - -(* Properties on atomic arity assignment for terms **************************) - -lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃d. ⦃G, L⦄ ⊢ T ▪[h, g] d. -#h #g #G #L #T #A #H elim H -G -L -T -A -[ #G #L #k elim (deg_total h g k) /3 width=2 by da_sort, ex_intro/ -| * #G #L #K #V #B #i #HLK #_ * /3 width=5 by da_ldef, da_ldec, ex_intro/ -| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by da_bind, ex_intro/ -| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by da_bind, ex_intro/ -| #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by da_flat, ex_intro/ -| #G #L #W #T #A #_ #_ #_ * /3 width=2 by da_flat, ex_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/da_da.ma b/matita/matita/contribs/lambdadelta/basic_2/static/da_da.ma deleted file mode 100644 index c670e0a5d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/da_da.ma +++ /dev/null @@ -1,38 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/da_lift.ma". - -(* DEGREE ASSIGNMENT FOR TERMS **********************************************) - -(* Main properties **********************************************************) - -theorem da_mono: ∀h,g,G,L,T,d1. ⦃G, L⦄ ⊢ T ▪[h, g] d1 → - ∀d2. ⦃G, L⦄ ⊢ T ▪[h, g] d2 → d1 = d2. -#h #g #G #L #T #d1 #H elim H -G -L -T -d1 -[ #G #L #k #d1 #Hkd1 #d2 #H - lapply (da_inv_sort … H) -G -L #Hkd2 - >(deg_mono … Hkd2 … Hkd1) -h -k -d2 // -| #G #L #K #V #i #d1 #HLK #_ #IHV #d2 #H - elim (da_inv_lref … H) -H * #K0 #V0 [| #d0 ] #HLK0 #HV0 [| #Hd0 ] - lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /2 width=1 by/ -| #G #L #K #W #i #d1 #HLK #_ #IHW #d2 #H - elim (da_inv_lref … H) -H * #K0 #W0 [| #d0 ] #HLK0 #HW0 [| #Hd0 ] - lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct /3 width=1 by eq_f/ -| #a #I #G #L #V #T #d1 #_ #IHT #d2 #H - lapply (da_inv_bind … H) -H /2 width=1 by/ -| #I #G #L #V #T #d1 #_ #IHT #d2 #H - lapply (da_inv_flat … H) -H /2 width=1 by/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma deleted file mode 100644 index 25cde075b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/da_lift.ma +++ /dev/null @@ -1,78 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/substitution/drop_drop.ma". -include "basic_2/static/da.ma". - -(* DEGREE ASSIGNMENT FOR TERMS **********************************************) - -(* Properties on relocation *************************************************) - -lemma da_lift: ∀h,g,G,L1,T1,d. ⦃G, L1⦄ ⊢ T1 ▪[h, g] d → - ∀L2,s,l,m. ⬇[s, l, m] L2 ≡ L1 → ∀T2. ⬆[l, m] T1 ≡ T2 → - ⦃G, L2⦄ ⊢ T2 ▪[h, g] d. -#h #g #G #L1 #T1 #d #H elim H -G -L1 -T1 -d -[ #G #L1 #k #d #Hkd #L2 #s #l #m #_ #X #H - >(lift_inv_sort1 … H) -X /2 width=1 by da_sort/ -| #G #L1 #K1 #V1 #i #d #HLK1 #_ #IHV1 #L2 #s #l #m #HL21 #X #H - elim (lift_inv_lref1 … H) * #Hil #H destruct - [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H - elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #V2 #HK21 #HV12 #H destruct - /3 width=9 by da_ldef/ - | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 - /3 width=8 by da_ldef, drop_inv_gen/ - ] -| #G #L1 #K1 #W1 #i #d #HLK1 #_ #IHW1 #L2 #s #l #m #HL21 #X #H - elim (lift_inv_lref1 … H) * #Hil #H destruct - [ elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by ylt_fwd_le/ #X #HLK2 #H - elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K2 #W2 #HK21 #HW12 #H destruct - /3 width=8 by da_ldec/ - | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 - /3 width=8 by da_ldec, drop_inv_gen/ - ] -| #a #I #G #L1 #V1 #T1 #d #_ #IHT1 #L2 #s #l #m #HL21 #X #H - elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct - /4 width=5 by da_bind, drop_skip/ -| #I #G #L1 #V1 #T1 #d #_ #IHT1 #L2 #s #l #m #HL21 #X #H - elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HU12 #H destruct - /3 width=5 by da_flat/ -] -qed. - -(* Inversion lemmas on relocation *******************************************) - -lemma da_inv_lift: ∀h,g,G,L2,T2,d. ⦃G, L2⦄ ⊢ T2 ▪[h, g] d → - ∀L1,s,l,m. ⬇[s, l, m] L2 ≡ L1 → ∀T1. ⬆[l, m] T1 ≡ T2 → - ⦃G, L1⦄ ⊢ T1 ▪[h, g] d. -#h #g #G #L2 #T2 #d #H elim H -G -L2 -T2 -d -[ #G #L2 #k #d #Hkd #L1 #s #l #m #_ #X #H - >(lift_inv_sort2 … H) -X /2 width=1 by da_sort/ -| #G #L2 #K2 #V2 #i #d #HLK2 #HV2 #IHV2 #L1 #s #l #m #HL21 #X #H - elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HV2 | -IHV2 ] - [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldef/ - | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldef/ - ] -| #G #L2 #K2 #W2 #i #d #HLK2 #HW2 #IHW2 #L1 #s #l #m #HL21 #X #H - elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HW2 | -IHW2 ] - [ elim (drop_conf_lt … HL21 … HLK2) -L2 /3 width=8 by da_ldec/ - | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 /2 width=4 by da_ldec/ - ] -| #a #I #G #L2 #V2 #T2 #d #_ #IHT2 #L1 #s #l #m #HL21 #X #H - elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /4 width=5 by da_bind, drop_skip/ -| #I #G #L2 #V2 #T2 #d #_ #IHT2 #L1 #s #l #m #HL21 #X #H - elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - /3 width=5 by da_flat/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma index e1869f384..613a8f321 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma @@ -100,21 +100,21 @@ lemma lsuba_refl: ∀G,L. G ⊢ L ⫃⁝ L. qed. (* Note: the constant 0 cannot be generalized *) -lemma lsuba_drop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K1,s,m. ⬇[s, 0, m] L1 ≡ K1 → - ∃∃K2. G ⊢ K1 ⫃⁝ K2 & ⬇[s, 0, m] L2 ≡ K2. +lemma lsuba_drop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K1,c,k. ⬇[c, 0, k] L1 ≡ K1 → + ∃∃K2. G ⊢ K1 ⫃⁝ K2 & ⬇[c, 0, k] L2 ≡ K2. #G #L1 #L2 #H elim H -L1 -L2 [ /2 width=3 by ex2_intro/ -| #I #L1 #L2 #V #_ #IHL12 #K1 #s #m #H +| #I #L1 #L2 #V #_ #IHL12 #K1 #c #k #H elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1 [ destruct - elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H + elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, drop_pair, ex2_intro/ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/ ] -| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K1 #s #m #H +| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K1 #c #k #H elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1 [ destruct - elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H + elim (IHL12 L1 c 0) -IHL12 // #X #HL12 #H <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_beta, drop_pair, ex2_intro/ | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/ ] @@ -122,21 +122,21 @@ lemma lsuba_drop_O1_conf: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K1,s,m. ⬇[s, 0 qed-. (* Note: the constant 0 cannot be generalized *) -lemma lsuba_drop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 → - ∃∃K1. G ⊢ K1 ⫃⁝ K2 & ⬇[s, 0, m] L1 ≡ K1. +lemma lsuba_drop_O1_trans: ∀G,L1,L2. G ⊢ L1 ⫃⁝ L2 → ∀K2,c,k. ⬇[c, 0, k] L2 ≡ K2 → + ∃∃K1. G ⊢ K1 ⫃⁝ K2 & ⬇[c, 0, k] L1 ≡ K1. #G #L1 #L2 #H elim H -L1 -L2 [ /2 width=3 by ex2_intro/ -| #I #L1 #L2 #V #_ #IHL12 #K2 #s #m #H +| #I #L1 #L2 #V #_ #IHL12 #K2 #c #k #H elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2 [ destruct - elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H + elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_pair, drop_pair, ex2_intro/ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/ ] -| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K2 #s #m #H +| #L1 #L2 #W #V #A #HV #HW #_ #IHL12 #K2 #c #k #H elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2 [ destruct - elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H + elim (IHL12 L2 c 0) -IHL12 // #X #HL12 #H <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsuba_beta, drop_pair, ex2_intro/ | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/ ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma deleted file mode 100644 index 29da4dad7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma +++ /dev/null @@ -1,151 +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/lrsubeqd_5.ma". -include "basic_2/static/lsubr.ma". -include "basic_2/static/da.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) - -inductive lsubd (h) (g) (G): relation lenv ≝ -| lsubd_atom: lsubd h g G (⋆) (⋆) -| lsubd_pair: ∀I,L1,L2,V. lsubd h g G L1 L2 → - lsubd h g G (L1.ⓑ{I}V) (L2.ⓑ{I}V) -| lsubd_beta: ∀L1,L2,W,V,d. ⦃G, L1⦄ ⊢ V ▪[h, g] d+1 → ⦃G, L2⦄ ⊢ W ▪[h, g] d → - lsubd h g G L1 L2 → lsubd h g G (L1.ⓓⓝW.V) (L2.ⓛW) -. - -interpretation - "local environment refinement (degree assignment)" - 'LRSubEqD h g G L1 L2 = (lsubd h g G L1 L2). - -(* Basic forward lemmas *****************************************************) - -lemma lsubd_fwd_lsubr: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → L1 ⫃ L2. -#h #g #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubr_pair, lsubr_beta/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact lsubd_inv_atom1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → L1 = ⋆ → L2 = ⋆. -#h #g #G #L1 #L2 * -L1 -L2 -[ // -| #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #W #V #d #_ #_ #_ #H destruct -] -qed-. - -lemma lsubd_inv_atom1: ∀h,g,G,L2. G ⊢ ⋆ ⫃▪[h, g] L2 → L2 = ⋆. -/2 width=6 by lsubd_inv_atom1_aux/ qed-. - -fact lsubd_inv_pair1_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → - ∀I,K1,X. L1 = K1.ⓑ{I}X → - (∃∃K2. G ⊢ K1 ⫃▪[h, g] K2 & L2 = K2.ⓑ{I}X) ∨ - ∃∃K2,W,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d & - G ⊢ K1 ⫃▪[h, g] K2 & - I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. -#h #g #G #L1 #L2 * -L1 -L2 -[ #J #K1 #X #H destruct -| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by ex2_intro, or_introl/ -| #L1 #L2 #W #V #d #HV #HW #HL12 #J #K1 #X #H destruct /3 width=9 by ex6_4_intro, or_intror/ -] -qed-. - -lemma lsubd_inv_pair1: ∀h,g,I,G,K1,L2,X. G ⊢ K1.ⓑ{I}X ⫃▪[h, g] L2 → - (∃∃K2. G ⊢ K1 ⫃▪[h, g] K2 & L2 = K2.ⓑ{I}X) ∨ - ∃∃K2,W,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d & - G ⊢ K1 ⫃▪[h, g] K2 & - I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. -/2 width=3 by lsubd_inv_pair1_aux/ qed-. - -fact lsubd_inv_atom2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → L2 = ⋆ → L1 = ⋆. -#h #g #G #L1 #L2 * -L1 -L2 -[ // -| #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #W #V #d #_ #_ #_ #H destruct -] -qed-. - -lemma lsubd_inv_atom2: ∀h,g,G,L1. G ⊢ L1 ⫃▪[h, g] ⋆ → L1 = ⋆. -/2 width=6 by lsubd_inv_atom2_aux/ qed-. - -fact lsubd_inv_pair2_aux: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → - ∀I,K2,W. L2 = K2.ⓑ{I}W → - (∃∃K1. G ⊢ K1 ⫃▪[h, g] K2 & L1 = K1.ⓑ{I}W) ∨ - ∃∃K1,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d & - G ⊢ K1 ⫃▪[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. -#h #g #G #L1 #L2 * -L1 -L2 -[ #J #K2 #U #H destruct -| #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3 by ex2_intro, or_introl/ -| #L1 #L2 #W #V #d #HV #HW #HL12 #J #K2 #U #H destruct /3 width=7 by ex5_3_intro, or_intror/ -] -qed-. - -lemma lsubd_inv_pair2: ∀h,g,I,G,L1,K2,W. G ⊢ L1 ⫃▪[h, g] K2.ⓑ{I}W → - (∃∃K1. G ⊢ K1 ⫃▪[h, g] K2 & L1 = K1.ⓑ{I}W) ∨ - ∃∃K1,V,d. ⦃G, K1⦄ ⊢ V ▪[h, g] d+1 & ⦃G, K2⦄ ⊢ W ▪[h, g] d & - G ⊢ K1 ⫃▪[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. -/2 width=3 by lsubd_inv_pair2_aux/ qed-. - -(* Basic properties *********************************************************) - -lemma lsubd_refl: ∀h,g,G,L. G ⊢ L ⫃▪[h, g] L. -#h #g #G #L elim L -L /2 width=1 by lsubd_pair/ -qed. - -(* Note: the constant 0 cannot be generalized *) -lemma lsubd_drop_O1_conf: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → - ∀K1,s,m. ⬇[s, 0, m] L1 ≡ K1 → - ∃∃K2. G ⊢ K1 ⫃▪[h, g] K2 & ⬇[s, 0, m] L2 ≡ K2. -#h #g #G #L1 #L2 #H elim H -L1 -L2 -[ /2 width=3 by ex2_intro/ -| #I #L1 #L2 #V #_ #IHL12 #K1 #s #m #H - elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1 - [ destruct - elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H - <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, drop_pair, ex2_intro/ - | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/ - ] -| #L1 #L2 #W #V #d #HV #HW #_ #IHL12 #K1 #s #m #H - elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK1 - [ destruct - elim (IHL12 L1 s 0) -IHL12 // #X #HL12 #H - <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_beta, drop_pair, ex2_intro/ - | elim (IHL12 … HLK1) -L1 /3 width=3 by drop_drop_lt, ex2_intro/ - ] -] -qed-. - -(* Note: the constant 0 cannot be generalized *) -lemma lsubd_drop_O1_trans: ∀h,g,G,L1,L2. G ⊢ L1 ⫃▪[h, g] L2 → - ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 → - ∃∃K1. G ⊢ K1 ⫃▪[h, g] K2 & ⬇[s, 0, m] L1 ≡ K1. -#h #g #G #L1 #L2 #H elim H -L1 -L2 -[ /2 width=3 by ex2_intro/ -| #I #L1 #L2 #V #_ #IHL12 #K2 #s #m #H - elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2 - [ destruct - elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H - <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_pair, drop_pair, ex2_intro/ - | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/ - ] -| #L1 #L2 #W #V #d #HV #HW #_ #IHL12 #K2 #s #m #H - elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK2 - [ destruct - elim (IHL12 L2 s 0) -IHL12 // #X #HL12 #H - <(drop_inv_O2 … H) in HL12; -H /3 width=3 by lsubd_beta, drop_pair, ex2_intro/ - | elim (IHL12 … HLK2) -L2 /3 width=3 by drop_drop_lt, ex2_intro/ - ] -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd_da.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd_da.ma deleted file mode 100644 index 111a0103a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd_da.ma +++ /dev/null @@ -1,65 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/da_da.ma". -include "basic_2/static/lsubd.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) - -(* Properties on degree assignment ******************************************) - -lemma lsubd_da_trans: ∀h,g,G,L2,T,d. ⦃G, L2⦄ ⊢ T ▪[h, g] d → - ∀L1. G ⊢ L1 ⫃▪[h, g] L2 → ⦃G, L1⦄ ⊢ T ▪[h, g] d. -#h #g #G #L2 #T #d #H elim H -G -L2 -T -d -[ /2 width=1 by da_sort/ -| #G #L2 #K2 #V #i #d #HLK2 #_ #IHV #L1 #HL12 - elim (lsubd_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 - elim (lsubd_inv_pair2 … H) -H * #K1 [ | -IHV -HLK1 ] - [ #HK12 #H destruct /3 width=4 by da_ldef/ - | #W #d0 #_ #_ #_ #H destruct - ] -| #G #L2 #K2 #W #i #d #HLK2 #HW #IHW #L1 #HL12 - elim (lsubd_drop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 - elim (lsubd_inv_pair2 … H) -H * #K1 [ -HW | -IHW ] - [ #HK12 #H destruct /3 width=4 by da_ldec/ - | #V #d0 #HV #H0W #_ #_ #H destruct - lapply (da_mono … H0W … HW) -H0W -HW #H destruct /3 width=7 by da_ldef, da_flat/ - ] -| /4 width=1 by lsubd_pair, da_bind/ -| /3 width=1 by da_flat/ -] -qed-. - -lemma lsubd_da_conf: ∀h,g,G,L1,T,d. ⦃G, L1⦄ ⊢ T ▪[h, g] d → - ∀L2. G ⊢ L1 ⫃▪[h, g] L2 → ⦃G, L2⦄ ⊢ T ▪[h, g] d. -#h #g #G #L1 #T #d #H elim H -G -L1 -T -d -[ /2 width=1 by da_sort/ -| #G #L1 #K1 #V #i #d #HLK1 #HV #IHV #L2 #HL12 - elim (lsubd_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2 - elim (lsubd_inv_pair1 … H) -H * #K2 [ -HV | -IHV ] - [ #HK12 #H destruct /3 width=4 by da_ldef/ - | #W0 #V0 #d0 #HV0 #HW0 #_ #_ #H1 #H2 destruct - lapply (da_inv_flat … HV) -HV #H0V0 - lapply (da_mono … H0V0 … HV0) -H0V0 -HV0 #H destruct /2 width=4 by da_ldec/ - ] -| #G #L1 #K1 #W #i #d #HLK1 #HW #IHW #L2 #HL12 - elim (lsubd_drop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2 - elim (lsubd_inv_pair1 … H) -H * #K2 [ -HW | -IHW ] - [ #HK12 #H destruct /3 width=4 by da_ldec/ - | #W0 #V0 #d0 #HV0 #HW0 #_ #H destruct - ] -| /4 width=1 by lsubd_pair, da_bind/ -| /3 width=1 by da_flat/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd_lsubd.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd_lsubd.ma deleted file mode 100644 index 270cdbd9d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd_lsubd.ma +++ /dev/null @@ -1,36 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/static/lsubd_da.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) - -(* Main properties **********************************************************) - -theorem lsubd_trans: ∀h,g,G. Transitive … (lsubd h g G). -#h #g #G #L1 #L #H elim H -L1 -L -[ #X #H >(lsubd_inv_atom1 … H) -H // -| #I #L1 #L #Y #HL1 #IHL1 #X #H - elim (lsubd_inv_pair1 … H) -H * #L2 - [ #HL2 #H destruct /3 width=1 by lsubd_pair/ - | #W #V #d #HV #HW #HL2 #H1 #H2 #H3 destruct - /3 width=3 by lsubd_beta, lsubd_da_trans/ - ] -| #L1 #L #W #V #d #HV #HW #HL1 #IHL1 #X #H - elim (lsubd_inv_pair1 … H) -H * #L2 - [ #HL2 #H destruct /3 width=5 by lsubd_beta, lsubd_da_conf/ - | #W0 #V0 #d0 #_ #_ #_ #H destruct - ] -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr.ma index 111da1751..5845bcb63 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr.ma @@ -13,9 +13,9 @@ (**************************************************************************) include "basic_2/notation/relations/lrsubeqc_2.ma". -include "basic_2/substitution/drop.ma". +include "basic_2/grammar/lenv.ma". -(* RESTRICTED LOCAL ENVIRONMENT REFINEMENT **********************************) +(* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************) inductive lsubr: relation lenv ≝ | lsubr_atom: ∀L. lsubr L (⋆) @@ -70,38 +70,3 @@ qed-. lemma lsubr_inv_abbr2: ∀L1,K2,W. L1 ⫃ K2.ⓓW → ∃∃K1. K1 ⫃ K2 & L1 = K1.ⓓW. /2 width=3 by lsubr_inv_abbr2_aux/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lsubr_fwd_length: ∀L1,L2. L1 ⫃ L2 → |L2| ≤ |L1|. -#L1 #L2 #H elim H -L1 -L2 /2 width=1 by monotonic_le_plus_l/ -qed-. - -lemma lsubr_fwd_drop2_pair: ∀L1,L2. L1 ⫃ L2 → - ∀I,K2,W,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W → - (∃∃K1. K1 ⫃ K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W) ∨ - ∃∃K1,V. K1 ⫃ K2 & ⬇[s, 0, i] L1 ≡ K1.ⓓⓝW.V & I = Abst. -#L1 #L2 #H elim H -L1 -L2 -[ #L #I #K2 #W #s #i #H - elim (drop_inv_atom1 … H) -H #H destruct -| #J #L1 #L2 #V #HL12 #IHL12 #I #K2 #W #s #i #H - elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] - [ /3 width=3 by drop_pair, ex2_intro, or_introl/ - | elim (IHL12 … HLK2) -IHL12 -HLK2 * - /4 width=4 by drop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/ - ] -| #L1 #L2 #V1 #V2 #HL12 #IHL12 #I #K2 #W #s #i #H - elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] - [ /3 width=4 by drop_pair, ex3_2_intro, or_intror/ - | elim (IHL12 … HLK2) -IHL12 -HLK2 * - /4 width=4 by drop_drop_lt, ex3_2_intro, ex2_intro, or_introl, or_intror/ - ] -] -qed-. - -lemma lsubr_fwd_drop2_abbr: ∀L1,L2. L1 ⫃ L2 → - ∀K2,V,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓓV → - ∃∃K1. K1 ⫃ K2 & ⬇[s, 0, i] L1 ≡ K1.ⓓV. -#L1 #L2 #HL12 #K2 #V #s #i #HLK2 elim (lsubr_fwd_drop2_pair … HL12 … HLK2) -L2 // * -#K1 #W #_ #_ #H destruct -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_drops.ma new file mode 100644 index 000000000..ca2154608 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_drops.ma @@ -0,0 +1,52 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/drops.ma". +include "basic_2/static/lsubr.ma". + +(* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************) + +(* Forward lemmas with generic slicing for local environments ***************) + +(* Basic_2A1: includes: lsubr_fwd_drop2_pair *) +lemma lsubr_fwd_drops2_pair: ∀L1,L2. L1 ⫃ L2 → + ∀I,K2,W,c,f. 𝐔⦃f⦄ → ⬇*[c, f] L2 ≡ K2.ⓑ{I}W → + (∃∃K1. K1 ⫃ K2 & ⬇*[c, f] L1 ≡ K1.ⓑ{I}W) ∨ + ∃∃K1,V. K1 ⫃ K2 & ⬇*[c, f] L1 ≡ K1.ⓓⓝW.V & I = Abst. +#L1 #L2 #H elim H -L1 -L2 +[ #L #I #K2 #W #c #f #_ #H + elim (drops_inv_atom1 … H) -H #H destruct +| #J #L1 #L2 #V #HL12 #IHL12 #I #K2 #W #c #f #Hf #H + elim (drops_inv_pair1_isuni … Hf H) -H * #g #HLK2 try #H destruct [ -IHL12 | -HL12 ] + [ /4 width=4 by drops_refl, ex2_intro, or_introl/ + | elim (IHL12 … HLK2) -IHL12 -HLK2 /2 width=3 by isuni_inv_next/ * + /4 width=4 by drops_drop, ex3_2_intro, ex2_intro, or_introl, or_intror/ + ] +| #L1 #L2 #V1 #V2 #HL12 #IHL12 #I #K2 #W #c #f #Hf #H + elim (drops_inv_pair1_isuni … Hf H) -H * #g #HLK2 try #H destruct [ -IHL12 | -HL12 ] + [ /4 width=4 by drops_refl, ex3_2_intro, or_intror/ + | elim (IHL12 … HLK2) -IHL12 -HLK2 /2 width=3 by isuni_inv_next/ * + /4 width=4 by drops_drop, ex3_2_intro, ex2_intro, or_introl, or_intror/ + ] +] +qed-. + +(* Basic_2A1: includes: lsubr_fwd_drop2_abbr *) +lemma lsubr_fwd_drops2_abbr: ∀L1,L2. L1 ⫃ L2 → + ∀K2,V,c,f. 𝐔⦃f⦄ → ⬇*[c, f] L2 ≡ K2.ⓓV → + ∃∃K1. K1 ⫃ K2 & ⬇*[c, f] L1 ≡ K1.ⓓV. +#L1 #L2 #HL12 #K2 #V #c #f #Hf #HLK2 +elim (lsubr_fwd_drops2_pair … HL12 … Hf HLK2) -L2 -Hf // * +#K1 #W #_ #_ #H destruct +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_length.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_length.ma new file mode 100644 index 000000000..e240002bf --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_length.ma @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/grammar/lenv_length.ma". +include "basic_2/static/lsubr.ma". + +(* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************) + +(* Forward lemmas with length for local environments ************************) + +lemma lsubr_fwd_length: ∀L1,L2. L1 ⫃ L2 → |L2| ≤ |L1|. +#L1 #L2 #H elim H -L1 -L2 /2 width=1 by O, le_S_S/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma index f0e171e0b..787a19377 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubr_lsubr.ma @@ -14,7 +14,7 @@ include "basic_2/static/lsubr.ma". -(* RESTRICTED LOCAL ENVIRONMENT REFINEMENT **********************************) +(* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************) (* Auxiliary inversion lemmas ***********************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma b/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma index 2bf43b5ee..bf7534529 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma @@ -19,113 +19,106 @@ include "basic_2/static/sh.ma". (* sort degree specification *) record sd (h:sh): Type[0] ≝ { deg : relation nat; (* degree of the sort *) - deg_total: ∀k. ∃d. deg k d; (* functional relation axioms *) - deg_mono : ∀k,d1,d2. deg k d1 → deg k d2 → d1 = d2; - deg_next : ∀k,d. deg k d → deg (next h k) (d - 1) (* compatibility condition *) + deg_total: ∀s. ∃d. deg s d; (* functional relation axioms *) + deg_mono : ∀s,d1,d2. deg s d1 → deg s d2 → d1 = d2; + deg_next : ∀s,d. deg s d → deg (next h s) (⫰d) (* compatibility condition *) }. (* Notable specifications ***************************************************) -definition deg_O: relation nat ≝ λk,d. d = 0. +definition deg_O: relation nat ≝ λs,d. d = 0. definition sd_O: ∀h. sd h ≝ λh. mk_sd h deg_O …. /2 width=2 by le_n_O_to_eq, le_n, ex_intro/ defined. -inductive deg_SO (h:sh) (k:nat) (k0:nat): predicate nat ≝ -| deg_SO_pos : ∀d0. (next h)^d0 k0 = k → deg_SO h k k0 (d0 + 1) -| deg_SO_zero: ((∃d0. (next h)^d0 k0 = k) → ⊥) → deg_SO h k k0 0 +(* Basic_2A1: includes: deg_SO_pos *) +inductive deg_SO (h:sh) (s:nat) (s0:nat): predicate nat ≝ +| deg_SO_succ : ∀n. (next h)^n s0 = s → deg_SO h s s0 (⫯n) +| deg_SO_zero: ((∃n. (next h)^n s0 = s) → ⊥) → deg_SO h s s0 0 . -fact deg_SO_inv_pos_aux: ∀h,k,k0,d0. deg_SO h k k0 d0 → ∀d. d0 = d + 1 → - (next h)^d k0 = k. -#h #k #k0 #d0 * -d0 -[ #d0 #Hd0 #d #H - lapply (injective_plus_l … H) -H #H destruct // -| #_ #d0 H -H #H - lapply (transitive_lt … H HK12) -k1 #H1 - lapply (nexts_le h k2 d) #H2 - lapply (le_to_lt_to_lt … H2 H1) -h -d #H +| #n #_ #H + lapply (next_lt h ((next h)^n s2)) >H -H #H + lapply (transitive_lt … H HK12) -s1 #H1 + lapply (nexts_le h s2 n) #H2 + lapply (le_to_lt_to_lt … H2 H1) -h -n #H elim (lt_refl_false … H) ] qed. -definition sd_SO: ∀h. nat → sd h ≝ λh,k. mk_sd h (deg_SO h k) …. -[ #k0 - lapply (nexts_dec h k0 k) * - [ * /3 width=2 by deg_SO_pos, ex_intro/ | /4 width=2 by deg_SO_zero, ex_intro/ ] -| #K0 #d1 #d2 * [ #d01 ] #H1 * [1,3: #d02 ] #H2 // +definition sd_SO: ∀h. nat → sd h ≝ λh,s. mk_sd h (deg_SO h s) …. +[ #s0 + lapply (nexts_dec h s0 s) * + [ * /3 width=2 by deg_SO_succ, ex_intro/ | /4 width=2 by deg_SO_zero, ex_intro/ ] +| #K0 #d1 #d2 * [ #n1 ] #H1 * [1,3: #n2 ] #H2 // [ < H2 in H1; -H2 #H lapply (nexts_inj … H) -H #H destruct // | elim H1 /2 width=2 by ex_intro/ | elim H2 /2 width=2 by ex_intro/ ] -| #k0 #d0 * +| #s0 #n * [ #d #H destruct elim d -d normalize - /2 width=1 by deg_SO_gt, deg_SO_pos, next_lt/ + /2 width=1 by deg_SO_gt, deg_SO_succ, next_lt/ | #H1 @deg_SO_zero * #d #H2 destruct - @H1 -H1 @(ex_intro … (S d)) /2 width=1 by sym_eq/ (**) (* explicit constructor *) + @H1 -H1 @(ex_intro … (⫯d)) /2 width=1 by sym_eq/ (**) (* explicit constructor *) ] ] defined. -rec definition sd_d (h:sh) (k:nat) (d:nat) on d : sd h ≝ +rec definition sd_d (h:sh) (s:nat) (d:nat) on d : sd h ≝ match d with [ O ⇒ sd_O h | S d ⇒ match d with - [ O ⇒ sd_SO h k - | _ ⇒ sd_d h (next h k) d + [ O ⇒ sd_SO h s + | _ ⇒ sd_d h (next h s) d ] ]. (* Basic inversion lemmas ***************************************************) -lemma deg_inv_pred: ∀h,g,k,d. deg h g (next h k) (d+1) → deg h g k (d+2). -#h #g #k #d #H1 -elim (deg_total h g k) #d0 #H0 +lemma deg_inv_pred: ∀h,o,s,d. deg h o (next h s) (⫯d) → deg h o s (⫯⫯d). +#h #o #s #d #H1 +elim (deg_total h o s) #n #H0 lapply (deg_next … H0) #H2 -lapply (deg_mono … H1 H2) -H1 -H2 #H -<(associative_plus d 1 1) >H H >S_pred /2 width=2 by ltn_to_ltO/ qed-. -lemma deg_inv_prec: ∀h,g,k,d,d0. deg h g ((next h)^d k) (d0+1) → deg h g k (d+d0+1). -#h #g #k #d @(nat_ind_plus … d) -d // -#d #IHd #d0 >iter_SO #H -lapply (deg_inv_pred … H) -H <(associative_plus d0 1 1) #H -lapply (IHd … H) -IHd -H // +lemma deg_inv_prec: ∀h,o,s,n,d. deg h o ((next h)^n s) (⫯d) → deg h o s (⫯(d+n)). +#h #o #s #n elim n -n normalize /3 width=1 by deg_inv_pred/ qed-. (* Basic properties *********************************************************) -lemma deg_iter: ∀h,g,k,d1,d2. deg h g k d1 → deg h g ((next h)^d2 k) (d1-d2). -#h #g #k #d1 #d2 @(nat_ind_plus … d2) -d2 [ iter_SO iter_SO -lapply (nexts_le h k d) #H +lemma nexts_lt: ∀h,s,n. s < (next h)^(⫯n) s. +#h #s #n normalize +lapply (nexts_le h s n) #H @(le_to_lt_to_lt … H) // qed. -axiom nexts_dec: ∀h,k1,k2. Decidable (∃d. (next h)^d k1 = k2). +axiom nexts_dec: ∀h,s1,s2. Decidable (∃n. (next h)^n s1 = s2). -axiom nexts_inj: ∀h,k,d1,d2. (next h)^d1 k = (next h)^d2 k → d1 = d2. +axiom nexts_inj: ∀h,s,n1,n2. (next h)^n1 s = (next h)^n2 s → n1 = n2. 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 b65bfda55..0d288fdc2 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 @@ -183,20 +183,14 @@ table { ] } ] - class "grass" +*) + class "green" [ { "static typing" * } { - [ { "local env. ref. for degree assignment" * } { - [ "lsubd ( ? ⊢ ? ⫃▪[?,?] ? )" "lsubd_da" + "lsubd_lsubd" * ] - } - ] - [ { "degree assignment" * } { - [ "da ( ⦃?,?⦄ ⊢ ? ▪[?,?] ? )" "da_lift" + "da_aaa" + "da_da" * ] - } - ] [ { "parameters" * } { [ "sh" "sd" * ] } ] +(* [ { "local env. ref. for atomic arity assignment" * } { [ "lsuba ( ? ⊢ ? ⫃⁝ ? )" "lsuba_aaa" + "lsuba_lsuba" * ] } @@ -205,74 +199,47 @@ table { [ "aaa ( ⦃?,?⦄ ⊢ ? ⁝ ? )" "aaa_lift" + "aaa_lifts" + "aaa_fqus" + "aaa_lleq" + "aaa_aaa" * ] } ] - [ { "restricted local env. ref." * } { - [ "lsubr ( ? ⫃ ? )" "lsubr_lsubr" * ] - } - ] - } - ] - - [ { "context-sensitive multiple rt-substitution" * } { - [ "cpys ( ⦃?,?⦄ ⊢ ? ▶*[?,?] ? )" "cpys_alt ( ⦃?,?⦄ ⊢ ? ▶▶*[?,?] ? )" "cpys_lift" + "cpys_cpys" * ] - } - ] - [ { "iterated structural successor for closures" * } { - [ "fqus ( ⦃?,?,?⦄ ⊐* ⦃?,?,?⦄ )" "fqus_alt" + "fqus_fqus" * ] - [ "fqup ( ⦃?,?,?⦄ ⊐+ ⦃?,?,?⦄ )" "fqup_fqup" * ] +*) + [ { "restricted ref. for local env." * } { + [ "lsubr ( ? ⫃ ? )" "lsubr_length" + "lsubr_drops" + "lsubr_lsubr" * ] } ] - [ { "pointwise union for local environments" * } { - [ "llor ( ? ⋓[?,?] ? ≡ ? )" "llor_alt" + "llor_drop" * ] + [ { "ranged equivalence for closures" * } { + [ "freq ( ⦃?,?,?⦄ ≡ ⦃?,?,?⦄ )" "freq_freq" * ] } ] - [ { "support for multiple relocation" * } { - [ "mr2 ( @⦃?,?⦄ ≡ ? )" "mr2_plus ( ? + ? )" "mr2_minus ( ? ▭ ? ≡ ? )" "mr2_mr2" * ] + [ { "context-sensitive free variables" * } { + [ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_weight" + "frees_lreq" + "frees_frees" * ] } ] - [ { "lazy pointwise extension of a relation" * } { - [ "llpx_sn" "llpx_sn_alt" + "llpx_sn_alt_rec" + "llpx_sn_tc" + "llpx_sn_lreq" + "llpx_sn_drop" + "llpx_sn_lpx_sn" + "llpx_sn_frees" + "llpx_sn_llor" * ] + } + ] + class "grass" + [ { "s-computation" * } { + [ { "" * } { + [ * ] } ] + } + ] + class "yellow" + [ { "s-transition" * } { [ { "structural successor for closures" * } { - [ "fquq ( ⦃?,?,?⦄ ⊐⸮ ⦃?,?,?⦄ )" "fquq_alt ( ⦃?,?,?⦄ ⊐⊐⸮ ⦃?,?,?⦄ )" * ] - [ "fqu ( ⦃?,?,?⦄ ⊐ ⦃?,?,?⦄ )" * ] - } - ] - [ { "global env. slicing" * } { - [ "gget ( ⬇[?] ? ≡ ? )" "gget_gget" * ] + [ "fquq ( ⦃?,?,?⦄ ⊐⸮ ⦃?,?,?⦄ )" "fquq_length" + "fquq_weight" * ] + [ "fqu ( ⦃?,?,?⦄ ⊐ ⦃?,?,?⦄ )" "fqu_length" + "fqu_weight" * ] } ] - [ { "context-sensitive ordinary rt-substitution" * } { - [ "cpy ( ⦃?,?⦄ ⊢ ? ▶[?,?] ? )" "cpy_lift" + "cpy_nlift" + "cpy_cpy" * ] - } - ] - [ { "local env. ref. for rt-substitution" * } { - [ "lsuby ( ? ⊆[?,?] ? )" "lsuby_lsuby" * ] - } - ] - [ { "pointwise extension of a relation" * } { - [ "lpx_sn" "lpx_sn_alt" + "lpx_sn_tc" + "lpx_sn_drop" + "lpx_sn_lpx_sn" * ] - } - ] - [ "lleq ( ? ≡[?,?] ? )" "lleq_alt" + "lleq_alt_rec" + "lleq_lreq" + "lleq_drop" + "lleq_fqus" + "lleq_llor" + "lleq_lleq" * ] -*) - class "yellow" + } + ] + class "orange" [ { "relocation" * } { - [ { "ranged equivalence for closures" * } { - [ "freq ( ⦃?,?,?⦄ ≡ ⦃?,?,?⦄ )" "freq_freq" * ] - } - ] - [ { "context-sensitive free variables" * } { - [ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_weight" + "frees_lreq" + "frees_frees" * ] - } - ] [ { "generic slicing for local environments" * } { [ "drops_vector ( ⬇*[?,?] ? ≡ ? )" * ] [ "drops ( ⬇*[?,?] ? ≡ ? )" "drops_lstar" + "drops_weight" + "drops_length" + "drops_ceq" + "drops_lexs" + "drops_lreq" + "drops_drops" * ] } ] [ { "generic relocation for terms" * } { - [ "lifts_vector ( ⬆*[?] ? ≡ ? )" "lifts_lift_vector" * ] + [ "lifts_vector ( ⬆*[?] ? ≡ ? )" "lifts_lifts_vector" * ] [ "lifts ( ⬆*[?] ? ≡ ? )" "lifts_simple" + "lifts_weight" + "lifts_lifts" * ] } ] @@ -286,16 +253,12 @@ table { ] } ] - class "orange" - [ { "" * } { - [ { "" * } { - [ * ] - } - ] - } - ] class "red" [ { "grammar" * } { + [ { "append for local environments" * } { + [ "append ( ? @@ ? )" "append_length" * ] + } + ] [ { "context-sensitive equivalences for terms" * } { [ "ceq" "ceq_ceq" * ] } @@ -310,7 +273,7 @@ table { ] [ { "internal syntax" * } { [ "genv" * ] - [ "lenv" "lenv_weight ( ♯{?} )" "lenv_length ( |?| )" "lenv_append ( ? @@ ? )" * ] + [ "lenv" "lenv_weight ( ♯{?} )" "lenv_length ( |?| )" * ] [ "term" "term_weight ( ♯{?} )" "term_simple ( 𝐒⦃?⦄ )" "term_vector ( Ⓐ?.? )" * ] [ "item" * ] } @@ -328,3 +291,47 @@ class "top" { * } class "capitalize italic" { 0 } class "italic" { 1 } +(* + [ { "local env. ref. for degree assignment" * } { + [ "lsubd ( ? ⊢ ? ⫃▪[?,?] ? )" "lsubd_da" + "lsubd_lsubd" * ] + } + ] + [ { "degree assignment" * } { + [ "da ( ⦃?,?⦄ ⊢ ? ▪[?,?] ? )" "da_lift" + "da_aaa" + "da_da" * ] + } + ] + [ { "context-sensitive multiple rt-substitution" * } { + [ "cpys ( ⦃?,?⦄ ⊢ ? ▶*[?,?] ? )" "cpys_alt ( ⦃?,?⦄ ⊢ ? ▶▶*[?,?] ? )" "cpys_lift" + "cpys_cpys" * ] + } + ] + [ { "iterated structural successor for closures" * } { + [ "fqus ( ⦃?,?,?⦄ ⊐* ⦃?,?,?⦄ )" "fqus_alt" + "fqus_fqus" * ] + [ "fqup ( ⦃?,?,?⦄ ⊐+ ⦃?,?,?⦄ )" "fqup_fqup" * ] + } + ] + [ { "pointwise union for local environments" * } { + [ "llor ( ? ⋓[?,?] ? ≡ ? )" "llor_alt" + "llor_drop" * ] + } + ] + [ { "lazy pointwise extension of a relation" * } { + [ "llpx_sn" "llpx_sn_alt" + "llpx_sn_alt_rec" + "llpx_sn_tc" + "llpx_sn_lreq" + "llpx_sn_drop" + "llpx_sn_lpx_sn" + "llpx_sn_frees" + "llpx_sn_llor" * ] + } + ] + [ { "global env. slicing" * } { + [ "gget ( ⬇[?] ? ≡ ? )" "gget_gget" * ] + } + ] + [ { "context-sensitive ordinary rt-substitution" * } { + [ "cpy ( ⦃?,?⦄ ⊢ ? ▶[?,?] ? )" "cpy_lift" + "cpy_nlift" + "cpy_cpy" * ] + } + ] + [ { "local env. ref. for rt-substitution" * } { + [ "lsuby ( ? ⊆[?,?] ? )" "lsuby_lsuby" * ] + } + ] + [ { "pointwise extension of a relation" * } { + [ "lpx_sn" "lpx_sn_alt" + "lpx_sn_tc" + "lpx_sn_drop" + "lpx_sn_lpx_sn" * ] + } + ] + [ "lleq ( ? ≡[?,?] ? )" "lleq_alt" + "lleq_alt_rec" + "lleq_lreq" + "lleq_drop" + "lleq_fqus" + "lleq_llor" + "lleq_lleq" * ] +*) diff --git a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma index 3cb9e9f30..f86f7a3bd 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -172,6 +172,10 @@ qed-. lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥. /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-. +lemma succ_inv_refl_sn: ∀x. ⫯x = x → ⊥. +#x #H @(lt_le_false x (⫯x)) // +qed-. + lemma lt_inv_O1: ∀n. 0 < n → ∃m. ⫯m = n. * /2 width=2 by ex_intro/ #H cases (lt_le_false … H) -H // diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma index 80401a3d1..222bc07f7 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma @@ -14,7 +14,6 @@ include "ground_2/notation/relations/rafter_3.ma". include "ground_2/relocation/rtmap_istot.ma". -include "ground_2/relocation/rtmap_uni.ma". (* RELOCATION MAP ***********************************************************) diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma index 06223bc53..a1f191a25 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "ground_2/notation/relations/rat_3.ma". -include "ground_2/relocation/rtmap_id.ma". +include "ground_2/relocation/rtmap_uni.ma". (* RELOCATION MAP ***********************************************************) @@ -320,3 +320,9 @@ lemma id_inv_at: ∀f. (∀i. @⦃i, f⦄ ≡ i) → 𝐈𝐝 ≗ f. lemma id_at: ∀i. @⦃i, 𝐈𝐝⦄ ≡ i. /2 width=1 by isid_inv_at/ qed. + +(* Properties with uniform relocations **************************************) + +lemma at_uni: ∀n,i. @⦃i,𝐔❴n❵⦄ ≡ n+i. +#n elim n -n /2 width=5 by at_next/ +qed.