From: Ferruccio Guidi Date: Sat, 21 Apr 2012 13:20:21 +0000 (+0000) Subject: - lambda_delta: static type assignment is defined X-Git-Tag: make_still_working~1807 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=1efc4c2c7be1e4aff0ccccabf905d45795b3865f;p=helm.git - lambda_delta: static type assignment is defined - predefined_virtuals: some additions - nat: we added a lemma --- diff --git a/matita/matita/contribs/lambda_delta/basic_2/basic_1.txt b/matita/matita/contribs/lambda_delta/basic_2/basic_1.txt index 92bd9860d..1b6aaa13f 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/basic_1.txt +++ b/matita/matita/contribs/lambda_delta/basic_2/basic_1.txt @@ -204,13 +204,6 @@ pr3/pr1 pr3_pr1 pr3/props pr3_eta pr3/subst1 pr3_gen_cabbr sn3/props sns3_lifts -sty0/fwd sty0_gen_sort -sty0/fwd sty0_gen_lref -sty0/fwd sty0_gen_bind -sty0/fwd sty0_gen_appl -sty0/fwd sty0_gen_cast -sty0/props sty0_lift -sty0/props sty0_correct sty1/cnt sty1_cnt sty1/props sty1_trans sty1/props sty1_bind diff --git a/matita/matita/contribs/lambda_delta/basic_2/dynamic/nta.ma b/matita/matita/contribs/lambda_delta/basic_2/dynamic/nta.ma new file mode 100644 index 000000000..99592cdd9 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/dynamic/nta.ma @@ -0,0 +1,51 @@ +(**************************************************************************) +(* ___ *) +(* ||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/sh.ma". +include "basic_2/equivalence/cpcs.ma". + +(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************) + +inductive nta (h:sh): lenv → relation term ≝ +| nta_sort: ∀L,k. nta h L (⋆k) (⋆(next h k)) +| nta_ldef: ∀L,K,V,W,U,i. ⇩[0, i] L ≡ K. ⓓV → nta h K V W → + ⇧[0, i + 1] W ≡ U → nta h L (#i) U +| nta_ldec: ∀L,K,W,V,U,i. ⇩[0, i] L ≡ K. ⓛW → nta h K W V → + ⇧[0, i + 1] W ≡ U → nta h L (#i) U +| nta_bind: ∀I,L,V,W,T,U. nta h L V W → nta h (L. ⓑ{I} V) T U → + nta h L (ⓑ{I}V.T) (ⓑ{I}V.U) +| nta_appl: ∀L,V,W,U,T1,T2. nta h L V W → nta h L W U → nta h (L.ⓛW) T1 T2 → + nta h L (ⓐV.ⓛW.T1) (ⓐV.ⓛW.T2) +| nta_pure: ∀L,V,W,T,U. nta h L T U → nta h L (ⓐV.U) W → + nta h L (ⓐV.T) (ⓐV.U) +| nta_cast: ∀L,T,U. nta h L T U → nta h L (ⓣU. T) U +| nta_conv: ∀L,T,U1,U2,V2. nta h L T U1 → L ⊢ U1 ⬌* U2 → nta h L U2 V2 → + nta h L T U2 +. + +interpretation "native type assignment (term)" + 'NativeType h L T U = (nta h L T U). + +(* Basic properties *********************************************************) + +(* Basic_1: was: ty3_cast *) +lemma nta_cast_old: ∀h,L,W,T,U. + ⦃h, L⦄ ⊢ T : U → ⦃h, L⦄ ⊢ U : W → ⦃h, L⦄ ⊢ ⓣU.T : ⓣW.U. +/4 width=3/ qed. + +(* Basic_1: was: ty3_typecheck *) +lemma nta_typecheck: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∃T0. ⦃h, L⦄ ⊢ ⓣU.T : T0. +/3 width=2/ qed. + +(* Basic_1: removed theorems 1: ty3_getl_subst0 *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/dynamic/nta_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/dynamic/nta_lift.ma new file mode 100644 index 000000000..a0ece8cbf --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/dynamic/nta_lift.ma @@ -0,0 +1,201 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/equivalence/cpcs_cpcs.ma". +include "basic_2/dynamic/nta.ma". + +(* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************) + +(* Advanced inversion lemmas ************************************************) + +fact nta_inv_sort1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∀k0. T = ⋆k0 → + L ⊢ ⋆(next h k0) ⬌* U. +#h #L #T #U #H elim H -L -T -U +[ #L #k #k0 #H destruct // +| #L #K #V #W #U #i #_ #_ #_ #_ #k0 #H destruct +| #L #K #W #V #U #i #_ #_ #_ #_ #k0 #H destruct +| #I #L #V #W #T #U #_ #_ #_ #_ #k0 #H destruct +| #L #V #W #U #T1 #T2 #_ #_ #_ #_ #_ #_ #k0 #H destruct +| #L #V #W #T #U #_ #_ #_ #_ #k0 #H destruct +| #L #T #U #_ #_ #k0 #H destruct +| #L #T #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #k0 #H destruct + lapply (IHTU1 ??) -IHTU1 [ // | skip ] #Hk0 + lapply (cpcs_trans … Hk0 … HU12) -U1 // +] +qed. + +(* Basic_1: was: ty3_gen_sort *) +lemma nta_inv_sort1: ∀h,L,U,k. ⦃h, L⦄ ⊢ ⋆k : U → L ⊢ ⋆(next h k) ⬌* U. +/2 width=3/ qed-. + +fact nta_inv_lref1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∀j. T = #j → + (∃∃K,V,W,U0. ⇩[0, j] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V : W & + ⇧[0, j + 1] W ≡ U0 & L ⊢ U0 ⬌* U + ) ∨ + (∃∃K,W,V,U0. ⇩[0, j] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W : V & + ⇧[0, j + 1] W ≡ U0 & L ⊢ U0 ⬌* U + ). +#h #L #T #U #H elim H -L -T -U +[ #L #k #j #H destruct +| #L #K #V #W #U #i #HLK #HVW #HWU #_ #j #H destruct /3 width=8/ +| #L #K #W #V #U #i #HLK #HWV #HWU #_ #j #H destruct /3 width=8/ +| #I #L #V #W #T #U #_ #_ #_ #_ #j #H destruct +| #L #V #W #U #T1 #T2 #_ #_ #_ #_ #_ #_ #j #H destruct +| #L #V #W #T #U #_ #_ #_ #_ #j #H destruct +| #L #T #U #_ #_ #j #H destruct +| #L #T #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #j #H destruct + elim (IHTU1 ??) -IHTU1 [4: // |2: skip ] * #K #V #W #U0 #HLK #HVW #HWU0 #HU01 + lapply (cpcs_trans … HU01 … HU12) -U1 /3 width=8/ +] +qed. + +(* Basic_1: was ty3_gen_lref *) +lemma nta_inv_lref1: ∀h,L,U,i. ⦃h, L⦄ ⊢ #i : U → + (∃∃K,V,W,U0. ⇩[0, i] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V : W & + ⇧[0, i + 1] W ≡ U0 & L ⊢ U0 ⬌* U + ) ∨ + (∃∃K,W,V,U0. ⇩[0, i] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W : V & + ⇧[0, i + 1] W ≡ U0 & L ⊢ U0 ⬌* U + ). +/2 width=3/ qed-. + +fact nta_inv_bind1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∀J,X,Y. T = ⓑ{J}Y.X → + ∃∃Z1,Z2. ⦃h, L⦄ ⊢ Y : Z1 & ⦃h, L.ⓑ{J}Y⦄ ⊢ X : Z2 & + L ⊢ ⓑ{J}Y.Z2 ⬌* U. +#h #L #T #U #H elim H -L -T -U +[ #L #k #J #X #Y #H destruct +| #L #K #V #W #U #i #_ #_ #_ #_ #J #X #Y #H destruct +| #L #K #W #V #U #i #_ #_ #_ #_ #J #X #Y #H destruct +| #I #L #V #W #T #U #HVW #HTU #_ #_ #J #X #Y #H destruct /2 width=3/ +| #L #V #W #U #T1 #T2 #_ #_ #_ #_ #_ #_ #J #X #Y #H destruct +| #L #V #W #T #U #_ #_ #_ #_ #J #X #Y #H destruct +| #L #T #U #_ #_ #J #X #Y #H destruct +| #L #T #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #J #X #Y #H destruct + elim (IHTU1 ????) -IHTU1 [5: // |2,3,4: skip ] #Z1 #Z2 #HZ1 #HZ2 #HU1 + lapply (cpcs_trans … HU1 … HU12) -U1 /2 width=3/ +] +qed. + +(* Basic_1: was: ty3_gen_bind *) +lemma nta_inv_bind1: ∀h,J,L,Y,X,U. ⦃h, L⦄ ⊢ ⓑ{J}Y.X : U → + ∃∃Z1,Z2. ⦃h, L⦄ ⊢ Y : Z1 & ⦃h, L.ⓑ{J}Y⦄ ⊢ X : Z2 & + L ⊢ ⓑ{J}Y.Z2 ⬌* U. +/2 width=3/ qed-. + +fact nta_inv_cast1_aux: ∀h,L,T,U. ⦃h, L⦄ ⊢ T : U → ∀X,Y. T = ⓣY.X → + ⦃h, L⦄ ⊢ X : Y ∧ L ⊢ Y ⬌* U. +#h #L #T #U #H elim H -L -T -U +[ #L #k #X #Y #H destruct +| #L #K #V #W #U #i #_ #_ #_ #_ #X #Y #H destruct +| #L #K #W #V #U #i #_ #_ #_ #_ #X #Y #H destruct +| #I #L #V #W #T #U #_ #_ #_ #_ #X #Y #H destruct +| #L #V #W #U #T1 #T2 #_ #_ #_ #_ #_ #_ #X #Y #H destruct +| #L #V #W #T #U #_ #_ #_ #_ #X #Y #H destruct +| #L #T #U #HTU #_ #X #Y #H destruct /2 width=1/ +| #L #T #U1 #U2 #V2 #_ #HU12 #_ #IHTU1 #_ #X #Y #H destruct + elim (IHTU1 ???) -IHTU1 [4: // |2,3: skip ] #HXY #HU1 + lapply (cpcs_trans … HU1 … HU12) -U1 /2 width=1/ +] +qed. + +(* Basic_1: was: ty3_gen_cast *) +lemma nta_inv_cast1: ∀h,L,X,Y,U. ⦃h, L⦄ ⊢ ⓣY.X : U → ⦃h, L⦄ ⊢ X : Y ∧ L ⊢ Y ⬌* U. +/2 width=3/ qed-. + +(* Properties on relocation *************************************************) + +(* Basic_1: was: ty3_lift *) +lemma nta_lift: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → + ∀T2. ⇧[d, e] T1 ≡ T2 → ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 : U2. +#h #L1 #T1 #U1 #H elim H -L1 -T1 -U1 +[ #L1 #k #L2 #d #e #HL21 #X1 #H1 #X2 #H2 + >(lift_inv_sort1 … H1) -X1 + >(lift_inv_sort1 … H2) -X2 // +| #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // #W2 #HW12 #HWU2 + elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H + elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct + /3 width=8/ + | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2 + lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ + ] +| #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // (lift_inv_sort1 … H1) -X1 - >(lift_inv_sort1 … H2) -X2 // -| #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2 - elim (lift_inv_lref1 … H) * #Hid #H destruct - [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // #W2 #HW12 #HWU2 - elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H - elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct - /3 width=8/ - | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2 - lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ - ] -| #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2 - elim (lift_inv_lref1 … H) * #Hid #H destruct - [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // (lift_inv_sort1 … H1) -X1 + >(lift_inv_sort1 … H2) -X2 // +| #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // #W2 #HW12 #HWU2 + elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H + elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct + /3 width=8/ + | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2 + lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ + ] +| #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // (lift_inv_sort2 … H) -X /2 width=3/ +| #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #d #e #HL21 #X #H + elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ] + [ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // #K1 #V1 #HLK1 #HK21 #HV12 + elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HVW1 #HW12 + elim (lift_trans_le … HW12 … HW2 ?) -W2 // >minus_plus minus_minus_m_m /2 width=1/ /3 width=6/ + | minus_plus minus_minus_m_m /2 width=1/ /3 width=6/ + | (sta_inv_sort1 … H) -X // +| #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H + elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2 + lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct + lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct + >(lift_mono … HWU1 … HW0U2) -W0 -U1 // +| #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H + elim (sta_inv_lref1 … H) -H * #K0 #W0 #V0 #HLK0 #HWV0 #HV0U2 + lapply (ldrop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct + lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct + >(lift_mono … HWU1 … HV0U2) -W -U1 // +| #I #L #V #T #U1 #_ #IHTU1 #X #H + elim (sta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1/ +| #L #V #T #U1 #_ #IHTU1 #X #H + elim (sta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1/ +| #L #T #U1 #_ #_ #U2 #H + elim (sta_inv_cast1 … H) -H // +] +qed-. diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index 31d5ced7f..d806697b8 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -666,6 +666,9 @@ lemma minus_le_minus_minus_comm: ∀b,c,a. c ≤ b → a - (b - c) = a + c - b. #b #c #a #H >(plus_minus_m_m b c) in ⊢ (? ? ?%); // qed. +lemma minus_minus_m_m: ∀m,n. n ≤ m → m - (m - n) = n. +/2 width=1/ qed. + (* Stilll more atomic conclusion ********************************************) (* le *) diff --git a/matita/matita/predefined_virtuals.ml b/matita/matita/predefined_virtuals.ml index b276fe4bb..cbdafb2e0 100644 --- a/matita/matita/predefined_virtuals.ml +++ b/matita/matita/predefined_virtuals.ml @@ -1503,6 +1503,7 @@ let load_predefined_virtuals () = ;; let predefined_classes = [ + ["."; "•"; "◦"; ]; ["#"; "⌘"; ]; ["-"; "÷"; "⊢"; ]; ["="; "≃"; "≈"; "≝"; "≡"; "≅"; "≐"; "≑"; ];