From: Ferruccio Guidi Date: Sat, 10 Mar 2012 14:46:11 +0000 (+0000) Subject: We are decapitalizing the contributions' names ... X-Git-Tag: make_still_working~1864 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=eb918fc784eacd2094e3986ba321ef47690d9983;hp=011cf6478141e69822a5b40933f2444d0522532f;p=helm.git We are decapitalizing the contributions' names ... --- diff --git a/matita/matita/contribs/lambda_delta/Apps_2/functional/lift.ma b/matita/matita/contribs/lambda_delta/Apps_2/functional/lift.ma deleted file mode 100644 index aefa1576d..000000000 --- a/matita/matita/contribs/lambda_delta/Apps_2/functional/lift.ma +++ /dev/null @@ -1,69 +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/tri.ma". -include "Basic_2/substitution/lift.ma". -include "Apps_2/functional/notation.ma". - -(* RELOCATION ***************************************************************) - -let rec flift d e U on U ≝ match U with -[ TAtom I ⇒ match I with - [ Sort _ ⇒ U - | LRef i ⇒ #(tri … i d i (i + e) (i + e)) - | GRef _ ⇒ U - ] -| TPair I V T ⇒ match I with - [ Bind2 I ⇒ ⓑ{I} (flift d e V). (flift (d+1) e T) - | Flat2 I ⇒ ⓕ{I} (flift d e V). (flift d e T) - ] -]. - -interpretation "functional relocation" 'Lift d e T = (flift d e T). - -(* Main properties **********************************************************) - -theorem flift_lift: ∀T,d,e. ⇧[d, e] T ≡ ↑[d, e] T. -#T elim T -T -[ * #i #d #e // - elim (lt_or_eq_or_gt i d) #Hid normalize - [ >(tri_lt ?????? Hid) /2 width=1/ - | /2 width=1/ - | >(tri_gt ?????? Hid) /3 width=2/ - ] -| * /2/ -] -qed. - -(* Main inversion properties ************************************************) - -theorem flift_inv_lift: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → ↑[d, e] T1 = T2. -#d #e #T1 #T2 #H elim H -d -e -T1 -T2 normalize // -[ #i #d #e #Hid >(tri_lt ?????? Hid) // -| #i #d #e #Hid - elim (le_to_or_lt_eq … Hid) -Hid #Hid - [ >(tri_gt ?????? Hid) // - | destruct // - ] -] -qed-. - -(* Derived properties *******************************************************) - -lemma flift_join: ∀e1,e2,T. ⇧[e1, e2] ↑[0, e1] T ≡ ↑[0, e1 + e2] T. -#e1 #e2 #T -lapply (flift_lift T 0 (e1+e2)) #H -elim (lift_split … H e1 e1 ? ? ?) -H // #U #H ->(flift_inv_lift … H) -H // -qed. diff --git a/matita/matita/contribs/lambda_delta/Apps_2/functional/notation.ma b/matita/matita/contribs/lambda_delta/Apps_2/functional/notation.ma deleted file mode 100644 index 48df845c1..000000000 --- a/matita/matita/contribs/lambda_delta/Apps_2/functional/notation.ma +++ /dev/null @@ -1,27 +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 "functional" COMPONENT ********************************) - -notation "hvbox( ↑ [ d , break e ] break T )" - non associative with precedence 55 - for @{ 'Lift $d $e $T }. - -notation "hvbox( [ d ← break V ] break T )" - non associative with precedence 55 - for @{ 'Subst $V $d $T }. - -notation "hvbox( T1 ⇨ break T2 )" - non associative with precedence 45 - for @{ 'SRed $T1 $T2 }. diff --git a/matita/matita/contribs/lambda_delta/Apps_2/functional/rtm.ma b/matita/matita/contribs/lambda_delta/Apps_2/functional/rtm.ma deleted file mode 100644 index 845e8a04d..000000000 --- a/matita/matita/contribs/lambda_delta/Apps_2/functional/rtm.ma +++ /dev/null @@ -1,85 +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/grammar/term_vector.ma". -include "Basic_2/grammar/genv.ma". - -(* REDUCTION AND TYPE MACHINE ***********************************************) - -(* machine local environment *) -inductive xenv: Type[0] ≝ -| XAtom: xenv (* empty *) -| XQuad: xenv → bind2 → nat → xenv → term → xenv (* entry *) -. - -interpretation "atom (ext. local environment)" - 'Star = XAtom. - -interpretation "environment construction (quad)" - 'DxItem4 L I u K V = (XQuad L I u K V). - -(* machine stack *) -definition stack: Type[0] ≝ list2 xenv term. - -(* machine status *) -record rtm: Type[0] ≝ -{ rg: genv; (* global environment *) - ru: nat; (* current de Bruijn's level *) - re: xenv; (* extended local environment *) - rs: stack; (* application stack *) - rt: term (* code *) -}. - -(* initial state *) -definition rtm_i: genv → term → rtm ≝ - λG,T. mk_rtm G 0 (⋆) (⟠) T. - -(* update code *) -definition rtm_t: rtm → term → rtm ≝ - λM,T. match M with - [ mk_rtm G u E _ _ ⇒ mk_rtm G u E (⟠) T - ]. - -(* update closure *) -definition rtm_u: rtm → xenv → term → rtm ≝ - λM,E,T. match M with - [ mk_rtm G u _ _ _ ⇒ mk_rtm G u E (⟠) T - ]. - -(* get global environment *) -definition rtm_g: rtm → genv ≝ - λM. match M with - [ mk_rtm G _ _ _ _ ⇒ G - ]. - -(* get local reference level *) -definition rtm_l: rtm → nat ≝ - λM. match M with - [ mk_rtm _ u E _ _ ⇒ match E with - [ XAtom ⇒ u - | XQuad _ _ u _ _ ⇒ u - ] - ]. - -(* get stack *) -definition rtm_s: rtm → stack ≝ - λM. match M with - [ mk_rtm _ _ _ S _ ⇒ S - ]. - -(* get code *) -definition rtm_c: rtm → term ≝ - λM. match M with - [ mk_rtm _ _ _ _ T ⇒ T - ]. diff --git a/matita/matita/contribs/lambda_delta/Apps_2/functional/rtm_step.ma b/matita/matita/contribs/lambda_delta/Apps_2/functional/rtm_step.ma deleted file mode 100644 index 5b2a4eb2c..000000000 --- a/matita/matita/contribs/lambda_delta/Apps_2/functional/rtm_step.ma +++ /dev/null @@ -1,57 +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 "Apps_2/functional/rtm.ma". - -(* REDUCTION AND TYPE MACHINE ***********************************************) - -(* transitions *) -inductive rtm_step: relation rtm ≝ -| rtm_ldrop : ∀G,u,E,I,t,F,V,S,i. - rtm_step (mk_rtm G u (E. ④{I} {t, F, V}) S (#(i + 1))) - (mk_rtm G u E S (#i)) -| rtm_ldelta: ∀G,u,E,t,F,V,S. - rtm_step (mk_rtm G u (E. ④{Abbr} {t, F, V}) S (#0)) - (mk_rtm G u F S V) -| rtm_ltype : ∀G,u,E,t,F,V,S. - rtm_step (mk_rtm G u (E. ④{Abst} {t, F, V}) S (#0)) - (mk_rtm G u F S V) -| rtm_gdrop : ∀G,I,V,u,E,S,p. p < |G| → - rtm_step (mk_rtm (G. ⓑ{I} V) u E S (§p)) - (mk_rtm G u E S (§p)) -| rtm_gdelta: ∀G,V,u,E,S,p. p = |G| → - rtm_step (mk_rtm (G. ⓓV) u E S (§p)) - (mk_rtm G u E S V) -| rtm_gtype : ∀G,V,u,E,S,p. p = |G| → - rtm_step (mk_rtm (G. ⓛV) u E S (§p)) - (mk_rtm G u E S V) -| rtm_tau : ∀G,u,E,S,W,T. - rtm_step (mk_rtm G u E S (ⓣW. T)) - (mk_rtm G u E S T) -| rtm_appl : ∀G,u,E,S,V,T. - rtm_step (mk_rtm G u E S (ⓐV. T)) - (mk_rtm G u E ({E, V} :: S) T) -| rtm_beta : ∀G,u,E,F,V,S,W,T. - rtm_step (mk_rtm G u E ({F, V} :: S) (ⓛW. T)) - (mk_rtm G u (E. ④{Abbr} {u, F, V}) S T) -| rtm_push : ∀G,u,E,W,T. - rtm_step (mk_rtm G u E ⟠ (ⓛW. T)) - (mk_rtm G (u + 1) (E. ④{Abst} {u, E, W}) ⟠ T) -| rtm_theta : ∀G,u,E,S,V,T. - rtm_step (mk_rtm G u E S (ⓓV. T)) - (mk_rtm G u (E. ④{Abbr} {u, E, V}) S T) -. - -interpretation "sequential reduction (RTM)" - 'SRed O1 O2 = (rtm_step O1 O2). diff --git a/matita/matita/contribs/lambda_delta/Apps_2/functional/subst.ma b/matita/matita/contribs/lambda_delta/Apps_2/functional/subst.ma deleted file mode 100644 index 6aa75625d..000000000 --- a/matita/matita/contribs/lambda_delta/Apps_2/functional/subst.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/unfold/delift_lift.ma". -include "Apps_2/functional/lift.ma". - -(* CORE SUBSTITUTION ********************************************************) - -let rec fsubst W d U on U ≝ match U with -[ TAtom I ⇒ match I with - [ Sort _ ⇒ U - | LRef i ⇒ tri … i d (#i) (↑[0, i] W) (#(i-1)) - | GRef _ ⇒ U - ] -| TPair I V T ⇒ match I with - [ Bind2 I ⇒ ⓑ{I} (fsubst W d V). (fsubst W (d+1) T) - | Flat2 I ⇒ ⓕ{I} (fsubst W d V). (fsubst W d T) - ] -]. - -interpretation "functional core substitution" 'Subst V d T = (fsubst V d T). - -(* Main properties **********************************************************) - -theorem fsubst_delift: ∀K,V,T,L,d. - ⇩[0, d] L ≡ K. ⓓV → L ⊢ T [d, 1] ≡ [d ← V] T. -#K #V #T elim T -T -[ * #i #L #d #HLK normalize in ⊢ (? ? ? ? ? %); /2 width=3/ - elim (lt_or_eq_or_gt i d) #Hid - [ -HLK >(tri_lt ?????? Hid) /3 width=3/ - | destruct >tri_eq /4 width=4 by tpss_strap, tps_subst, le_n, ex2_1_intro/ (**) (* too slow without trace *) - | -HLK >(tri_gt ?????? Hid) /3 width=3/ - ] -| * /3 width=1/ /4 width=1/ -] -qed. - -(* Main inversion properties ************************************************) - -theorem fsubst_inv_delift: ∀K,V,T1,L,T2,d. ⇩[0, d] L ≡ K. ⓓV → - L ⊢ T1 [d, 1] ≡ T2 → [d ← V] T1 = T2. -#K #V #T1 elim T1 -T1 -[ * #i #L #T2 #d #HLK #H - [ -HLK >(delift_fwd_sort1 … H) -H // - | elim (lt_or_eq_or_gt i d) #Hid normalize - [ -HLK >(delift_fwd_lref1_lt … H) -H // /2 width=1/ - | destruct - elim (delift_fwd_lref1_be … H ? ?) -H // #K0 #V0 #V2 #HLK0 - lapply (ldrop_mono … HLK0 … HLK) -HLK0 -HLK #H >minus_plus (delift_inv_refl_O2 … HV2) -V >(flift_inv_lift … HVT2) -V2 // - | -HLK >(delift_fwd_lref1_ge … H) -H // /2 width=1/ - ] - | -HLK >(delift_fwd_gref1 … H) -H // - ] -| * #I #V1 #T1 #IHV1 #IHT1 #L #X #d #HLK #H - [ elim (delift_fwd_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - <(IHV1 … HV12) -IHV1 -HV12 // <(IHT1 … HT12) -IHT1 -HT12 // /2 width=1/ - | elim (delift_fwd_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - <(IHV1 … HV12) -IHV1 -HV12 // <(IHT1 … HT12) -IHT1 -HT12 // - ] -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/Basic_1.txt b/matita/matita/contribs/lambda_delta/Basic_2/Basic_1.txt deleted file mode 100644 index bc4582ede..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/Basic_1.txt +++ /dev/null @@ -1,388 +0,0 @@ -# waiting #################################################################### - -aplus/props aplus_reg_r -aplus/props aplus_assoc -aplus/props aplus_asucc -aplus/props aplus_sort_O_S_simpl -aplus/props aplus_sort_S_S_simpl -aplus/props aplus_asort_O_simpl -aplus/props aplus_asort_le_simpl -aplus/props aplus_asort_simpl -aplus/props aplus_ahead_simpl -aplus/props aplus_asucc_false -aplus/props aplus_inj -aprem/fwd aprem_gen_sort -aprem/fwd aprem_gen_head_O -aprem/fwd aprem_gen_head_S -aprem/props aprem_repl -aprem/props aprem_asucc -arity/aprem arity_aprem -arity/cimp arity_cimp_conf -arity/fwd arity_gen_sort -arity/fwd arity_gen_lref -arity/fwd arity_gen_bind -arity/fwd arity_gen_abst -arity/fwd arity_gen_appl -arity/fwd arity_gen_cast -arity/fwd arity_gen_appls -arity/fwd arity_gen_lift -arity/lift1 arity_lift1 -arity/pr3 arity_sred_wcpr0_pr0 -arity/pr3 arity_sred_wcpr0_pr1 -arity/pr3 arity_sred_pr2 -arity/pr3 arity_sred_pr3 -arity/props node_inh -arity/props arity_lift -arity/props arity_mono -arity/props arity_repellent -arity/props arity_appls_cast -arity/props arity_appls_abbr -arity/props arity_appls_bind -arity/subst0 arity_gen_cvoid_subst0 -arity/subst0 arity_gen_cvoid -arity/subst0 arity_fsubst0 -arity/subst0 arity_subst0 -asucc/fwd asucc_gen_sort -asucc/fwd asucc_gen_head -cnt/props cnt_lift -C/props clt_wf__q_ind -C/props clt_wf_ind -C/props chead_ctail -C/props clt_thead (ctail) -C/props c_tail_ind -csuba/arity csuba_arity -csuba/arity csuba_arity_rev -csuba/arity arity_appls_appl -csuba/clear csuba_clear_conf -csuba/clear csuba_clear_trans -csuba/drop csuba_drop_abbr -csuba/drop csuba_drop_abst -csuba/drop csuba_drop_abst_rev -csuba/drop csuba_drop_abbr_rev -csuba/fwd csuba_gen_abbr -csuba/fwd csuba_gen_void -csuba/fwd csuba_gen_abst -csuba/fwd csuba_gen_flat -csuba/fwd csuba_gen_bind -csuba/fwd csuba_gen_abst_rev -csuba/fwd csuba_gen_void_rev -csuba/fwd csuba_gen_abbr_rev -csuba/fwd csuba_gen_flat_rev -csuba/fwd csuba_gen_bind_rev -csuba/getl csuba_getl_abbr -csuba/getl csuba_getl_abst -csuba/getl csuba_getl_abst_rev -csuba/getl csuba_getl_abbr_rev -csuba/props csuba_refl -csubc/arity csubc_arity_conf -csubc/arity csubc_arity_trans -csubc/csuba csubc_csuba -csubc/drop1 drop1_csubc_trans -csubc/drop drop_csubc_trans -csubt/clear csubt_clear_conf -csubt/csuba csubt_csuba -csubt/drop csubt_drop_flat -csubt/drop csubt_drop_abbr -csubt/drop csubt_drop_abst -csubt/fwd csubt_gen_abbr -csubt/fwd csubt_gen_abst -csubt/fwd csubt_gen_flat -csubt/fwd csubt_gen_bind -csubt/getl csubt_getl_abbr -csubt/getl csubt_getl_abst -csubt/pc3 csubt_pr2 -csubt/pc3 csubt_pc3 -csubt/props csubt_refl -csubt/ty3 csubt_ty3 -csubt/ty3 csubt_ty3_ld -csubv/clear csubv_clear_conf -csubv/clear csubv_clear_conf_void -csubv/drop csubv_drop_conf -csubv/getl csubv_getl_conf -csubv/getl csubv_getl_conf_void -csubv/props csubv_bind_same -csubv/props csubv_refl -drop1/props drop1_cons_tail -drop/props drop_ctail -ex0/props aplus_gz_le -ex0/props aplus_gz_ge -ex0/props next_plus_gz -ex0/props leqz_leq -ex0/props leq_leqz -ex1/props ex1__leq_sort_SS -ex1/props ex1_arity -ex1/props ex1_ty3 -ex2/props ex2_nf2 -ex2/props ex2_arity -fsubst0/fwd fsubst0_gen_base -leq/asucc asucc_repl -leq/asucc asucc_inj -leq/asucc leq_asucc -leq/asucc leq_ahead_asucc_false -leq/asucc leq_asucc_false -leq/fwd leq_gen_sort1 -leq/fwd leq_gen_head1 -leq/fwd leq_gen_sort2 -leq/fwd leq_gen_head2 -leq/props ahead_inj_snd -leq/props leq_refl -leq/props leq_eq -leq/props leq_sym -leq/props leq_trans -leq/props leq_ahead_false_1 -leq/props leq_ahead_false_2 -lift1/fwd lift1_cons_tail - -# check ############################################################# - -lift1/fwd lifts1_flat -lift1/fwd lifts1_nil -lift1/fwd lifts1_cons -lift1/props lift1_free -lift/props thead_x_lift_y_y -lift/props lifts_tapp - -# waiting #################################################################### - -lift/props lifts_inj -llt/props lweight_repl -llt/props llt_repl -llt/props llt_trans -llt/props llt_head_sx -llt/props llt_head_dx -llt/props llt_wf__q_ind -llt/props llt_wf_ind -next_plus/props next_plus_assoc -next_plus/props next_plus_next -next_plus/props next_plus_lt -nf2/arity arity_nf2_inv_all -nf2/dec nf2_dec - -nf2/fwd nf2_gen_lref -nf2/fwd nf2_gen_abst -nf2/fwd nf2_gen_cast -nf2/fwd nf2_gen_beta -nf2/fwd nf2_gen_flat -nf2/fwd nf2_gen__nf2_gen_aux -nf2/fwd nf2_gen_abbr - -nf2/fwd nf2_gen_void -nf2/iso nf2_iso_appls_lref -nf2/pr3 nf2_pr3_unfold -nf2/pr3 nf2_pr3_confluence - -nf2/props nfs2_tapp -nf2/props nf2_appls_lref - -pc1/props pc1_pr0_r -pc1/props pc1_pr0_x -pc1/props pc1_refl -pc1/props pc1_pr0_u -pc1/props pc1_s -pc1/props pc1_head_1 -pc1/props pc1_head_2 -pc1/props pc1_t -pc1/props pc1_pr0_u2 -pc1/props pc1_head - -pc3/dec pc3_dec -pc3/dec pc3_abst_dec -pc3/fsubst0 pc3_pr2_fsubst0 -pc3/fsubst0 pc3_pr2_fsubst0_back -pc3/fsubst0 pc3_fsubst0 -pc3/fwd pc3_gen_sort -pc3/fwd pc3_gen_abst -pc3/fwd pc3_gen_abst_shift -pc3/fwd pc3_gen_lift -pc3/fwd pc3_gen_not_abst -pc3/fwd pc3_gen_lift_abst -pc3/fwd pc3_gen_sort_abst -pc3/left pc3_ind_left__pc3_left_pr3 -pc3/left pc3_ind_left__pc3_left_trans -pc3/left pc3_ind_left__pc3_left_sym -pc3/left pc3_ind_left__pc3_left_pc3 -pc3/left pc3_ind_left__pc3_pc3_left -pc3/left pc3_ind_left -pc3/nf2 pc3_nf2 -pc3/nf2 pc3_nf2_unfold -pc3/pc1 pc3_pc1 -pc3/props pc3_pr2_r -pc3/props pc3_pr2_x -pc3/props pc3_pr3_r -pc3/props pc3_pr3_x -pc3/props pc3_pr3_t -pc3/props pc3_s -pc3/props pc3_thin_dx -pc3/props pc3_head_1 -pc3/props pc3_head_2 -pc3/props pc3_pr2_u -pc3/props pc3_t -pc3/props pc3_pr2_u2 -pc3/props pc3_pr3_conf -pc3/props pc3_head_12 -pc3/props pc3_head_21 -pc3/props pc3_pr0_pr2_t -pc3/props pc3_pr2_pr2_t -pc3/props pc3_pr2_pr3_t -pc3/props pc3_pr3_pc3_t -pc3/props pc3_lift -pc3/props pc3_eta -pc3/subst1 pc3_gen_cabbr -pc3/wcpr0 pc3_wcpr0__pc3_wcpr0_t_aux -pc3/wcpr0 pc3_wcpr0_t -pc3/wcpr0 pc3_wcpr0 -pr0/fwd pr0_gen_void - -# check ############################################################# - -pr0/dec nf0_dec -pr0/subst1 pr0_subst1_back -pr0/subst1 pr0_subst1_fwd -pr1/pr1 pr1_strip -pr1/pr1 pr1_confluence -pr1/props pr1_pr0 -pr1/props pr1_t -pr1/props pr1_head_1 -pr1/props pr1_head_2 -pr1/props pr1_comp -pr1/props pr1_eta -pr2/clen pr2_gen_ctail -pr2/fwd pr2_gen_void -pr2/props pr2_ctail -pr2/subst1 pr2_gen_cabbr - -pr3/fwd pr3_gen_abst -pr3/fwd pr3_gen_lref -pr3/fwd pr3_gen_void -pr3/fwd pr3_gen_appl -pr3/fwd pr3_gen_bind -pr3/iso pr3_iso_appls_abbr -pr3/iso pr3_iso_appl_bind -pr3/iso pr3_iso_appls_appl_bind -pr3/iso pr3_iso_appls_bind -pr3/iso pr3_iso_beta -pr3/iso pr3_iso_appls_beta -pr3/pr1 pr3_pr1 -pr3/pr3 pr3_strip -pr3/props pr3_thin_dx -pr3/props pr3_head_1 -pr3/props pr3_head_2 -pr3/props pr3_head_21 -pr3/props pr3_head_12 -pr3/props pr3_flat -pr3/props pr3_eta -pr3/subst1 pr3_subst1 -pr3/subst1 pr3_gen_cabbr -sn3/props sn3_cpr3_trans - -sn3/props sn3_shift -sn3/props sn3_change -sn3/props sn3_gen_def -sn3/props sn3_cdelta -sn3/props sn3_appl_lref -sn3/props sn3_appl_abbr -sn3/props sn3_appl_cast -sn3/props sn3_appl_beta -sn3/props sn3_appl_appls -sn3/props sn3_appls_lref -sn3/props sn3_appls_cast -sn3/props sn3_appls_bind -sn3/props sn3_appls_beta -sn3/props sn3_appls_abbr -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 -sty1/props sty1_appl -sty1/props sty1_lift -sty1/props sty1_correct -sty1/props sty1_abbr -sty1/props sty1_cast2 -subst/fwd subst_sort -subst/fwd subst_lref_lt -subst/fwd subst_lref_eq -subst/fwd subst_lref_gt -subst/fwd subst_head -subst/props subst_lift_SO -subst/props subst_subst0 -T/dec binder_dec -T/dec abst_dec -tlist/props tslt_wf__q_ind -tlist/props tslt_wf_ind -tlist/props theads_tapp -tlist/props tcons_tapp_ex -tlist/props tlist_ind_rev -ty3/arity ty3_arity -ty3/arity_props ty3_predicative -ty3/arity_props ty3_repellent -ty3/arity_props ty3_acyclic -ty3/arity_props ty3_sn3 -ty3/dec ty3_inference -ty3/fsubst0 ty3_fsubst0 -ty3/fsubst0 ty3_csubst0 -ty3/fsubst0 ty3_subst0 -ty3/fwd ty3_gen_sort -ty3/fwd ty3_gen_lref -ty3/fwd ty3_gen_bind -ty3/fwd ty3_gen_appl -ty3/fwd ty3_gen_cast -ty3/fwd tys3_gen_nil -ty3/fwd tys3_gen_cons -ty3/fwd_nf2 ty3_gen_appl_nf2 -ty3/fwd_nf2 ty3_inv_lref_nf2_pc3 -ty3/fwd_nf2 ty3_inv_lref_nf2 -ty3/fwd_nf2 ty3_inv_appls_lref_nf2 -ty3/fwd_nf2 ty3_inv_lref_lref_nf2 -ty3/nf2 ty3_nf2_inv_abst_premise_csort -ty3/nf2 ty3_nf2_inv_all -ty3/nf2 ty3_nf2_inv_sort -ty3/nf2 ty3_nf2_gen__ty3_nf2_inv_abst_aux -ty3/nf2 ty3_nf2_inv_abst -ty3/pr3 ty3_sred_wcpr0_pr0 -ty3/pr3 ty3_sred_pr0 -ty3/pr3 ty3_sred_pr1 -ty3/pr3 ty3_sred_pr2 -ty3/pr3 ty3_sred_pr3 -ty3/pr3_props ty3_cred_pr2 -ty3/pr3_props ty3_cred_pr3 -ty3/pr3_props ty3_gen_lift -ty3/pr3_props ty3_tred -ty3/pr3_props ty3_sconv_pc3 -ty3/pr3_props ty3_sred_back -ty3/pr3_props ty3_sconv -ty3/props ty3_lift -ty3/props ty3_correct -ty3/props ty3_unique -ty3/props ty3_gen_abst_abst -ty3/props ty3_typecheck -ty3/props ty3_getl_subst0 -ty3/sty0 ty3_sty0 -ty3/subst1 ty3_gen_cabbr -ty3/subst1 ty3_gen_cvoid -wf3/clear wf3_clear_conf -wf3/clear clear_wf3_trans -wf3/fwd wf3_gen_sort1 -wf3/fwd wf3_gen_bind1 -wf3/fwd wf3_gen_flat1 -wf3/fwd wf3_gen_head2 -wf3/getl wf3_getl_conf -wf3/getl getl_wf3_trans -wf3/props wf3_mono -wf3/props wf3_total -wf3/props ty3_shift1 -wf3/props wf3_idem -wf3/props wf3_ty3 -wf3/ty3 wf3_pr2_conf -wf3/ty3 wf3_pr3_conf -wf3/ty3 wf3_pc3_conf -wf3/ty3 wf3_ty3_conf - -# check ###################################################################### diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/acp.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/acp.ma deleted file mode 100644 index 8be75a31f..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/acp.ma +++ /dev/null @@ -1,55 +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/unfold/ldrops.ma". - -(* ABSTRACT COMPUTATION PROPERTIES ******************************************) - -definition CP1 ≝ λRR:lenv→relation term. λRS:relation term. - ∀L,k. NF … (RR L) RS (⋆k). - -definition CP2 ≝ λRR:lenv→relation term. λRS:relation term. - ∀L,K,W,i. ⇩[0,i] L ≡ K. ⓛW → NF … (RR L) RS (#i). - -definition CP3 ≝ λRR:lenv→relation term. λRP:lenv→predicate term. - ∀L,V,k. RP L (ⓐ⋆k.V) → RP L V. - -definition CP4 ≝ λRR:lenv→relation term. λRS:relation term. - ∀L0,L,T,T0,d,e. NF … (RR L) RS T → - ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → NF … (RR L0) RS T0. - -definition CP4s ≝ λRR:lenv→relation term. λRS:relation term. - ∀L0,L,des. ⇩*[des] L0 ≡ L → - ∀T,T0. ⇧*[des] T ≡ T0 → - NF … (RR L) RS T → NF … (RR L0) RS T0. - -(* requirements for abstract computation properties *) -record acp (RR:lenv->relation term) (RS:relation term) (RP:lenv→predicate term) : Prop ≝ -{ cp1: CP1 RR RS; - cp2: CP2 RR RS; - cp3: CP3 RR RP; - cp4: CP4 RR RS -}. - -(* Basic properties *********************************************************) - -(* Basic_1: was: nf2_lift1 *) -lemma acp_lifts: ∀RR,RS. CP4 RR RS → CP4s RR RS. -#RR #RS #HRR #L1 #L2 #des #H elim H -L1 -L2 -des -[ #L #T1 #T2 #H #HT1 - <(lifts_inv_nil … H) -H // -| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2 - elim (lifts_inv_cons … H) -H /3 width=9/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/acp_aaa.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/acp_aaa.ma deleted file mode 100644 index 1cab5d4b8..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/acp_aaa.ma +++ /dev/null @@ -1,101 +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/unfold/lifts_lifts.ma". -include "Basic_2/unfold/ldrops_ldrops.ma". -include "Basic_2/static/aaa_lifts.ma". -include "Basic_2/static/aaa_aaa.ma". -include "Basic_2/computation/lsubc_ldrops.ma". - -(* ABSTRACT COMPUTATION PROPERTIES ******************************************) - -(* Main propertis ***********************************************************) - -(* Basic_1: was: sc3_arity_csubc *) -theorem aacr_aaa_csubc_lifts: ∀RR,RS,RP. - acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L1,T,A. L1 ⊢ T ÷ A → ∀L0,des. ⇩*[des] L0 ≡ L1 → - ∀T0. ⇧*[des] T ≡ T0 → ∀L2. L2 [RP] ⊑ L0 → - ⦃L2, T0⦄ [RP] ϵ 〚A〛. -#RR #RS #RP #H1RP #H2RP #L1 #T #A #H elim H -L1 -T -A -[ #L #k #L0 #des #HL0 #X #H #L2 #HL20 - >(lifts_inv_sort1 … H) -H - lapply (aacr_acr … H1RP H2RP ⓪) #HAtom - @(s2 … HAtom … ◊) // /2 width=2/ -| #I #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #des #HL01 #X #H #L2 #HL20 - lapply (aacr_acr … H1RP H2RP B) #HB - elim (lifts_inv_lref1 … H) -H #i1 #Hi1 #H destruct - lapply (ldrop_fwd_ldrop2 … HLK1) #HK1b - elim (ldrops_ldrop_trans … HL01 … HLK1) #X #des1 #i0 #HL0 #H #Hi0 #Hdes1 - >(at_mono … Hi1 … Hi0) -i1 - elim (ldrops_inv_skip2 … Hdes1 … H) -des1 #K0 #V0 #des0 #Hdes0 #HK01 #HV10 #H destruct - elim (lsubc_ldrop_O1_trans … HL20 … HL0) -HL0 #X #HLK2 #H - elim (lsubc_inv_pair2 … H) -H * - [ #K2 #HK20 #H destruct - generalize in match HLK2; generalize in match I; -HLK2 -I * #HLK2 - [ elim (lift_total V0 0 (i0 +1)) #V #HV0 - elim (lifts_lift_trans … Hi0 … Hdes0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2 - @(s4 … HB … ◊ … HV0 HLK2) /3 width=7/ (* uses IHB HL20 V2 HV0 *) - | @(s2 … HB … ◊) // /2 width=3/ - ] - | -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hdes0 - #K2 #V2 #A2 #HKV2A #HKV0A #_ #H1 #H2 destruct - lapply (ldrop_fwd_ldrop2 … HLK2) #HLK2b - lapply (aaa_lifts … HK01 … HV10 HKV1B) -HKV1B -HK01 -HV10 #HKV0B - >(aaa_mono … HKV0A … HKV0B) in HKV2A; -HKV0A -HKV0B #HKV2B - elim (lift_total V2 0 (i0 +1)) #V #HV2 - @(s4 … HB … ◊ … HV2 HLK2) - @(s7 … HB … HKV2B) // - ] -| #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20 - elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct - lapply (aacr_acr … H1RP H2RP A) #HA - lapply (aacr_acr … H1RP H2RP B) #HB - lapply (s1 … HB) -HB #HB - @(s5 … HA … ◊ ◊) // /3 width=5/ -| #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL02 - elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct - @(aacr_abst … H1RP H2RP) - [ lapply (aacr_acr … H1RP H2RP B) #HB - @(s1 … HB) /2 width=5/ - | -IHB - #L3 #V3 #T3 #des3 #HL32 #HT03 #HB - elim (lifts_total des3 W0) #W2 #HW02 - elim (ldrops_lsubc_trans … H1RP H2RP … HL32 … HL02) -L2 #L2 #HL32 #HL20 - lapply (aaa_lifts … L2 W2 … (des @ des3) … HLWB) -HLWB /2 width=3/ #HLW2B - @(IHA (L2. ⓛW2) … (des + 1 @ des3 + 1)) -IHA - /2 width=3/ /3 width=5/ - ] -| #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20 - elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct - /3 width=10/ -| #L #V #T #A #_ #_ #IH1A #IH2A #L0 #des #HL0 #X #H #L2 #HL20 - elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct - lapply (aacr_acr … H1RP H2RP A) #HA - lapply (s1 … HA) #H - @(s6 … HA … ◊) /2 width=5/ /3 width=5/ -] -qed. - -(* Basic_1: was: sc3_arity *) -lemma aacr_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L,T,A. L ⊢ T ÷ A → ⦃L, T⦄ [RP] ϵ 〚A〛. -/2 width=8/ qed. - -lemma acp_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L,T,A. L ⊢ T ÷ A → RP L T. -#RR #RS #RP #H1RP #H2RP #L #T #A #HT -lapply (aacr_acr … H1RP H2RP A) #HA -@(s1 … HA) /2 width=4/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/acp_cr.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/acp_cr.ma deleted file mode 100644 index 6eb71054b..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/acp_cr.ma +++ /dev/null @@ -1,174 +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/grammar/aarity.ma". -include "Basic_2/unfold/gr2_gr2.ma". -include "Basic_2/unfold/lifts_lift_vector.ma". -include "Basic_2/unfold/ldrops_ldrop.ma". -include "Basic_2/computation/acp.ma". - -(* ABSTRACT COMPUTATION PROPERTIES ******************************************) - -(* Note: this is Girard's CR1 *) -definition S1 ≝ λRP,C:lenv→predicate term. - ∀L,T. C L T → RP L T. - -(* Note: this is Tait's iii, or Girard's CR4 *) -definition S2 ≝ λRR:lenv→relation term. λRS:relation term. λRP,C:lenv→predicate term. - ∀L,Vs. all … (RP L) Vs → - ∀T. 𝐒[T] → NF … (RR L) RS T → C L (ⒶVs.T). - -(* Note: this is Tait's ii *) -definition S3 ≝ λRP,C:lenv→predicate term. - ∀L,Vs,V,T,W. C L (ⒶVs. ⓓV. T) → RP L W → C L (ⒶVs. ⓐV. ⓛW. T). - -definition S4 ≝ λRP,C:lenv→predicate term. ∀L,K,Vs,V1,V2,i. - C L (ⒶVs. V2) → ⇧[0, i + 1] V1 ≡ V2 → - ⇩[0, i] L ≡ K. ⓓV1 → C L (Ⓐ Vs. #i). - -definition S5 ≝ λRP,C:lenv→predicate term. - ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀V,T. C (L. ⓓV) (ⒶV2s. T) → RP L V → C L (ⒶV1s. ⓓV. T). - -definition S6 ≝ λRP,C:lenv→predicate term. - ∀L,Vs,T,W. C L (ⒶVs. T) → RP L W → C L (ⒶVs. ⓣW. T). - -definition S7 ≝ λC:lenv→predicate term. ∀L2,L1,T1,d,e. - C L1 T1 → ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → C L2 T2. - -definition S7s ≝ λC:lenv→predicate term. - ∀L1,L2,des. ⇩*[des] L2 ≡ L1 → - ∀T1,T2. ⇧*[des] T1 ≡ T2 → C L1 T1 → C L2 T2. - -(* properties of the abstract candidate of reducibility *) -record acr (RR:lenv->relation term) (RS:relation term) (RP,C:lenv→predicate term) : Prop ≝ -{ s1: S1 RP C; - s2: S2 RR RS RP C; - s3: S3 RP C; - s4: S4 RP C; - s5: S5 RP C; - s6: S6 RP C; - s7: S7 C -}. - -(* the abstract candidate of reducibility associated to an atomic arity *) -let rec aacr (RP:lenv→predicate term) (A:aarity) (L:lenv) on A: predicate term ≝ -λT. match A with -[ AAtom ⇒ RP L T -| APair B A ⇒ ∀L0,V0,T0,des. aacr RP B L0 V0 → ⇩*[des] L0 ≡ L → ⇧*[des] T ≡ T0 → - aacr RP A L0 (ⓐV0. T0) -]. - -interpretation - "candidate of reducibility of an atomic arity (abstract)" - 'InEInt RP L T A = (aacr RP A L T). - -(* Basic properties *********************************************************) - -(* Basic_1: was: sc3_lift1 *) -lemma acr_lifts: ∀C. S7 C → S7s C. -#C #HC #L1 #L2 #des #H elim H -L1 -L2 -des -[ #L #T1 #T2 #H #HT1 - <(lifts_inv_nil … H) -H // -| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2 - elim (lifts_inv_cons … H) -H /3 width=9/ -] -qed. - -lemma rp_lifts: ∀RR,RS,RP. acr RR RS RP (λL,T. RP L T) → - ∀des,L0,L,V,V0. ⇩*[des] L0 ≡ L → ⇧*[des] V ≡ V0 → - RP L V → RP L0 V0. -#RR #RS #RP #HRP #des #L0 #L #V #V0 #HL0 #HV0 #HV -@acr_lifts /width=6/ -@(s7 … HRP) -qed. - -(* Basic_1: was only: sns3_lifts1 *) -lemma rp_liftsv_all: ∀RR,RS,RP. acr RR RS RP (λL,T. RP L T) → - ∀des,L0,L,Vs,V0s. ⇧*[des] Vs ≡ V0s → ⇩*[des] L0 ≡ L → - all … (RP L) Vs → all … (RP L0) V0s. -#RR #RS #RP #HRP #des #L0 #L #Vs #V0s #H elim H -Vs -V0s normalize // -#T1s #T2s #T1 #T2 #HT12 #_ #IHT2s #HL0 * #HT1 #HT1s -@conj /2 width=1/ /2 width=6 by rp_lifts/ -qed. - -(* Basic_1: was: - sc3_sn3 sc3_abst sc3_appl sc3_abbr sc3_bind sc3_cast sc3_lift -*) -lemma aacr_acr: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀A. acr RR RS RP (aacr RP A). -#RR #RS #RP #H1RP #H2RP #A elim A -A normalize // -#B #A #IHB #IHA @mk_acr normalize -[ #L #T #H - lapply (H ? (⋆0) ? ⟠ ? ? ?) -H - [1,3: // |2,4: skip - | @(s2 … IHB … ◊) // /2 width=2/ - | #H @(cp3 … H1RP … 0) @(s1 … IHA) // - ] -| #L #Vs #HVs #T #H1T #H2T #L0 #V0 #X #des #HB #HL0 #H - elim (lifts_inv_applv1 … H) -H #V0s #T0 #HV0s #HT0 #H destruct - lapply (s1 … IHB … HB) #HV0 - @(s2 … IHA … (V0 :: V0s)) /2 width=4 by lifts_simple_dx/ /3 width=6/ -| #L #Vs #U #T #W #HA #HW #L0 #V0 #X #des #HB #HL0 #H - elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct - elim (lifts_inv_flat1 … HY) -HY #U0 #X #HU0 #HX #H destruct - elim (lifts_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT0 #H destruct - @(s3 … IHA … (V0 :: V0s)) /2 width=6 by rp_lifts/ /4 width=5/ -| #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #des #HB #HL0 #H - elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct - elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct - elim (ldrops_ldrop_trans … HL0 … HLK) #X #des0 #i1 #HL02 #H #Hi1 #Hdes0 - >(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02 - elim (ldrops_inv_skip2 … Hdes0 … H) -H -des0 #L2 #W1 #des0 #Hdes0 #HLK #HVW1 #H destruct - elim (lift_total W1 0 (i0 + 1)) #W2 #HW12 - elim (lifts_lift_trans … Hdes0 … HVW1 … HW12) // -Hdes0 -Hi0 #V3 #HV13 #HVW2 - >(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2 - @(s4 … IHA … (V0 :: V0s) … HW12 HL02) /3 width=4/ -| #L #V1s #V2s #HV12s #V #T #HA #HV #L0 #V10 #X #des #HB #HL0 #H - elim (lifts_inv_applv1 … H) -H #V10s #Y #HV10s #HY #H destruct - elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct - elim (lift_total V10 0 1) #V20 #HV120 - elim (liftv_total 0 1 V10s) #V20s #HV120s - @(s5 … IHA … (V10 :: V10s) (V20 :: V20s)) /2 width=1/ /2 width=6 by rp_lifts/ - @(HA … (des + 1)) /2 width=1/ - [ @(s7 … IHB … HB … HV120) /2 width=1/ - | @lifts_applv // - elim (liftsv_liftv_trans_le … HV10s … HV120s) -V10s #V10s #HV10s #HV120s - >(liftv_mono … HV12s … HV10s) -V1s // - ] -| #L #Vs #T #W #HA #HW #L0 #V0 #X #des #HB #HL0 #H - elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct - elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct - @(s6 … IHA … (V0 :: V0s)) /2 width=6 by rp_lifts/ /3 width=4/ -| /3 width=7/ -] -qed. - -lemma aacr_abst: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L,W,T,A,B. RP L W → ( - ∀L0,V0,T0,des. ⇩*[des] L0 ≡ L → ⇧*[des + 1] T ≡ T0 → - ⦃L0, V0⦄ [RP] ϵ 〚B〛 → ⦃L0. ⓓV0, T0⦄ [RP] ϵ 〚A〛 - ) → - ⦃L, ⓛW. T⦄ [RP] ϵ 〚②B. A〛. -#RR #RS #RP #H1RP #H2RP #L #W #T #A #B #HW #HA #L0 #V0 #X #des #HB #HL0 #H -lapply (aacr_acr … H1RP H2RP A) #HCA -lapply (aacr_acr … H1RP H2RP B) #HCB -elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct -lapply (s1 … HCB) -HCB #HCB -@(s3 … HCA … ◊) /2 width=6 by rp_lifts/ -@(s5 … HCA … ◊ ◊) // /2 width=1/ /2 width=3/ -qed. - -(* Basic_1: removed theorems 2: sc3_arity_gen sc3_repl *) -(* Basic_1: removed local theorems 1: sc3_sn3_abst *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs.ma deleted file mode 100644 index 12cd0791f..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs.ma +++ /dev/null @@ -1,85 +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/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Basic_1: includes: pr3_pr2 *) -definition cprs: lenv → relation term ≝ - λL. TC … (cpr L). - -interpretation "context-sensitive parallel computation (term)" - 'PRedStar L T1 T2 = (cprs L T1 T2). - -(* Basic eliminators ********************************************************) - -lemma cprs_ind: ∀L,T1. ∀R:predicate term. R T1 → - (∀T,T2. L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → R T → R T2) → - ∀T2. L ⊢ T1 ➡* T2 → R T2. -#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // -qed-. - -axiom cprs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 → - (∀T1,T. L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → R T → R T1) → - ∀T1. L ⊢ T1 ➡* T2 → R T1. - -(* Basic properties *********************************************************) - -(* Basic_1: was: pr3_refl *) -lemma cprs_refl: ∀L,T. L ⊢ T ➡* T. -/2 width=1/ qed. - -lemma cprs_strap1: ∀L,T1,T,T2. - L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → L ⊢ T1 ➡* T2. -/2 width=3/ qed. - -(* Basic_1: was: pr3_step *) -lemma cprs_strap2: ∀L,T1,T,T2. - L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → L ⊢ T1 ➡* T2. -/2 width=3/ qed. - -(* Note: it does not hold replacing |L1| with |L2| *) -lemma cprs_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡* T2 → - ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡* T2. -/3 width=3/ -qed. - -lemma cprs_flat_dx: ∀I,L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L ⊢ T1 ➡* T2 → - L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. -#I #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind … HT12) -T2 /3 width=1/ -#T #T2 #_ #HT2 #IHT2 -@(cprs_strap1 … IHT2) -IHT2 /2 width=1/ -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_1: was: pr3_gen_sort *) -lemma cprs_inv_sort1: ∀L,U2,k. L ⊢ ⋆k ➡* U2 → U2 = ⋆k. -#L #U2 #k #H @(cprs_ind … H) -U2 // -#U2 #U #_ #HU2 #IHU2 destruct ->(cpr_inv_sort1 … HU2) -HU2 // -qed-. - -(* Basic_1: was: pr3_gen_cast *) -lemma cprs_inv_cast1: ∀L,W1,T1,U2. L ⊢ ⓣW1.T1 ➡* U2 → L ⊢ T1 ➡* U2 ∨ - ∃∃W2,T2. L ⊢ W1 ➡* W2 & L ⊢ T1 ➡* T2 & U2 = ⓣW2.T2. -#L #W1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ -#U2 #U #_ #HU2 * /3 width=3/ * -#W #T #HW1 #HT1 #H destruct -elim (cpr_inv_cast1 … HU2) -HU2 /3 width=3/ * -#W2 #T2 #HW2 #HT2 #H destruct /4 width=5/ -qed-. - -(* Basic_1: removed theorems 2: clear_pr3_trans pr3_cflat *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_cprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_cprs.ma deleted file mode 100644 index 16c1a6256..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_cprs.ma +++ /dev/null @@ -1,91 +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/reducibility/cpr_lift.ma". -include "Basic_2/reducibility/cpr_cpr.ma". -include "Basic_2/reducibility/lcpr_cpr.ma". -include "Basic_2/computation/cprs_lcpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Advanced properties ******************************************************) - -lemma cprs_abbr_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 → - L ⊢ ⓓV1. T1 ➡* ⓓV2. T2. -#L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1 -[ /3 width=5/ -| #T1 #T #HT1 #_ #IHT1 - @(cprs_strap2 … IHT1) -IHT1 /2 width=1/ -] -qed. - -lemma cpr_abbr: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡ T2 → - L ⊢ ⓓV1. T1 ➡* ⓓV2. T2. -/3 width=1/ qed. - -lemma cpr_abbr2: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡ T2 → - L ⊢ ⓓV1. T1 ➡* ⓓV2. T2. -#L #V1 #V2 #HV12 #T1 #T2 #HT12 -lapply (lcpr_cpr_trans (L. ⓓV1) … HT12) /2 width=1/ -qed. - -(* Advanced inversion lemmas ************************************************) - -(* Basic_1: was pr3_gen_appl *) -lemma cprs_inv_appl1: ∀L,V1,T1,U2. L ⊢ ⓐV1. T1 ➡* U2 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 & - U2 = ⓐV2. T2 - | ∃∃V2,W,T. L ⊢ V1 ➡* V2 & - L ⊢ T1 ➡* ⓛW. T & L ⊢ ⓓV2. T ➡* U2 - | ∃∃V0,V2,V,T. L ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 & - L ⊢ T1 ➡* ⓓV. T & L ⊢ ⓓV. ⓐV2. T ➡* U2. -#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ -#U #U2 #_ #HU2 * * -[ #V0 #T0 #HV10 #HT10 #H destruct - elim (cpr_inv_appl1 … HU2) -HU2 * - [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/ - | #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct /4 width=7/ - | #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct - @or3_intro2 @(ex4_4_intro … HV2 HT10) /2 width=3/ /3 width=1/ (**) (* explicit constructor. /5 width=8/ is too slow because TC_transitive gets in the way *) - ] -| /4 width=8/ -| /4 width=10/ -] -qed-. - -(* Main propertis ***********************************************************) - -(* Basic_1: was: pr3_confluence *) -theorem cprs_conf: ∀L,T1,T. L ⊢ T ➡* T1 → ∀T2. L ⊢ T ➡* T2 → - ∃∃T0. L ⊢ T1 ➡* T0 & L ⊢ T2 ➡* T0. -/3 width=3/ qed. - -(* Basic_1: was: pr3_t *) -theorem cprs_trans: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T ➡* T2 → L ⊢ T1 ➡* T2. -/2 width=3/ qed. - -lemma cprs_flat: ∀I,L,T1,T2. L ⊢ T1 ➡* T2 → ∀V1,V2. L ⊢ V1 ➡* V2 → - L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. -#I #L #T1 #T2 #HT12 #V1 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/ -#V #V2 #_ #HV2 #IHV1 -@(cprs_trans … IHV1) -IHV1 /2 width=1/ -qed. - -(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *) -lemma lcpr_cprs_trans: ∀L1,L2. L1 ⊢ ➡ L2 → - ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2. -#L1 #L2 #HL12 #T1 #T2 #H @(cprs_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT2 -@(cprs_trans … IHT2) /2 width=3/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lcpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lcpr.ma deleted file mode 100644 index 7816989a7..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lcpr.ma +++ /dev/null @@ -1,46 +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/reducibility/ltpr_tps.ma". -include "Basic_2/reducibility/cpr_ltpss.ma". -include "Basic_2/reducibility/lcpr.ma". -include "Basic_2/computation/cprs.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Properties concerning context-sensitive parallel reduction on lenv's *****) - -lemma ltpr_tpss_trans: ∀L1,L2. L1 ➡ L2 → ∀T1,T2,d,e. L2 ⊢ T1 [d, e] ▶* T2 → - ∃∃T. L1 ⊢ T1 [d, e] ▶* T & L1 ⊢ T ➡* T2. -#L1 #L2 #HL12 #T1 #T2 #d #e #H @(tpss_ind … H) -T2 -[ /2 width=3/ -| #T #T2 #_ #HT2 * #T0 #HT10 #HT0 - elim (ltpr_tps_trans … HT2 … HL12) -L2 #T3 #HT3 #HT32 - @(ex2_1_intro … HT10) -T1 (**) (* explicit constructors *) - @(cprs_strap1 … T3 …) /2 width=1/ -HT32 - @(cprs_strap1 … HT0) -HT0 /3 width=3/ -] -qed. - -(* Basic_1: was just: pr3_pr0_pr2_t *) -lemma ltpr_cpr_trans: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡* T2. -#L1 #L2 #HL12 #T1 #T2 * #T #HT1 -<(ltpr_fwd_length … HL12) #HT2 -elim (ltpr_tpss_trans … HL12 … HT2) -L2 /3 width=3/ -qed. - -(* Basic_1: was just: pr3_pr2_pr2_t *) -lemma lcpr_cpr_trans: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡* T2. -#L1 #L2 * /3 width=7/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lcprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lcprs.ma deleted file mode 100644 index 6d4f5ec8c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lcprs.ma +++ /dev/null @@ -1,52 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/computation/cprs_cprs.ma". -include "Basic_2/computation/lcprs_lcprs.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Properties exploiting context-senstive computation on local environments *) - -(* Basic_1: was just: pr3_pr3_pr3_t *) -lemma lcprs_cprs_trans: ∀L1,L2. L1 ⊢ ➡* L2 → - ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2. -#L1 #L2 #HL12 @(lcprs_ind … HL12) -L2 // /3 width=3/ -qed. - -(* Advanced inversion lemmas ************************************************) - -(* Basic_1: was pr3_gen_abbr *) -lemma cprs_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡* U2 → - (∃∃V2,T2. L ⊢ V1 ➡* V2 & L. ⓓV1 ⊢ T1 ➡* T2 & - U2 = ⓓV2. T2 - ) ∨ - ∃∃U. ⇧[0, 1] U2 ≡ U & L. ⓓV1 ⊢ T1 ➡* U. -#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ -#U0 #U2 #_ #HU02 * * -[ #V0 #T0 #HV10 #HT10 #H destruct - elim (cpr_inv_abbr1 … HU02) -HU02 * - [ #V #V2 #T2 #HV0 #HV2 #HT02 #H destruct - lapply (cpr_intro … HV0 … HV2) -HV2 #HV02 - lapply (ltpr_cpr_trans (L.ⓓV0) … HT02) /2 width=1/ -V #HT02 - lapply (lcprs_cprs_trans (L. ⓓV1) … HT02) -HT02 /2 width=1/ /4 width=5/ - | -V0 #T2 #HT20 #HTU2 - elim (lift_total U2 0 1) #U0 #HU20 - lapply (cpr_lift (L.ⓓV1) … HT20 … HU20 HTU2) -T2 /2 width=1/ /4 width=5/ - ] -| #U1 #HU01 #HTU1 - elim (lift_total U2 0 1) #U #HU2 - lapply (cpr_lift (L.ⓓV1) … HU01 … HU2 HU02) -U0 /2 width=1/ /4 width=5/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lift.ma deleted file mode 100644 index ccff2f1fe..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_lift.ma +++ /dev/null @@ -1,42 +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/reducibility/cpr_lift.ma". -include "Basic_2/computation/cprs.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Relocation properties ****************************************************) - -(* Basic_1: was: pr3_lift *) -lemma cprs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 → - ∀T2. K ⊢ T1 ➡* T2 → ∀U2. ⇧[d, e] T2 ≡ U2 → - L ⊢ U1 ➡* U2. -#L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #HT12 @(cprs_ind … HT12) -T2 -[ -HLK #T2 #HT12 - <(lift_mono … HTU1 … HT12) -T1 // -| -HTU1 #T #T2 #_ #HT2 #IHT2 #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (cpr_lift … HLK … HTU … HTU2 … HT2) -T2 -HLK /3 width=3/ -] -qed. - -(* Basic_1: was: pr3_gen_lift *) -lemma cprs_inv_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → - ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀U2. L ⊢ U1 ➡* U2 → - ∃∃T2. ⇧[d, e] T2 ≡ U2 & K ⊢ T1 ➡* T2. -#L #K #d #e #HLK #T1 #U1 #HTU1 #U2 #HU12 @(cprs_ind … HU12) -U2 /2 width=3/ --HTU1 #U #U2 #_ #HU2 * #T #HTU #HT1 -elim (cpr_inv_lift … HLK … HTU … HU2) -U -HLK /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_tstc.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_tstc.ma deleted file mode 100644 index e623c7b90..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_tstc.ma +++ /dev/null @@ -1,56 +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/grammar/tstc.ma". -(* -include "Basic_2/reducibility/cpr_lift.ma". -include "Basic_2/reducibility/lcpr_cpr.ma". -*) -include "Basic_2/computation/cprs_cprs.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Forward lemmas involving same top term constructor ***********************) -(* -lemma cpr_fwd_beta: ∀L,V,W,T,U. L ⊢ ⓐV. ⓛW. T ➡ U → - ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⓓV. T ➡* U. -#L #V #W #T #U #H -elim (cpr_inv_appl1 … H) -H * -[ #V0 #X #_ #_ #H destruct /2 width=1/ -| #V0 #W0 #T1 #T2 #HV0 #HT12 #H1 #H2 destruct - lapply (lcpr_cpr_trans (L. ⓓV) … HT12) -HT12 /2 width=1/ /3 width=1/ -| #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #H destruct -] -qed-. - -lemma cpr_fwd_theta: ∀L,V1,V,T,U. L ⊢ ⓐV1. ⓓV. T ➡ U → - ∀V2. ⇧[0, 1] V1 ≡ V2 → ⓐV1. ⓓV. T ≃ U ∨ - L ⊢ ⓓV. ⓐV2. T ➡* U. -#L #V1 #V #T #U #H #V2 #HV12 -elim (cpr_inv_appl1 … H) -H * -[ -HV12 #V0 #X #_ #_ #H destruct /2 width=1/ -| -HV12 #V0 #W #T1 #T2 #_ #_ #H destruct -| #V0 #V3 #W1 #W2 #T1 #T2 #HV10 #HW12 #HT12 #HV03 #H1 #H2 destruct - lapply (cpr_lift (L.ⓓW1) … HV12 … HV03 … HV10) -V0 -HV12 /2 width=1/ #HV23 - lapply (lcpr_cpr_trans (L. ⓓW1) … HT12) -HT12 /2 width=1/ #HT12 - /4 width=1/ -] -qed-. -*) -lemma cprs_fwd_tau: ∀L,W,T,U. L ⊢ ⓣW. T ➡* U → - ⓣW. T ≃ U ∨ L ⊢ T ➡* U. -#L #W #T #U #H -elim (cprs_inv_cast1 … H) -H /2 width=1/ * -#W0 #T0 #_ #_ #H destruct /2 width=1/ -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_tstc_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_tstc_vector.ma deleted file mode 100644 index 3e6f4a20d..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/cprs_tstc_vector.ma +++ /dev/null @@ -1,69 +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/grammar/tstc_vector.ma". -include "Basic_2/substitution/lift_vector.ma". -include "Basic_2/computation/cprs_tstc.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) - -(* Vector form of forward lemmas involving same top term constructor ********) -(* -lemma cpr_fwd_beta_vector: ∀L,V,W,T,U,Vs. L ⊢ ⒶVs. ⓐV. ⓛW. T ➡ U → - ⒶVs. ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⒶVs. ⓓV. T ➡* U. -#L #V #W #T #U * /2 width=1 by cpr_fwd_beta/ -#V0 #Vs #H -elim (cpr_inv_appl1_simple … H ?) -H -[ #V1 #T1 #_ #_ #H destruct /2 width=1/ -| elim Vs -Vs // -] -qed-. - -lemma cpr_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡ U → - ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U. -#L #V1s #V2s * -V1s -V2s /3 width=1/ -#V1s #V2s #V1a #V2a #HV12a * -V1s -V2s /2 width=1 by cpr_fwd_theta/ -HV12a -#V1s #V2s #V1b #V2b #_ #_ #V #U #T #H -elim (cpr_inv_appl1_simple … H ?) -H // -#V0 #T0 #_ #_ #H destruct /2 width=1/ -qed-. -*) - -(* Basic_1: was: pr3_iso_appls_cast *) -lemma cprs_fwd_tau_vector: ∀L,Vs,W,T,U. L ⊢ ⒶVs. ⓣW. T ➡* U → - ⒶVs. ⓣW. T ≃ U ∨ L ⊢ ⒶVs. T ➡* U. -#L #Vs elim Vs -Vs /2 width=1 by cprs_fwd_tau/ -#V #Vs #IHVs #W #T #U #H -elim (cprs_inv_appl1 … H) -H * -[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/ -| #V0 #W0 #T0 #HV0 #HT0 #HU - elim (IHVs … HT0) -IHVs -HT0 #HT0 - [ elim (tstc_inv_bind_appls_simple … HT0 ?) // - | @or_intror - @(cprs_trans … HU) -HU - @(cprs_strap1 … (ⓐV0.ⓛW0.T0)) /2 width=1/ -HV0 -HT0 /3 width=1/ - ] -| #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU - elim (IHVs … HT0) -IHVs -HT0 #HT0 - [ elim (tstc_inv_bind_appls_simple … HT0 ?) // - | @or_intror - @(cprs_trans … HU) -HU - @(cprs_strap1 … (ⓐV0.ⓓV2.T0)) /2 width=1/ -HV0 -HT0 /3 width=3/ -] -qed-. - -axiom cprs_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡* U → - ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn.ma deleted file mode 100644 index 043fa9a27..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn.ma +++ /dev/null @@ -1,85 +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/reducibility/cpr.ma". -include "Basic_2/reducibility/cnf.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) - -definition csn: lenv → predicate term ≝ λL. SN … (cpr L) (eq …). - -interpretation - "context-sensitive strong normalization (term)" - 'SN L T = (csn L T). - -(* Basic eliminators ********************************************************) - -lemma csn_ind: ∀L. ∀R:predicate term. - (∀T1. L ⊢ ⬇* T1 → - (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → R T2) → - R T1 - ) → - ∀T. L ⊢ ⬇* T → R T. -#L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1 -@H0 -H0 /3 width=1/ -IHT1 /4 width=1/ -qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: sn3_pr2_intro *) -lemma csn_intro: ∀L,T1. - (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇* T2) → L ⊢ ⬇* T1. -#L #T1 #H -@(SN_intro … H) -qed. - -(* Basic_1: was: sn3_nf2 *) -lemma csn_cnf: ∀L,T. L ⊢ 𝐍[T] → L ⊢ ⬇* T. -/2 width=1/ qed. - -lemma csn_cpr_trans: ∀L,T1. L ⊢ ⬇* T1 → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ ⬇* T2. -#L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12 -@csn_intro #T #HLT2 #HT2 -elim (term_eq_dec T1 T2) #HT12 -[ -IHT1 -HLT12 destruct /3 width=1/ -| -HT1 -HT2 /3 width=4/ -qed. - -(* Basic_1: was: sn3_cast *) -lemma csn_cast: ∀L,W. L ⊢ ⬇* W → ∀T. L ⊢ ⬇* T → L ⊢ ⬇* ⓣW. T. -#L #W #HW elim HW -W #W #_ #IHW #T #HT @(csn_ind … HT) -T #T #HT #IHT -@csn_intro #X #H1 #H2 -elim (cpr_inv_cast1 … H1) -H1 -[ * #W0 #T0 #HLW0 #HLT0 #H destruct - elim (eq_false_inv_tpair_sn … H2) -H2 - [ /3 width=3/ - | -HLW0 * #H destruct /3 width=1/ - ] -| /3 width=3/ -] -qed. - -(* Basic forward lemmas *****************************************************) - -fact csn_fwd_flat_dx_aux: ∀L,U. L ⊢ ⬇* U → ∀I,V,T. U = ⓕ{I} V. T → L ⊢ ⬇* T. -#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct -@csn_intro #T2 #HLT2 #HT2 -@(IH (ⓕ{I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/ -qed. - -(* Basic_1: was: sn3_gen_flat *) -lemma csn_fwd_flat_dx: ∀I,L,V,T. L ⊢ ⬇* ⓕ{I} V. T → L ⊢ ⬇* T. -/2 width=5/ qed-. - -(* Basic_1: removed theorems 3: sn3_gen_cflat sn3_cflat sn3_bind *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_aaa.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_aaa.ma deleted file mode 100644 index 421a3a7f7..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_aaa.ma +++ /dev/null @@ -1,25 +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/computation/acp_aaa.ma". -include "Basic_2/computation/csn_lcpr_vector.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) - -(* Properties concerning atomic arity assignment ****************************) - -lemma csn_aaa: ∀L,T,A. L ⊢ T ÷ A → L ⊢ ⬇* T. -#L #T #A #H -@(acp_aaa … csn_acp csn_acr … H) -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cpr.ma deleted file mode 100644 index 11647e9dc..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cpr.ma +++ /dev/null @@ -1,41 +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/reducibility/cpr_cpr.ma". -include "Basic_2/computation/csn.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) - -(* Advanced forvard lemmas **************************************************) - -fact csn_fwd_pair_sn_aux: ∀L,U. L ⊢ ⬇* U → ∀I,V,T. U = ②{I} V. T → L ⊢ ⬇* V. -#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct -@csn_intro #V2 #HLV2 #HV2 -@(IH (②{I} V2. T)) -IH // /2 width=1/ -HLV2 #H destruct /2 width=1/ -qed. - -(* Basic_1: was: sn3_gen_head *) -lemma csn_fwd_pair_sn: ∀I,L,V,T. L ⊢ ⬇* ②{I} V. T → L ⊢ ⬇* V. -/2 width=5/ qed. - -fact csn_fwd_bind_dx_aux: ∀L,U. L ⊢ ⬇* U → - ∀I,V,T. U = ⓑ{I} V. T → L. ⓑ{I} V ⊢ ⬇* T. -#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct -@csn_intro #T2 #HLT2 #HT2 -@(IH (ⓑ{I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/ -qed. - -(* Basic_1: was: sn3_gen_bind *) -lemma csn_fwd_bind_dx: ∀I,L,V,T. L ⊢ ⬇* ⓑ{I} V. T → L. ⓑ{I} V ⊢ ⬇* T. -/2 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cprs.ma deleted file mode 100644 index bf6683d3c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_cprs.ma +++ /dev/null @@ -1,111 +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/computation/cprs.ma". -include "Basic_2/computation/csn.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) - -(* Properties concerning context-sensitive computation on terms *************) - -definition csns: lenv → predicate term ≝ λL. SN … (cprs L) (eq …). - -interpretation - "context-sensitive strong normalization (term)" - 'SNStar L T = (csns L T). - -(* Basic eliminators ********************************************************) - -lemma csns_ind: ∀L. ∀R:predicate term. - (∀T1. L ⊢ ⬇** T1 → - (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → R T2) → R T1 - ) → - ∀T. L ⊢ ⬇** T → R T. -#L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1 -@H0 -H0 /3 width=1/ -IHT1 /4 width=1/ -qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: sn3_intro *) -lemma csns_intro: ∀L,T1. - (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → L ⊢ ⬇** T2) → L ⊢ ⬇** T1. -#L #T1 #H -@(SN_intro … H) -qed. - -fact csns_intro_aux: ∀L,T1. - (∀T,T2. L ⊢ T ➡* T2 → T1 = T → (T1 = T2 → False) → L ⊢ ⬇** T2) → L ⊢ ⬇** T1. -/4 width=3/ qed-. - -(* Basic_1: was: sn3_pr3_trans (old version) *) -lemma csns_cprs_trans: ∀L,T1. L ⊢ ⬇** T1 → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ ⬇** T2. -#L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12 -@csns_intro #T #HLT2 #HT2 -elim (term_eq_dec T1 T2) #HT12 -[ -IHT1 -HLT12 destruct /3 width=1/ -| -HT1 -HT2 /3 width=4/ -qed. - -(* Basic_1: was: sn3_pr2_intro (old version) *) -lemma csns_intro_cpr: ∀L,T1. - (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇** T2) → - L ⊢ ⬇** T1. -#L #T1 #H -@csns_intro_aux #T #T2 #H @(cprs_ind_dx … H) -T -[ -H #H destruct #H - elim (H ?) // -| #T0 #T #HLT1 #HLT2 #IHT #HT10 #HT12 destruct - elim (term_eq_dec T0 T) #HT0 - [ -HLT1 -HLT2 -H /3 width=1/ - | -IHT -HT12 /4 width=3/ - ] -] -qed. - -(* Main properties **********************************************************) - -theorem csn_csns: ∀L,T. L ⊢ ⬇* T → L ⊢ ⬇** T. -#L #T #H @(csn_ind … H) -T /4 width=1/ -qed. - -theorem csns_csn: ∀L,T. L ⊢ ⬇** T → L ⊢ ⬇* T. -#L #T #H @(csns_ind … H) -T /4 width=1/ -qed. - -(* Basic_1: was: sn3_pr3_trans *) -lemma csn_cprs_trans: ∀L,T1. L ⊢ ⬇* T1 → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ ⬇* T2. -/4 width=3/ qed. - -(* Basic_1: was: nf2_sn3 *) -lemma csn_cwn: ∀L,T1. L ⊢ ⬇* T1 → - ∃∃T2. L ⊢ T1 ➡* T2 & L ⊢ 𝐍[T2]. -#L #T1 #H elim H -T1 #T1 #_ #IHT1 -elim (cnf_dec L T1) -[ -IHT1 /2 width=3/ -| * #T2 #HLT12 #HT12 - elim (IHT1 T2 ? ?) -IHT1 // /2 width=1/ -HT12 /3 width=3/ -] -qed. - -(* Main eliminators *********************************************************) - -lemma csn_ind_cprs: ∀L. ∀R:predicate term. - (∀T1. L ⊢ ⬇* T1 → - (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → R T2) → R T1 - ) → - ∀T. L ⊢ ⬇* T → R T. -#L #R #H0 #T1 #H @(csns_ind … (csn_csns … H)) -T1 #T1 #HT1 #IHT1 -@H0 -H0 /2 width=1/ -HT1 /3 width=1/ -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma deleted file mode 100644 index 303be5997..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr.ma +++ /dev/null @@ -1,146 +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/grammar/tstc_tstc.ma". -include "Basic_2/computation/cprs_cprs.ma". -include "Basic_2/computation/csn_lift.ma". -include "Basic_2/computation/csn_cpr.ma". -include "Basic_2/computation/csn_cprs.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) - -(* Advanced properties ******************************************************) - -lemma csn_lcpr_conf: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T. L1 ⊢ ⬇* T → L2 ⊢ ⬇* T. -#L1 #L2 #HL12 #T #H @(csn_ind_cprs … H) -T #T #_ #IHT -@csn_intro #T0 #HLT0 #HT0 -@IHT /2 width=2/ -IHT -HT0 /2 width=3/ -qed. - -lemma csn_abbr: ∀L,V. L ⊢ ⬇* V → ∀T. L. ⓓV ⊢ ⬇* T → L ⊢ ⬇* ⓓV. T. -#L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_cprs … HT) -T #T #HT #IHT -@csn_intro #X #H1 #H2 -elim (cpr_inv_abbr1 … H1) -H1 * -[ #V0 #V1 #T1 #HLV0 #HLV01 #HLT1 #H destruct - lapply (cpr_intro … HLV0 HLV01) -HLV01 #HLV1 - lapply (ltpr_cpr_trans (L. ⓓV) … HLT1) /2 width=1/ -V0 #HLT1 - elim (eq_false_inv_tpair_sn … H2) -H2 - [ #HV1 @IHV // /2 width=1/ -HV1 - @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -HLV1 /2 width=3/ - | -IHV -HLV1 * #H destruct /3 width=1/ - ] -| -IHV -IHT -H2 #T0 #HT0 #HLT0 - lapply (csn_inv_lift … HT … HT0) -HT /2 width=3/ -] -qed. - -fact csn_appl_beta_aux: ∀L,W. L ⊢ ⬇* W → ∀U. L ⊢ ⬇* U → - ∀V,T. U = ⓓV. T → L ⊢ ⬇* ⓐV. ⓛW. T. -#L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V #T #H destruct -lapply (csn_fwd_pair_sn … HVT) #HV -lapply (csn_fwd_bind_dx … HVT) #HT -HVT -@csn_intro #X #H #H2 -elim (cpr_inv_appl1 … H) -H * -[ #V0 #Y #HLV0 #H #H0 destruct - elim (cpr_inv_abst1 … H Abbr V) -H #W0 #T0 #HLW0 #HLT0 #H destruct - elim (eq_false_inv_beta … H2) -H2 - [ -IHVT #HW0 @IHW -IHW [1,5: // |3: skip ] -HLW0 /2 width=1/ -HW0 - @csn_abbr /2 width=3/ -HV - @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/ - | -IHW -HLW0 -HV -HT * #H #HVT0 destruct - @(IHVT … HVT0) -IHVT -HVT0 // /2 width=1/ - ] -| -IHW -IHVT -H2 #V0 #W0 #T0 #T1 #HLV0 #HLT01 #H1 #H2 destruct - lapply (lcpr_cpr_trans (L. ⓓV) … HLT01) -HLT01 /2 width=1/ #HLT01 - @csn_abbr /2 width=3/ -HV - @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/ -| -IHW -IHVT -HV -HT -H2 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #_ #H destruct -] -qed. - -(* Basic_1: was: sn3_beta *) -lemma csn_appl_beta: ∀L,W. L ⊢ ⬇* W → ∀V,T. L ⊢ ⬇* (ⓓV. T) → (**) - L ⊢ ⬇* ⓐV. ⓛW. T. -/2 width=3/ qed. - -fact csn_appl_theta_aux: ∀L,U. L ⊢ ⬇* U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → - ∀V,T. U = ⓓV. ⓐV2. T → L ⊢ ⬇* ⓐV1. ⓓV. T. -#L #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct -lapply (csn_fwd_pair_sn … HVT) #HV -lapply (csn_fwd_bind_dx … HVT) -HVT #HVT -@csn_intro #X #HL #H -elim (cpr_inv_appl1 … HL) -HL * -[ -HV #V0 #Y #HLV10 #HL #H0 destruct - elim (cpr_inv_abbr1 … HL) -HL * - [ #V3 #V4 #T3 #HV3 #HLV34 #HLT3 #H0 destruct - lapply (cpr_intro … HV3 HLV34) -HLV34 #HLV34 - elim (lift_total V0 0 1) #V5 #HV05 - elim (term_eq_dec (ⓓV.ⓐV2.T) (ⓓV4.ⓐV5.T3)) - [ -IHVT #H0 destruct - elim (eq_false_inv_tpair_sn … H) -H - [ -HLV10 -HLV34 -HV3 -HLT3 -HVT - >(lift_inj … HV12 … HV05) -V5 - #H elim (H ?) // - | * #_ #H elim (H ?) // - ] - | -H -HVT #H - lapply (cpr_lift (L. ⓓV) … HV12 … HV05 HLV10) -HLV10 -HV12 /2 width=1/ #HV25 - lapply (ltpr_cpr_trans (L. ⓓV) … HLT3) /2 width=1/ -HLT3 #HLT3 - @(IHVT … H … HV05) -IHVT // -H -HV05 /3 width=1/ - ] - | -H -IHVT #T0 #HT0 #HLT0 - @(csn_cpr_trans … (ⓐV1.T0)) /2 width=1/ -V0 -Y - @(csn_inv_lift … 0 1 HVT) /2 width=1/ - ] -| -HV -HV12 -HVT -IHVT -H #V0 #W0 #T0 #T1 #_ #_ #H destruct -| -IHVT -H #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HLW01 #HLT01 #HV03 #H1 #H2 destruct - lapply (cpr_lift (L. ⓓW0) … HV12 … HV03 HLV10) -HLV10 -HV12 -HV03 /2 width=1/ #HLV23 - lapply (lcpr_cpr_trans (L. ⓓW0) … HLT01) -HLT01 /2 width=1/ #HLT01 - @csn_abbr /2 width=3/ -HV - @(csn_lcpr_conf (L. ⓓW0)) /2 width=1/ -W1 - @(csn_cprs_trans … HVT) -HVT /2 width=1/ -] -qed. - -(* Basic_1: was: sn3_appl_bind *) -lemma csn_appl_theta: ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → - ∀L,V,T. L ⊢ ⬇* ⓓV. ⓐV2. T → L ⊢ ⬇* ⓐV1. ⓓV. T. -/2 width=5/ qed. - -(* Basic_1: was only: sn3_appl_appl *) -lemma csn_appl_simple_tstc: ∀L,V. L ⊢ ⬇* V → ∀T1. - L ⊢ ⬇* T1 → - (∀T2. L ⊢ T1 ➡* T2 → (T1 ≃ T2 → False) → L ⊢ ⬇* ⓐV. T2) → - 𝐒[T1] → L ⊢ ⬇* ⓐV. T1. -#L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #H @(csn_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1 -@csn_intro #X #HL #H -elim (cpr_inv_appl1_simple … HL ?) -HL // -#V0 #T0 #HLV0 #HLT10 #H0 destruct -elim (eq_false_inv_tpair_sn … H) -H -[ -IHT1 #HV0 - @(csn_cpr_trans … (ⓐV0.T1)) /2 width=1/ -HLT10 - @IHV -IHV // -H1T1 -H3T1 /2 width=1/ -HV0 - #T2 #HLT12 #HT12 - @(csn_cpr_trans … (ⓐV.T2)) /2 width=1/ -HLV0 - @H2T1 -H2T1 // -HLT12 /2 width=1/ -| -IHV -H1T1 -HLV0 * #H #H1T10 destruct - elim (tstc_dec T1 T0) #H2T10 - [ @IHT1 -IHT1 // /2 width=1/ -H1T10 /2 width=3/ -H3T1 - #T2 #HLT02 #HT02 - @H2T1 -H2T1 /2 width=3/ -HLT10 -HLT02 /3 width=3/ - | -IHT1 -H3T1 -H1T10 - @H2T1 -H2T1 /2 width=1/ - ] -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr_vector.ma deleted file mode 100644 index 4ff8c039b..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lcpr_vector.ma +++ /dev/null @@ -1,85 +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/computation/acp_cr.ma". -include "Basic_2/computation/cprs_tstc_vector.ma". -include "Basic_2/computation/csn_lcpr.ma". -include "Basic_2/computation/csn_vector.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERM VECTORS **********************) - -(* Advanced properties ******************************************************) -(* -(* Basic_1: was only: sn3_appl_appls *) -lemma csn_appl_appls_simple_tstc: ∀L,Vs,V,T1. L ⊢ ⬇* V → L ⊢ ⬇* T1 → - (∀T2. L ⊢ ⒶVs.T1 ➡* T2 → (ⒶVs.T1 ≃ T2 → False) → L ⊢ ⬇* ⓐV. T2) → - 𝐒[T1] → L ⊢ ⬇* ⓐV. ⒶVs. T1. -#L * -[ #V #T1 #HV - @csn_appl_simple_tstc // -| #V0 #Vs #V #T1 #HV #H1T1 #H2T1 #H3T1 - @csn_appl_simple_tstc // -HV - [ @H2T1 -] -qed. -*) -lemma csn_applv_theta: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀V,T. L ⊢ ⬇* ⓓV. ⒶV2s. T → L ⊢ ⬇* V → - L ⊢ ⬇* ⒶV1s. ⓓV. T. -#L #V1s #V2s * -V1s -V2s /2 width=1/ -#V1s #V2s #V1 #V2 #HV12 #H -generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1 -elim H -V1s -V2s /2 width=3/ -#V1s #V2s #V1 #V2 #HV12 #HV12s #IHV12s #W1 #W2 #HW12 #V #T #H #HV -lapply (csn_appl_theta … HW12 … H) -H -HW12 #H -lapply (csn_fwd_pair_sn … H) #HW1 -lapply (csn_fwd_flat_dx … H) #H1 -@csn_appl_simple_tstc // -HW1 /2 width=3/ -IHV12s -HV -H1 #X #H1 #H2 -elim (cprs_fwd_theta_vector … (V2::V2s) … H1) -H1 /2 width=1/ -HV12s -HV12 -[ -H #H elim (H2 ?) -H2 // -| -H2 #H1 @(csn_cprs_trans … H) -H /2 width=1/ -] -qed. - -lemma csn_applv_tau: ∀L,W. L ⊢ ⬇* W → - ∀Vs,T. L ⊢ ⬇* ⒶVs. T → - L ⊢ ⬇* ⒶVs. ⓣW. T. -#L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW -#V #Vs #IHV #T #H1T -lapply (csn_fwd_pair_sn … H1T) #HV -lapply (csn_fwd_flat_dx … H1T) #H2T -@csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T -[ #X #H #H0 - elim (cprs_fwd_tau_vector … H) -H #H - [ -H1T elim (H0 ?) -H0 // - | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ - ] -| -H1T elim Vs -Vs // -] -qed. -(* -theorem csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬇* T). -@mk_acr // -[ -| -| -| #L #V1 #V2 #HV12 #V #T #H #HVT - @(csn_applv_theta … HV12) -HV12 // - @(csn_abbr) // -| /2 width=1/ -| @csn_lift -] -qed. -*) -axiom csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬇* T). diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lift.ma deleted file mode 100644 index 01d983c0c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_lift.ma +++ /dev/null @@ -1,100 +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/reducibility/cnf_lift.ma". -include "Basic_2/computation/acp.ma". -include "Basic_2/computation/csn.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) - -(* Relocation properties ****************************************************) - -(* Basic_1: was: sn3_lift *) -lemma csn_lift: ∀L2,L1,T1,d,e. L1 ⊢ ⬇* T1 → - ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → L2 ⊢ ⬇* T2. -#L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12 -@csn_intro #T #HLT2 #HT2 -elim (cpr_inv_lift … HL21 … HT12 … HLT2) -HLT2 #T0 #HT0 #HLT10 -@(IHT1 … HLT10) // -L1 -L2 #H destruct ->(lift_mono … HT0 … HT12) in HT2; -T0 /2 width=1/ -qed. - -(* Basic_1: was: sn3_gen_lift *) -lemma csn_inv_lift: ∀L2,L1,T1,d,e. L1 ⊢ ⬇* T1 → - ∀T2. ⇩[d, e] L1 ≡ L2 → ⇧[d, e] T2 ≡ T1 → L2 ⊢ ⬇* T2. -#L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21 -@csn_intro #T #HLT2 #HT2 -elim (lift_total T d e) #T0 #HT0 -lapply (cpr_lift … HL12 … HT21 … HT0 HLT2) -HLT2 #HLT10 -@(IHT1 … HLT10) // -L1 -L2 #H destruct ->(lift_inj … HT0 … HT21) in HT2; -T0 /2 width=1/ -qed. - -(* Advanced properties ******************************************************) - -(* Basic_1: was: sn3_abbr *) -lemma csn_lref_abbr: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓓV → K ⊢ ⬇* V → L ⊢ ⬇* #i. -#L #K #V #i #HLK #HV -@csn_intro #X #H #Hi -elim (cpr_inv_lref1 … H) -H -[ #H destruct elim (Hi ?) // -| -Hi * #K0 #V0 #V1 #HLK0 #HV01 #HV1 #_ - lapply (ldrop_mono … HLK0 … HLK) -HLK #H destruct - lapply (ldrop_fwd_ldrop2 … HLK0) -HLK0 #HLK - @(csn_lift … HLK HV1) -HLK -HV1 - @(csn_cpr_trans … HV) -HV - @(cpr_intro … HV01) -HV01 // -] -qed. - -lemma csn_abst: ∀L,W. L ⊢ ⬇* W → ∀I,V,T. L. ⓑ{I} V ⊢ ⬇* T → L ⊢ ⬇* ⓛW. T. -#L #W #HW elim HW -W #W #_ #IHW #I #V #T #HT @(csn_ind … HT) -T #T #HT #IHT -@csn_intro #X #H1 #H2 -elim (cpr_inv_abst1 … H1 I V) -H1 -#W0 #T0 #HLW0 #HLT0 #H destruct -elim (eq_false_inv_tpair_sn … H2) -H2 -[ /3 width=5/ -| -HLW0 * #H destruct /3 width=1/ -] -qed. - -lemma csn_appl_simple: ∀L,V. L ⊢ ⬇* V → ∀T1. - (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇* ⓐV. T2) → - 𝐒[T1] → L ⊢ ⬇* ⓐV. T1. -#L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1 -@csn_intro #X #H1 #H2 -elim (cpr_inv_appl1_simple … H1 ?) // -H1 -#V0 #T0 #HLV0 #HLT10 #H destruct -elim (eq_false_inv_tpair_dx … H2) -H2 -[ -IHV -HT1 #HT10 - @(csn_cpr_trans … (ⓐV.T0)) /2 width=1/ -HLV0 - @IHT1 -IHT1 // /2 width=1/ -| -HLT10 * #H #HV0 destruct - @IHV -IHV // -HT1 /2 width=1/ -HV0 - #T2 #HLT02 #HT02 - @(csn_cpr_trans … (ⓐV.T2)) /2 width=1/ -HLV0 - @IHT1 -IHT1 // -HLT02 /2 width=1/ -] -qed. - -(* Main properties **********************************************************) - -theorem csn_acp: acp cpr (eq …) (csn …). -@mk_acp -[ /2 width=1/ -| /2 width=3/ -| /2 width=5/ -| @cnf_lift -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_vector.ma deleted file mode 100644 index bddaaa89d..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/csn_vector.ma +++ /dev/null @@ -1,50 +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/grammar/term_vector.ma". -include "Basic_2/computation/csn.ma". - -(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERM VECTORS **********************) - -inductive csnv (L:lenv): predicate (list term) ≝ -| csnv_nil: csnv L ◊ -| csn_cons: ∀Ts,T. L ⊢ ⬇* T → csnv L Ts → csnv L (T :: Ts) -. - -interpretation - "context-sensitive strong normalization (term vector)" - 'SN L Ts = (csnv L Ts). - -(* Basic inversion lemmas ***************************************************) - -fact csnv_inv_cons_aux: ∀L,Ts. L ⊢ ⬇* Ts → ∀U,Us. Ts = U :: Us → - L ⊢ ⬇* U ∧ L ⊢ ⬇* Us. -#L #Ts * -Ts -[ #U #Us #H destruct -| #Ts #T #HT #HTs #U #Us #H destruct /2 width=1/ -] -qed. - -lemma csnv_inv_cons: ∀L,T,Ts. L ⊢ ⬇* T :: Ts → L ⊢ ⬇* T ∧ L ⊢ ⬇* Ts. -/2 width=3/ qed-. - -include "Basic_2/computation/csn_cpr.ma". - -lemma csn_fwd_applv: ∀L,T,Vs. L ⊢ ⬇* Ⓐ Vs. T → L ⊢ ⬇* Vs ∧ L ⊢ ⬇* T. -#L #T #Vs elim Vs -Vs /2 width=1/ -#V #Vs #IHVs #HVs -lapply (csn_fwd_pair_sn … HVs) #HV -lapply (csn_fwd_flat_dx … HVs) -HVs #HVs -elim (IHVs HVs) -IHVs -HVs /3 width=1/ -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs.ma deleted file mode 100644 index 627a0f58d..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs.ma +++ /dev/null @@ -1,44 +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/reducibility/lcpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************) - -definition lcprs: relation lenv ≝ TC … lcpr. - -interpretation - "context-sensitive parallel computation (environment)" - 'CPRedStar L1 L2 = (lcprs L1 L2). - -(* Basic eliminators ********************************************************) - -lemma lcprs_ind: ∀L1. ∀R:predicate lenv. R L1 → - (∀L,L2. L1 ⊢ ➡* L → L ⊢ ➡ L2 → R L → R L2) → - ∀L2. L1 ⊢ ➡* L2 → R L2. -#L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // -qed-. - -(* Basic properties *********************************************************) - -lemma lcprs_refl: ∀L. L ⊢ ➡* L. -/2 width=1/ qed. - -lemma lcprs_strap1: ∀L1,L,L2. - L1 ⊢ ➡* L → L ⊢ ➡ L2 → L1 ⊢ ➡* L2. -/2 width=3/ qed. - -lemma lcprs_strap2: ∀L1,L,L2. - L1 ⊢ ➡ L → L ⊢ ➡* L2 → L1 ⊢ ➡* L2. -/2 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs_cprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs_cprs.ma deleted file mode 100644 index a45e8ee9f..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs_cprs.ma +++ /dev/null @@ -1,27 +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/reducibility/lcpr_cpr.ma". -include "Basic_2/computation/cprs.ma". -include "Basic_2/computation/lcprs.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************) - -(* Advanced properties ******************************************************) - -lemma lcprs_pair_dx: ∀I,L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. L2 ⊢ V1 ➡* V2 → - L1. ⓑ{I} V1 ⊢ ➡* L2. ⓑ{I} V2. -#I #L1 #L2 #HL12 #V1 #V2 #H @(cprs_ind … H) -V2 -/3 width=1/ /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs_lcprs.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs_lcprs.ma deleted file mode 100644 index 935b83d24..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lcprs_lcprs.ma +++ /dev/null @@ -1,29 +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/computation/lcprs_cprs.ma". - -(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************) - -(* Main properties **********************************************************) - -theorem lcprs_trans: ∀L1,L. L1 ⊢ ➡* L → ∀L2. L ⊢ ➡* L2 → L1 ⊢ ➡* L2. -/2 width=3/ qed. - -lemma lcprs_pair: ∀L1,L2. L1 ⊢ ➡* L2 → ∀V1,V2. L2 ⊢ V1 ➡* V2 → - ∀I. L1. ⓑ{I} V1 ⊢ ➡* L2. ⓑ{I} V2. -#L1 #L2 #H @(lcprs_ind … H) -L2 /2 width=1/ -#L #L2 #_ #HL2 #IHL1 #V1 #V2 #HV12 #I -@(lcprs_trans … (L.ⓑ{I}V1)) /2 width=1/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc.ma deleted file mode 100644 index e61c2081d..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc.ma +++ /dev/null @@ -1,104 +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.ma". -include "Basic_2/computation/acp_cr.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) - -inductive lsubc (RP:lenv→predicate term): relation lenv ≝ -| lsubc_atom: lsubc RP (⋆) (⋆) -| lsubc_pair: ∀I,L1,L2,V. lsubc RP L1 L2 → lsubc RP (L1. ⓑ{I} V) (L2. ⓑ{I} V) -| lsubc_abbr: ∀L1,L2,V,W,A. ⦃L1, V⦄ [RP] ϵ 〚A〛 → L2 ⊢ W ÷ A → - lsubc RP L1 L2 → lsubc RP (L1. ⓓV) (L2. ⓛW) -. - -interpretation - "local environment refinement (abstract candidates of reducibility)" - 'CrSubEq L1 RP L2 = (lsubc RP L1 L2). - -(* Basic inversion lemmas ***************************************************) - -fact lsubc_inv_atom1_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → L1 = ⋆ → L2 = ⋆. -#RP #L1 #L2 * -L1 -L2 -[ // -| #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V #W #A #_ #_ #_ #H destruct -] -qed. - -(* Basic_1: was: csubc_gen_sort_r *) -lemma lsubc_inv_atom1: ∀RP,L2. ⋆ [RP] ⊑ L2 → L2 = ⋆. -/2 width=4/ qed-. - -fact lsubc_inv_pair1_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀I,K1,V. L1 = K1. ⓑ{I} V → - (∃∃K2. K1 [RP] ⊑ K2 & L2 = K2. ⓑ{I} V) ∨ - ∃∃K2,W,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & - K1 [RP] ⊑ K2 & - L2 = K2. ⓛW & I = Abbr. -#RP #L1 #L2 * -L1 -L2 -[ #I #K1 #V #H destruct -| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3/ -| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K1 #V #H destruct /3 width=7/ -] -qed. - -(* Basic_1: was: csubc_gen_head_r *) -lemma lsubc_inv_pair1: ∀RP,I,K1,L2,V. K1. ⓑ{I} V [RP] ⊑ L2 → - (∃∃K2. K1 [RP] ⊑ K2 & L2 = K2. ⓑ{I} V) ∨ - ∃∃K2,W,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & - K1 [RP] ⊑ K2 & - L2 = K2. ⓛW & I = Abbr. -/2 width=3/ qed-. - -fact lsubc_inv_atom2_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → L2 = ⋆ → L1 = ⋆. -#RP #L1 #L2 * -L1 -L2 -[ // -| #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V #W #A #_ #_ #_ #H destruct -] -qed. - -(* Basic_1: was: csubc_gen_sort_l *) -lemma lsubc_inv_atom2: ∀RP,L1. L1 [RP] ⊑ ⋆ → L1 = ⋆. -/2 width=4/ qed-. - -fact lsubc_inv_pair2_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀I,K2,W. L2 = K2. ⓑ{I} W → - (∃∃K1. K1 [RP] ⊑ K2 & L1 = K1. ⓑ{I} W) ∨ - ∃∃K1,V,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & - K1 [RP] ⊑ K2 & - L1 = K1. ⓓV & I = Abst. -#RP #L1 #L2 * -L1 -L2 -[ #I #K2 #W #H destruct -| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3/ -| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K2 #W #H destruct /3 width=7/ -] -qed. - -(* Basic_1: was: csubc_gen_head_l *) -lemma lsubc_inv_pair2: ∀RP,I,L1,K2,W. L1 [RP] ⊑ K2. ⓑ{I} W → - (∃∃K1. K1 [RP] ⊑ K2 & L1 = K1. ⓑ{I} W) ∨ - ∃∃K1,V,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & - K1 [RP] ⊑ K2 & - L1 = K1. ⓓV & I = Abst. -/2 width=3/ qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: csubc_refl *) -lemma lsubc_refl: ∀RP,L. L [RP] ⊑ L. -#RP #L elim L -L // /2 width=1/ -qed. - -(* Basic_1: removed theorems 2: csubc_clear_conf csubc_getl_conf *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_ldrop.ma deleted file mode 100644 index 03ea7b628..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_ldrop.ma +++ /dev/null @@ -1,68 +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/computation/acp_cr.ma". -include "Basic_2/computation/lsubc.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) - -(* Properties concerning basic local environment slicing ********************) - -(* Basic_1: was: csubc_drop_conf_O *) -(* Note: the constant (0) can not be generalized *) -lemma lsubc_ldrop_O1_trans: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → - ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 [RP] ⊑ K2. -#RP #L1 #L2 #H elim H -L1 -L2 -[ #X #e #H - >(ldrop_inv_atom1 … H) -H /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #X #e #H - elim (ldrop_inv_O1 … H) -H * #He #H destruct - [ elim (IHL12 L2 0 ?) -IHL12 // #X #H <(ldrop_inv_refl … H) -H /3 width=3/ - | elim (IHL12 … H) -L2 /3 width=3/ - ] -| #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #X #e #H - elim (ldrop_inv_O1 … H) -H * #He #H destruct - [ elim (IHL12 L2 0 ?) -IHL12 // #X #H <(ldrop_inv_refl … H) -H /3 width=7/ - | elim (IHL12 … H) -L2 /3 width=3/ - ] -qed-. - -(* Basic_1: was: csubc_drop_conf_rev *) -lemma ldrop_lsubc_trans: ∀RR,RS,RP. - acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 [RP] ⊑ K2 → - ∃∃L2. L1 [RP] ⊑ L2 & ⇩[d, e] L2 ≡ K2. -#RR #RS #RP #Hacp #Hacr #L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ #d #e #X #H - >(lsubc_inv_atom1 … H) -H /2 width=3/ -| #L1 #I #V1 #X #H - elim (lsubc_inv_pair1 … H) -H * - [ #K1 #HLK1 #H destruct /3 width=3/ - | #K1 #W1 #A #HV1 #HW1 #HLK1 #H1 #H2 destruct /3 width=3/ - ] -| #L1 #K1 #I #V1 #e #_ #IHLK1 #K2 #HK12 - elim (IHLK1 … HK12) -K1 /3 width=5/ -| #L1 #K1 #I #V1 #V2 #d #e #HLK1 #HV21 #IHLK1 #X #H - elim (lsubc_inv_pair1 … H) -H * - [ #K2 #HK12 #H destruct - elim (IHLK1 … HK12) -K1 /3 width=5/ - | #K2 #W2 #A #HV2 #HW2 #HK12 #H1 #H2 destruct - elim (IHLK1 … HK12) #K #HL1K #HK2 - lapply (aacr_acr … Hacp Hacr A) -Hacp -Hacr #HA - lapply (s7 … HA … HV2 … HLK1 HV21) -HV2 - elim (lift_total W2 d e) /4 width=9/ - ] -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_ldrops.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_ldrops.ma deleted file mode 100644 index 95879a234..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_ldrops.ma +++ /dev/null @@ -1,32 +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/computation/lsubc_ldrop.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) - -(* Properties concerning generic local environment slicing ******************) - -(* Basic_1: was: csubc_drop1_conf_rev *) -lemma ldrops_lsubc_trans: ∀RR,RS,RP. - acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L1,K1,des. ⇩*[des] L1 ≡ K1 → ∀K2. K1 [RP] ⊑ K2 → - ∃∃L2. L1 [RP] ⊑ L2 & ⇩*[des] L2 ≡ K2. -#RR #RS #RP #Hacp #Hacr #L1 #K1 #des #H elim H -L1 -K1 -des -[ /2 width=3/ -| #L1 #L #K1 #des #d #e #_ #HLK1 #IHL #K2 #HK12 - elim (ldrop_lsubc_trans … Hacp Hacr … HLK1 … HK12) -Hacp -Hacr -K1 #K #HLK #HK2 - elim (IHL … HLK) -IHL -HLK /3 width=5/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_lsuba.ma b/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_lsuba.ma deleted file mode 100644 index 41c906285..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/computation/lsubc_lsuba.ma +++ /dev/null @@ -1,26 +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/lsuba.ma". -include "Basic_2/computation/acp_aaa.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) - -(* properties concerning lenv refinement for atomic arity assignment ********) - -lemma lsubc_lsuba: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L1,L2. L1 ÷⊑ L2 → L1 [RP] ⊑ L2. -#RR #RS #RP #H1RP #H2RP #L1 #L2 #H elim H -L1 -L2 -// /2 width=1/ /3 width=4/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/conversion/cpc.ma b/matita/matita/contribs/lambda_delta/Basic_2/conversion/cpc.ma deleted file mode 100644 index e3058349b..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/conversion/cpc.ma +++ /dev/null @@ -1,37 +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/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************) - -definition cpc: lenv → relation term ≝ - λL,T1,T2. L ⊢ T1 ➡ T2 ∨ L ⊢ T2 ➡ T1. - -interpretation - "context-sensitive parallel conversion (term)" - 'PConv L T1 T2 = (cpc L T1 T2). - -(* Basic properties *********************************************************) - -lemma cpc_refl: ∀L,T. L ⊢ T ⬌ T. -/2 width=1/ qed. - -lemma cpc_sym: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → L ⊢ T2 ⬌ T1. -#L #T1 #T2 * /2 width=1/ -qed. - -lemma cpc_cpr: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → ∃∃T. L ⊢ T1 ➡ T & L ⊢ T2 ➡ T. -#L #T1 #T2 * /2 width=3/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/conversion/cpc_cpc.ma b/matita/matita/contribs/lambda_delta/Basic_2/conversion/cpc_cpc.ma deleted file mode 100644 index b4920c35c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/conversion/cpc_cpc.ma +++ /dev/null @@ -1,23 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/conversion/cpc.ma". - -(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************) - -(* Main properties **********************************************************) - -theorem cpc_conf: ∀L,T0,T1,T2. L ⊢ T0 ⬌ T1 → L ⊢ T0 ⬌ T2 → - ∃∃T. L ⊢ T1 ⬌ T & L ⊢ T2 ⬌ T. -/3 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/equivalence/cpcs.ma b/matita/matita/contribs/lambda_delta/Basic_2/equivalence/cpcs.ma deleted file mode 100644 index 6004ee539..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/equivalence/cpcs.ma +++ /dev/null @@ -1,47 +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/conversion/cpc.ma". - -(* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************) - -definition cpcs: lenv → relation term ≝ - λL. TC … (cpc L). - -interpretation "context-sensitive parallel equivalence (term)" - 'PConvStar L T1 T2 = (cpcs L T1 T2). - -(* Basic eliminators ********************************************************) - -lemma cpcs_ind: ∀L,T1. ∀R:predicate term. R T1 → - (∀T,T2. L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → R T → R T2) → - ∀T2. L ⊢ T1 ⬌* T2 → R T2. -#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // -qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: pc3_refl *) -lemma cpcs_refl: ∀L,T. L ⊢ T ⬌* T. -/2 width=1/ qed. - -lemma cpcs_strap1: ∀L,T1,T,T2. - L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → L ⊢ T1 ⬌* T2. -/2 width=3/ qed. - -lemma cpcs_strap2: ∀L,T1,T,T2. - L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. -/2 width=3/ qed. - -(* Basic_1: removed theorems 1: clear_pc3_trans *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma deleted file mode 100644 index 15418582b..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/aarity.ma +++ /dev/null @@ -1,77 +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 *) -(* *) -(**************************************************************************) - -(* THE FORMAL SYSTEM λδ - MATITA SOURCE FILES - * Support for abstract candidates of reducibility closed: 2012 January 27 - * Confluence of context-sensitive parallel reduction closed: 2011 September 21 - * Confluence of context-free parallel reduction closed: 2011 September 6 - * Specification started: 2011 April 17 - * - Patience on me to gain peace and perfection! - - * [ suggested invocation to start formal specifications with ] - *) - -include "Ground_2/star.ma". -include "Basic_2/notation.ma". - -(* ATOMIC ARITY *************************************************************) - -inductive aarity: Type[0] ≝ - | AAtom: aarity (* atomic aarity construction *) - | APair: aarity → aarity → aarity (* binary aarity construction *) -. - -interpretation "aarity construction (atomic)" - 'Item0 = AAtom. - -interpretation "aarity construction (binary)" - 'SnItem2 A1 A2 = (APair A1 A2). - -(* Basic inversion lemmas ***************************************************) - -lemma discr_apair_xy_x: ∀A,B. ②B. A = B → False. -#A #B elim B -B -[ #H destruct -| #Y #X #IHY #_ #H destruct - -H >e0 in e1; normalize (**) (* destruct: one quality is not simplified, the destucted equality is not erased *) - /2 width=1/ -] -qed-. - -lemma discr_tpair_xy_y: ∀B,A. ②B. A = A → False. -#B #A elim A -A -[ #H destruct -| #Y #X #_ #IHX #H destruct - -H (**) (* destruct: the destucted equality is not erased *) - /2 width=1/ -] -qed-. - -(* Basic properties *********************************************************) - -lemma aarity_eq_dec: ∀A1,A2:aarity. Decidable (A1 = A2). -#A1 elim A1 -A1 -[ #A2 elim A2 -A2 /2 width=1/ - #B2 #A2 #_ #_ @or_intror #H destruct -| #B1 #A1 #IHB1 #IHA1 #A2 elim A2 -A2 - [ -IHB1 -IHA1 @or_intror #H destruct - | #B2 #A2 #_ #_ elim (IHB1 B2) -IHB1 - [ #H destruct elim (IHA1 A2) -IHA1 - [ #H destruct /2 width=1/ - | #HA12 @or_intror #H destruct /2 width=1/ - ] - | -IHA1 #HB12 @or_intror #H destruct /2 width=1/ - ] - ] -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_shift.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_shift.ma deleted file mode 100644 index 3a259cecf..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_shift.ma +++ /dev/null @@ -1,24 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/grammar/lenv.ma". - -(* SHIFT OF A CLOSURE *******************************************************) - -let rec shift L T on L ≝ match L with -[ LAtom ⇒ T -| LPair L I V ⇒ shift L (ⓑ{I} V. T) -]. - -interpretation "shift (closure)" 'Append L T = (shift L T). diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma deleted file mode 100644 index 925b84a34..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma +++ /dev/null @@ -1,46 +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/grammar/lenv_weight.ma". -include "Basic_2/grammar/cl_shift.ma". - -(* WEIGHT OF A CLOSURE ******************************************************) - -definition cw: lenv → term → ? ≝ λL,T. #[L] + #[T]. - -interpretation "weight (closure)" 'Weight L T = (cw L T). - -(* Basic properties *********************************************************) - -(* Basic_1: was: flt_wf__q_ind *) - -(* Basic_1: was: flt_wf_ind *) -axiom cw_wf_ind: ∀R:lenv→predicate term. - (∀L2,T2. (∀L1,T1. #[L1,T1] < #[L2,T2] → R L1 T1) → R L2 T2) → - ∀L,T. R L T. - -(* Basic_1: was: flt_shift *) -lemma cw_shift: ∀K,I,V,T. #[K. ⓑ{I} V, T] < #[K, ②{I} V. T]. -normalize // -qed. - -lemma tw_shift: ∀L,T. #[L, T] ≤ #[L @ T]. -#L elim L // -#K #I #V #IHL #T -@transitive_le [3: @IHL |2: /2 width=2/ | skip ] -qed. - -(* Basic_1: removed theorems 6: - flt_thead_sx flt_thead_dx flt_arith0 flt_arith1 flt_arith2 flt_trans -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/genv.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/genv.ma deleted file mode 100644 index 349f8708a..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/genv.ma +++ /dev/null @@ -1,40 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Ground_2/list.ma". -include "Basic_2/grammar/term.ma". - -(* GLOBAL ENVIRONMENTS ******************************************************) - -(* global environments *) -definition genv ≝ list2 bind2 term. - -interpretation "sort (global environment)" - 'Star = (nil2 bind2 term). - -interpretation "environment construction (binary)" - 'DxItem2 L I T = (cons2 bind2 term I T L). - -interpretation "environment binding construction (binary)" - 'DxBind2 L I T = (cons2 bind2 term I T L). - -interpretation "abbreviation (global environment)" - 'DxAbbr L T = (cons2 bind2 term Abbr T L). - -interpretation "abstraction (global environment)" - 'DxAbst L T = (cons2 bind2 term Abst T L). - -(* Basic properties *********************************************************) - -axiom genv_eq_dec: ∀T1,T2:genv. Decidable (T1 = T2). diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/item.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/item.ma deleted file mode 100644 index 7e4ee3839..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/item.ma +++ /dev/null @@ -1,67 +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/arith.ma". -include "Basic_2/notation.ma". - -(* ITEMS ********************************************************************) - -(* atomic items *) -inductive item0: Type[0] ≝ - | Sort: nat → item0 (* sort: starting at 0 *) - | LRef: nat → item0 (* reference by index: starting at 0 *) - | GRef: nat → item0 (* reference by position: starting at 0 *) -. - -(* binary binding items *) -inductive bind2: Type[0] ≝ - | Abbr: bind2 (* abbreviation *) - | Abst: bind2 (* abstraction *) -. - -(* binary non-binding items *) -inductive flat2: Type[0] ≝ - | Appl: flat2 (* application *) - | Cast: flat2 (* explicit type annotation *) -. - -(* binary items *) -inductive item2: Type[0] ≝ - | Bind2: bind2 → item2 (* binding item *) - | Flat2: flat2 → item2 (* non-binding item *) -. - -coercion Bind2. - -coercion Flat2. - -(* Basic properties *********************************************************) - -axiom item0_eq_dec: ∀I1,I2:item0. Decidable (I1 = I2). - -(* Basic_1: was: bind_dec *) -axiom bind2_eq_dec: ∀I1,I2:bind2. Decidable (I1 = I2). - -(* Basic_1: was: flat_dec *) -axiom flat2_eq_dec: ∀I1,I2:flat2. Decidable (I1 = I2). - -(* Basic_1: was: kind_dec *) -axiom item2_eq_dec: ∀I1,I2:item2. Decidable (I1 = I2). - -(* Basic_1: removed theorems 21: - s_S s_plus s_plus_sym s_minus minus_s_s s_le s_lt s_inj s_inc - s_arith0 s_arith1 - r_S r_plus r_plus_sym r_minus r_dis s_r r_arith0 r_arith1 - not_abbr_abst bind_dec_not -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv.ma deleted file mode 100644 index 5aacaf57c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv.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/grammar/term.ma". - -(* LOCAL ENVIRONMENTS *******************************************************) - -(* local environments *) -inductive lenv: Type[0] ≝ -| LAtom: lenv (* empty *) -| LPair: lenv → bind2 → term → lenv (* binary binding construction *) -. - -interpretation "sort (local environment)" - 'Star = LAtom. - -interpretation "environment construction (binary)" - 'DxItem2 L I T = (LPair L I T). - -interpretation "environment binding construction (binary)" - 'DxBind2 L I T = (LPair L I T). - -interpretation "abbreviation (local environment)" - 'DxAbbr L T = (LPair L Abbr T). - -interpretation "abstraction (local environment)" - 'DxAbst L T = (LPair L Abst T). diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv_length.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv_length.ma deleted file mode 100644 index 357301333..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv_length.ma +++ /dev/null @@ -1,24 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/grammar/lenv.ma". - -(* LENGTH OF A LOCAL ENVIRONMENT ********************************************) - -let rec length L ≝ match L with -[ LAtom ⇒ 0 -| LPair L _ _ ⇒ length L + 1 -]. - -interpretation "length (local environment)" 'card L = (length L). diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv_weight.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv_weight.ma deleted file mode 100644 index 414d3f1ab..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lenv_weight.ma +++ /dev/null @@ -1,27 +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/grammar/term_weight.ma". -include "Basic_2/grammar/lenv.ma". - -(* WEIGHT OF A LOCAL ENVIRONMENT ********************************************) - -let rec lw L ≝ match L with -[ LAtom ⇒ 0 -| LPair L _ V ⇒ lw L + #[V] -]. - -interpretation "weight (local environment)" 'Weight L = (lw L). - -(* Basic_1: removed theorems 2: clt_cong clt_head *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lsubs.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/lsubs.ma deleted file mode 100644 index 2c27f0f9e..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/lsubs.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/grammar/lenv_length.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR SUBSTITUTION ****************************) - -inductive lsubs: nat → nat → relation lenv ≝ -| lsubs_sort: ∀d,e. lsubs d e (⋆) (⋆) -| lsubs_OO: ∀L1,L2. lsubs 0 0 L1 L2 -| lsubs_abbr: ∀L1,L2,V,e. lsubs 0 e L1 L2 → - lsubs 0 (e + 1) (L1. ⓓV) (L2.ⓓV) -| lsubs_abst: ∀L1,L2,I,V1,V2,e. lsubs 0 e L1 L2 → - lsubs 0 (e + 1) (L1. ⓛV1) (L2.ⓑ{I} V2) -| lsubs_skip: ∀L1,L2,I1,I2,V1,V2,d,e. - lsubs d e L1 L2 → lsubs (d + 1) e (L1. ⓑ{I1} V1) (L2. ⓑ{I2} V2) -. - -interpretation - "local environment refinement (substitution)" - 'SubEq L1 d e L2 = (lsubs d e L1 L2). - -definition lsubs_conf: ∀S. (lenv → relation S) → Prop ≝ λS,R. - ∀L1,s1,s2. R L1 s1 s2 → - ∀L2,d,e. L1 [d, e] ≼ L2 → R L2 s1 s2. - -(* Basic properties *********************************************************) - -lemma TC_lsubs_conf: ∀S,R. lsubs_conf S R → lsubs_conf S (λL. (TC … (R L))). -#S #R #HR #L1 #s1 #s2 #H elim H -s2 -[ /3 width=5/ -| #s #s2 #_ #Hs2 #IHs1 #L2 #d #e #HL12 - lapply (HR … Hs2 … HL12) -HR -Hs2 -HL12 /3 width=3/ -] -qed. - -lemma lsubs_bind_eq: ∀L1,L2,e. L1 [0, e] ≼ L2 → ∀I,V. - L1. ⓑ{I} V [0, e + 1] ≼ L2.ⓑ{I} V. -#L1 #L2 #e #HL12 #I #V elim I -I /2 width=1/ -qed. - -lemma lsubs_refl: ∀d,e,L. L [d, e] ≼ L. -#d elim d -d -[ #e elim e -e // #e #IHe #L elim L -L // /2 width=1/ -| #d #IHd #e #L elim L -L // /2 width=1/ -] -qed. - -lemma lsubs_skip_lt: ∀L1,L2,d,e. L1 [d - 1, e] ≼ L2 → 0 < d → - ∀I1,I2,V1,V2. L1. ⓑ{I1} V1 [d, e] ≼ L2. ⓑ{I2} V2. - -#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) // /2 width=1/ -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Basic forward lemmas ***************************************************) - -fact lsubs_fwd_length_full1_aux: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → - d = 0 → e = |L1| → |L1| ≤ |L2|. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize -[ // -| /2 width=1/ -| /3 width=1/ -| /3 width=1/ -| #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma lsubs_fwd_length_full1: ∀L1,L2. L1 [0, |L1|] ≼ L2 → |L1| ≤ |L2|. -/2 width=5/ qed-. - -fact lsubs_fwd_length_full2_aux: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → - d = 0 → e = |L2| → |L2| ≤ |L1|. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize -[ // -| /2 width=1/ -| /3 width=1/ -| /3 width=1/ -| #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma lsubs_fwd_length_full2: ∀L1,L2. L1 [0, |L2|] ≼ L2 → |L2| ≤ |L1|. -/2 width=5/ qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma deleted file mode 100644 index 6e8370d49..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma +++ /dev/null @@ -1,117 +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/grammar/item.ma". - -(* TERMS ********************************************************************) - -(* terms *) -inductive term: Type[0] ≝ - | TAtom: item0 → term (* atomic item construction *) - | TPair: item2 → term → term → term (* binary item construction *) -. - -interpretation "term construction (atomic)" - 'Item0 I = (TAtom I). - -interpretation "term construction (binary)" - 'SnItem2 I T1 T2 = (TPair I T1 T2). - -interpretation "term binding construction (binary)" - 'SnBind2 I T1 T2 = (TPair (Bind2 I) T1 T2). - -interpretation "term flat construction (binary)" - 'SnFlat2 I T1 T2 = (TPair (Flat2 I) T1 T2). - -interpretation "sort (term)" - 'Star k = (TAtom (Sort k)). - -interpretation "local reference (term)" - 'LRef i = (TAtom (LRef i)). - -interpretation "global reference (term)" - 'GRef p = (TAtom (GRef p)). - -interpretation "abbreviation (term)" - 'SnAbbr T1 T2 = (TPair (Bind2 Abbr) T1 T2). - -interpretation "abstraction (term)" - 'SnAbst T1 T2 = (TPair (Bind2 Abst) T1 T2). - -interpretation "application (term)" - 'SnAppl T1 T2 = (TPair (Flat2 Appl) T1 T2). - -interpretation "native type annotation (term)" - 'SnCast T1 T2 = (TPair (Flat2 Cast) T1 T2). - -(* Basic properties *********************************************************) - -(* Basic_1: was: term_dec *) -axiom term_eq_dec: ∀T1,T2:term. Decidable (T1 = T2). - -(* Basic inversion lemmas ***************************************************) - -lemma discr_tpair_xy_x: ∀I,T,V. ②{I} V. T = V → False. -#I #T #V elim V -V -[ #J #H destruct -| #J #W #U #IHW #_ #H destruct - -H >e0 in e1; normalize (**) (* destruct: one quality is not simplified, the destucted equality is not erased *) - /2 width=1/ -] -qed-. - -(* Basic_1: was: thead_x_y_y *) -lemma discr_tpair_xy_y: ∀I,V,T. ②{I} V. T = T → False. -#I #V #T elim T -T -[ #J #H destruct -| #J #W #U #_ #IHU #H destruct - -H (**) (* destruct: the destucted equality is not erased *) - /2 width=1/ -] -qed-. - -lemma eq_false_inv_tpair_sn: ∀I,V1,T1,V2,T2. - (②{I} V1. T1 = ②{I} V2. T2 → False) → - (V1 = V2 → False) ∨ (V1 = V2 ∧ (T1 = T2 → False)). -#I #V1 #T1 #V2 #T2 #H -elim (term_eq_dec V1 V2) /3 width=1/ #HV12 destruct -@or_intror @conj // #HT12 destruct /2 width=1/ -qed-. - -lemma eq_false_inv_tpair_dx: ∀I,V1,T1,V2,T2. - (②{I} V1. T1 = ②{I} V2. T2 → False) → - (T1 = T2 → False) ∨ (T1 = T2 ∧ (V1 = V2 → False)). -#I #V1 #T1 #V2 #T2 #H -elim (term_eq_dec T1 T2) /3 width=1/ #HT12 destruct -@or_intror @conj // #HT12 destruct /2 width=1/ -qed-. - -lemma eq_false_inv_beta: ∀V1,V2,W1,W2,T1,T2. - (ⓐV1. ⓛW1. T1 = ⓐV2. ⓛW2 .T2 →False) → - (W1 = W2 → False) ∨ - (W1 = W2 ∧ (ⓓV1. T1 = ⓓV2. T2 → False)). -#V1 #V2 #W1 #W2 #T1 #T2 #H -elim (eq_false_inv_tpair_sn … H) -H -[ #HV12 elim (term_eq_dec W1 W2) /3 width=1/ - #H destruct @or_intror @conj // #H destruct /2 width=1/ -| * #H1 #H2 destruct - elim (eq_false_inv_tpair_sn … H2) -H2 /3 width=1/ - * #H #HT12 destruct - @or_intror @conj // #H destruct /2 width=1/ -] -qed. - -(* Basic_1: removed theorems 3: - not_void_abst not_abbr_void not_abst_void -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma deleted file mode 100644 index 9b6c08270..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma +++ /dev/null @@ -1,41 +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/grammar/term.ma". - -(* SIMPLE (NEUTRAL) TERMS ***************************************************) - -inductive simple: predicate term ≝ - | simple_atom: ∀I. simple (⓪{I}) - | simple_flat: ∀I,V,T. simple (ⓕ{I} V. T) -. - -interpretation "simple (term)" 'Simple T = (simple T). - -(* Basic inversion lemmas ***************************************************) - -fact simple_inv_bind_aux: ∀T. 𝐒[T] → ∀J,W,U. T = ⓑ{J} W. U → False. -#T * -T -[ #I #J #W #U #H destruct -| #I #V #T #J #W #U #H destruct -] -qed. - -lemma simple_inv_bind: ∀I,V,T. 𝐒[ⓑ{I} V. T] → False. -/2 width=6/ qed-. - -lemma simple_inv_pair: ∀I,V,T. 𝐒[②{I}V.T] → ∃J. I = Flat2 J. -* /2 width=2/ #I #V #T #H -elim (simple_inv_bind … H) -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_vector.ma deleted file mode 100644 index 2b3f8880e..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_vector.ma +++ /dev/null @@ -1,27 +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/list.ma". -include "Basic_2/grammar/term.ma". - -(* TERMS ********************************************************************) - -let rec applv Vs T on Vs ≝ - match Vs with - [ nil ⇒ T - | cons hd tl ⇒ ⓐhd. (applv tl T) - ]. - -interpretation "application o vevtor (term)" - 'SnApplV Vs T = (applv Vs T). diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_weight.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_weight.ma deleted file mode 100644 index 1cf414b38..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term_weight.ma +++ /dev/null @@ -1,43 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/grammar/term.ma". - -(* WEIGHT OF A TERM *********************************************************) - -let rec tw T ≝ match T with -[ TAtom _ ⇒ 1 -| TPair _ V T ⇒ tw V + tw T + 1 -]. - -interpretation "weight (term)" 'Weight T = (tw T). - -(* Basic properties *********************************************************) - -(* Basic_1: was: tweight_lt *) -lemma tw_pos: ∀T. 1 ≤ #[T]. -#T elim T -T // -qed. - -(* Basic eliminators ********************************************************) - -axiom tw_wf_ind: ∀R:predicate term. - (∀T2. (∀T1. #[T1] < #[T2] → R T1) → R T2) → - ∀T. R T. - -(* Basic_1: removed theorems 11: - wadd_le wadd_lt wadd_O weight_le weight_eq weight_add_O - weight_add_S tlt_trans tlt_head_sx tlt_head_dx tlt_wf_ind - removed local theorems 1: q_ind -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tshf.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/tshf.ma deleted file mode 100644 index 364a530e1..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tshf.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/grammar/term_simple.ma". - -(* SAME HEAD TERM FORMS *****************************************************) - -inductive tshf: relation term ≝ - | tshf_atom: ∀I. tshf (⓪{I}) (⓪{I}) - | tshf_abst: ∀V1,V2,T1,T2. tshf (ⓛV1. T1) (ⓛV2. T2) - | tshf_appl: ∀V1,V2,T1,T2. tshf T1 T2 → 𝐒[T1] → 𝐒[T2] → - tshf (ⓐV1. T1) (ⓐV2. T2) -. - -interpretation "same head form (term)" 'napart T1 T2 = (tshf T1 T2). - -(* Basic properties *********************************************************) - -lemma tshf_sym: ∀T1,T2. T1 ≈ T2 → T2 ≈ T1. -#T1 #T2 #H elim H -T1 -T2 /2 width=1/ -qed. - -lemma tshf_refl2: ∀T1,T2. T1 ≈ T2 → T2 ≈ T2. -#T1 #T2 #H elim H -T1 -T2 // /2 width=1/ -qed. - -lemma tshf_refl1: ∀T1,T2. T1 ≈ T2 → T1 ≈ T1. -/3 width=2/ qed. - -lemma simple_tshf_repl_dx: ∀T1,T2. T1 ≈ T2 → 𝐒[T1] → 𝐒[T2]. -#T1 #T2 #H elim H -T1 -T2 // -#V1 #V2 #T1 #T2 #H -elim (simple_inv_bind … H) -qed. (**) (* remove from index *) - -lemma simple_tshf_repl_sn: ∀T1,T2. T1 ≈ T2 → 𝐒[T2] → 𝐒[T1]. -/3 width=3/ qed-. - -(* Basic inversion lemmas ***************************************************) - -fact tshf_inv_bind1_aux: ∀T1,T2. T1 ≈ T2 → ∀I,W1,U1. T1 = ⓑ{I}W1.U1 → - ∃∃W2,U2. I = Abst & T2 = ⓛW2. U2. -#T1 #T2 * -T1 -T2 -[ #J #I #W1 #U1 #H destruct -| #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/ -| #V1 #V2 #T1 #T2 #H_ #_ #_ #I #W1 #U1 #H destruct -] -qed. - -lemma tshf_inv_bind1: ∀I,W1,U1,T2. ⓑ{I}W1.U1 ≈ T2 → - ∃∃W2,U2. I = Abst & T2 = ⓛW2. U2. -/2 width=5/ qed-. - -fact tshf_inv_flat1_aux: ∀T1,T2. T1 ≈ T2 → ∀I,W1,U1. T1 = ⓕ{I}W1.U1 → - ∃∃W2,U2. U1 ≈ U2 & 𝐒[U1] & 𝐒[U2] & - I = Appl & T2 = ⓐW2. U2. -#T1 #T2 * -T1 -T2 -[ #J #I #W1 #U1 #H destruct -| #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct -| #V1 #V2 #T1 #T2 #HT12 #HT1 #HT2 #I #W1 #U1 #H destruct /2 width=5/ -] -qed. - -lemma tshf_inv_flat1: ∀I,W1,U1,T2. ⓕ{I}W1.U1 ≈ T2 → - ∃∃W2,U2. U1 ≈ U2 & 𝐒[U1] & 𝐒[U2] & - I = Appl & T2 = ⓐW2. U2. -/2 width=4/ qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc.ma deleted file mode 100644 index c4c44c9f3..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc.ma +++ /dev/null @@ -1,107 +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/grammar/term_simple.ma". - -(* SAME TOP TERM CONSTRUCTOR ************************************************) - -inductive tstc: relation term ≝ - | tstc_atom: ∀I. tstc (⓪{I}) (⓪{I}) - | tstc_pair: ∀I,V1,V2,T1,T2. tstc (②{I} V1. T1) (②{I} V2. T2) -. - -interpretation "same top constructor (term)" 'Iso T1 T2 = (tstc T1 T2). - -(* Basic inversion lemmas ***************************************************) - -fact tstc_inv_atom1_aux: ∀T1,T2. T1 ≃ T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I}. -#T1 #T2 * -T1 -T2 // -#J #V1 #V2 #T1 #T2 #I #H destruct -qed. - -(* Basic_1: was: iso_gen_sort iso_gen_lref *) -lemma tstc_inv_atom1: ∀I,T2. ⓪{I} ≃ T2 → T2 = ⓪{I}. -/2 width=3/ qed-. - -fact tstc_inv_pair1_aux: ∀T1,T2. T1 ≃ T2 → ∀I,W1,U1. T1 = ②{I}W1.U1 → - ∃∃W2,U2. T2 = ②{I}W2. U2. -#T1 #T2 * -T1 -T2 -[ #J #I #W1 #U1 #H destruct -| #J #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/ -] -qed. - -(* Basic_1: was: iso_gen_head *) -lemma tstc_inv_pair1: ∀I,W1,U1,T2. ②{I}W1.U1 ≃ T2 → - ∃∃W2,U2. T2 = ②{I}W2. U2. -/2 width=5/ qed-. - -fact tstc_inv_atom2_aux: ∀T1,T2. T1 ≃ T2 → ∀I. T2 = ⓪{I} → T1 = ⓪{I}. -#T1 #T2 * -T1 -T2 // -#J #V1 #V2 #T1 #T2 #I #H destruct -qed. - -lemma tstc_inv_atom2: ∀I,T1. T1 ≃ ⓪{I} → T1 = ⓪{I}. -/2 width=3/ qed-. - -fact tstc_inv_pair2_aux: ∀T1,T2. T1 ≃ T2 → ∀I,W2,U2. T2 = ②{I}W2.U2 → - ∃∃W1,U1. T1 = ②{I}W1. U1. -#T1 #T2 * -T1 -T2 -[ #J #I #W2 #U2 #H destruct -| #J #V1 #V2 #T1 #T2 #I #W2 #U2 #H destruct /2 width=3/ -] -qed. - -lemma tstc_inv_pair2: ∀I,T1,W2,U2. T1 ≃ ②{I}W2.U2 → - ∃∃W1,U1. T1 = ②{I}W1. U1. -/2 width=5/ qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: iso_refl *) -lemma tstc_refl: ∀T. T ≃ T. -#T elim T -T // -qed. - -lemma tstc_sym: ∀T1,T2. T1 ≃ T2 → T2 ≃ T1. -#T1 #T2 #H elim H -T1 -T2 // -qed. - -lemma tstc_dec: ∀T1,T2. Decidable (T1 ≃ T2). -* #I1 [2: #V1 #T1 ] * #I2 [2,4: #V2 #T2 ] -[ elim (item2_eq_dec I1 I2) #HI12 - [ destruct /2 width=1/ - | @or_intror #H - elim (tstc_inv_pair1 … H) -H #V #T #H destruct /2 width=1/ - ] -| @or_intror #H - lapply (tstc_inv_atom1 … H) -H #H destruct -| @or_intror #H - lapply (tstc_inv_atom2 … H) -H #H destruct -| elim (item0_eq_dec I1 I2) #HI12 - [ destruct /2 width=1/ - | @or_intror #H - lapply (tstc_inv_atom2 … H) -H #H destruct /2 width=1/ - ] -] -qed. - -lemma simple_tstc_repl_dx: ∀T1,T2. T1 ≃ T2 → 𝐒[T1] → 𝐒[T2]. -#T1 #T2 * -T1 -T2 // -#I #V1 #V2 #T1 #T2 #H -elim (simple_inv_pair … H) -H #J #H destruct // -qed. (**) (* remove from index *) - -lemma simple_tstc_repl_sn: ∀T1,T2. T1 ≃ T2 → 𝐒[T2] → 𝐒[T1]. -/3 width=3/ qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc_tstc.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc_tstc.ma deleted file mode 100644 index a32d045b5..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc_tstc.ma +++ /dev/null @@ -1,32 +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/grammar/tstc.ma". - -(* SAME TOP TERM CONSTRUCTOR ************************************************) - -(* Main properties **********************************************************) - -(* Basic_1: was: iso_trans *) -theorem tstc_trans: ∀T1,T. T1 ≃ T → ∀T2. T ≃ T2 → T1 ≃ T2. -#T1 #T * -T1 -T // -#I #V1 #V #T1 #T #X #H -elim (tstc_inv_pair1 … H) -H #V2 #T2 #H destruct // -qed. - -theorem tstc_canc_sn: ∀T,T1. T ≃ T1 → ∀T2. T ≃ T2 → T1 ≃ T2. -/3 width=3/ qed. - -theorem tstc_canc_dx: ∀T1,T. T1 ≃ T → ∀T2. T2 ≃ T → T1 ≃ T2. -/3 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc_vector.ma deleted file mode 100644 index a3b38fcbf..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/tstc_vector.ma +++ /dev/null @@ -1,33 +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/grammar/term_vector.ma". -include "Basic_2/grammar/tstc.ma". - -(* SAME TOP TERM CONSTRUCTOR ************************************************) - -(* Advanced inversion lemmas ************************************************) - -(* Basic_1: was only: iso_flats_lref_bind_false iso_flats_flat_bind_false *) -lemma tstc_inv_bind_appls_simple: ∀I,Vs,V2,T1,T2. ⒶVs.T1 ≃ ⓑ{I} V2. T2 → - 𝐒[T1] → False. -#I #Vs #V2 #T1 #T2 #H -elim (tstc_inv_pair2 … H) -H #V0 #T0 -elim Vs -Vs normalize -[ #H destruct #H - @(simple_inv_bind … H) -| #V #Vs #_ #H destruct -] -qed. - diff --git a/matita/matita/contribs/lambda_delta/Basic_2/names.txt b/matita/matita/contribs/lambda_delta/Basic_2/names.txt deleted file mode 100644 index dfa83ebdd..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/names.txt +++ /dev/null @@ -1,40 +0,0 @@ -NAMING CONVENTIONS FOR METAVARIABLES - -A,B : arity -C,D : candidate of reducibility -E,F : RTM environment -G : global environment -H : reserved: transient premise -IH : reserved: inductive premise -I,J : item -K,L : local environment -M,N : reserved: future use -O : -P,Q : reserved: future use -R : generic predicate (relation) -S : RTM stack -T,U,V,W: term -X,Y,Z : reserved: transient objet denoted by a capital letter - -d : relocation depth -e : relocation height -h : sort hierarchy parameter -i,j : local reference position index (de Bruijn's) -k : sort index -p,q : global reference position -t,u : local reference position level (de Bruijn's) -x,y,z : reserved: transient objet denoted by a small letter - -NAMING CONVENTIONS FOR CONSTRUCTORS - -0: atomic -2: binary - -A: application to vector - -a: application -b: binder -d: abbreviation -f: flat -l: abstraction -t: native type annotation diff --git a/matita/matita/contribs/lambda_delta/Basic_2/notation.ma b/matita/matita/contribs/lambda_delta/Basic_2/notation.ma deleted file mode 100644 index a9fca999c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/notation.ma +++ /dev/null @@ -1,307 +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 λδ ****************************************) - -(* Grammar ******************************************************************) - -notation "hvbox( ⓪ )" - non associative with precedence 90 - for @{ 'Item0 }. - -notation "hvbox( ⓪ { I } )" - non associative with precedence 90 - for @{ 'Item0 $I }. - -notation "hvbox( ⋆ )" - non associative with precedence 90 - for @{ 'Star }. - -notation "hvbox( ⋆ term 90 k )" - non associative with precedence 90 - for @{ 'Star $k }. - -notation "hvbox( # term 90 i )" - non associative with precedence 90 - for @{ 'LRef $i }. - -notation "hvbox( § term 90 p )" - non associative with precedence 90 - for @{ 'GRef $p }. - -notation "hvbox( ② term 90 T1 . break term 90 T )" - non associative with precedence 90 - for @{ 'SnItem2 $T1 $T }. - -notation "hvbox( ② { I } break term 90 T1 . break term 90 T )" - non associative with precedence 90 - for @{ 'SnItem2 $I $T1 $T }. - -notation "hvbox( ⓑ { I } break term 90 T1 . break term 90 T )" - non associative with precedence 90 - for @{ 'SnBind2 $I $T1 $T }. - -notation "hvbox( ⓕ { I } break term 90 T1 . break term 90 T )" - non associative with precedence 90 - for @{ 'SnFlat2 $I $T1 $T }. - -notation "hvbox( ⓓ term 90 T1 . break term 90 T2 )" - non associative with precedence 90 - for @{ 'SnAbbr $T1 $T2 }. - -notation "hvbox( ⓛ term 90 T1 . break term 90 T2 )" - non associative with precedence 90 - for @{ 'SnAbst $T1 $T2 }. - -notation "hvbox( ⓐ term 90 T1 . break term 90 T2 )" - non associative with precedence 90 - for @{ 'SnAppl $T1 $T2 }. - -notation "hvbox( ⓣ term 90 T1 . break term 90 T2 )" - non associative with precedence 90 - for @{ 'SnCast $T1 $T2 }. - -notation "hvbox( Ⓐ term 90 T1 . break term 90 T )" - non associative with precedence 90 - for @{ 'SnApplV $T1 $T }. - -notation > "hvbox( T . break ②{ I } break term 47 T1 )" - non associative with precedence 46 - for @{ 'DxBind2 $T $I $T1 }. - -notation "hvbox( T . break ⓑ { I } break term 90 T1 )" - non associative with precedence 89 - for @{ 'DxBind2 $T $I $T1 }. - -notation "hvbox( T1 . break ⓓ T2 )" - left associative with precedence 48 - for @{ 'DxAbbr $T1 $T2 }. - -notation "hvbox( T1 . break ⓛ T2 )" - left associative with precedence 49 - for @{ 'DxAbst $T1 $T2 }. - -notation "hvbox( T . break ④ { I } break { T1 , break T2 , break T3 } )" - non associative with precedence 47 - for @{ 'DxItem4 $T $I $T1 $T2 $T3 }. - -notation "hvbox( # [ x ] )" - non associative with precedence 90 - for @{ 'Weight $x }. - -notation "hvbox( # [ x , break y ] )" - non associative with precedence 90 - for @{ 'Weight $x $y }. - -notation "hvbox( 𝐒 [ T ] )" - non associative with precedence 45 - for @{ 'Simple $T }. - -notation "hvbox( L ⊢ break term 90 T1 ≈ break T2 )" - non associative with precedence 45 - for @{ 'Hom $L $T1 $T2 }. - -notation "hvbox( T1 ≃ break T2 )" - non associative with precedence 45 - for @{ 'Iso $T1 $T2 }. - -notation "hvbox( T1 break [ d , break e ] ≼ break T2 )" - non associative with precedence 45 - for @{ 'SubEq $T1 $d $e $T2 }. - -(* Substitution *************************************************************) - -notation "hvbox( L ⊢ break [ d , break e ] break 𝐑 [ T ] )" - non associative with precedence 45 - for @{ 'Reducible $L $d $e $T }. - -notation "hvbox( L ⊢ break [ d , break e ] break 𝐈 [ T ] )" - non associative with precedence 45 - for @{ 'NotReducible $L $d $e $T }. - -notation "hvbox( L ⊢ break [ d , break e ] break 𝐍 [ T ] )" - non associative with precedence 45 - for @{ 'Normal $L $d $e $T }. - -notation "hvbox( ⇧ [ d , break e ] break T1 ≡ break T2 )" - non associative with precedence 45 - for @{ 'RLift $d $e $T1 $T2 }. - -notation "hvbox( ⇩ [ e ] break L1 ≡ break L2 )" - non associative with precedence 45 - for @{ 'RDrop $e $L1 $L2 }. - -notation "hvbox( ⇩ [ d , break e ] break L1 ≡ break L2 )" - non associative with precedence 45 - for @{ 'RDrop $d $e $L1 $L2 }. - -notation "hvbox( T1 break [ d , break e ] ▶ break T2 )" - non associative with precedence 45 - for @{ 'PSubst $T1 $d $e $T2 }. - -notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ▶ break T2 )" - non associative with precedence 45 - for @{ 'PSubst $L $T1 $d $e $T2 }. - -(* Unfold *******************************************************************) - -notation "hvbox( @ [ T1 ] break f ≡ break T2 )" - non associative with precedence 45 - for @{ 'RAt $T1 $f $T2 }. - -notation "hvbox( T1 ▭ break T2 ≡ break T )" - non associative with precedence 45 - for @{ 'RMinus $T1 $T2 $T }. - -notation "hvbox( ⇧ * [ e ] break T1 ≡ break T2 )" - non associative with precedence 45 - for @{ 'RLiftStar $e $T1 $T2 }. - -notation "hvbox( ⇩ * [ e ] break L1 ≡ break L2 )" - non associative with precedence 45 - for @{ 'RDropStar $e $L1 $L2 }. - -notation "hvbox( T1 break [ d , break e ] ▶* break T2 )" - non associative with precedence 45 - for @{ 'PSubstStar $T1 $d $e $T2 }. - -notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ▶* break T2 )" - non associative with precedence 45 - for @{ 'PSubstStar $L $T1 $d $e $T2 }. - -notation "hvbox( T1 break [ d , break e ] ≡ break T2 )" - non associative with precedence 45 - for @{ 'TSubst $T1 $d $e $T2 }. - -notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ≡ break T2 )" - non associative with precedence 45 - for @{ 'TSubst $L $T1 $d $e $T2 }. - -(* Static Typing ************************************************************) - -notation "hvbox( L ⊢ break term 90 T ÷ break A )" - non associative with precedence 45 - for @{ 'AtomicArity $L $T $A }. - -notation "hvbox( T1 ÷ ⊑ break T2 )" - non associative with precedence 45 - for @{ 'CrSubEqA $T1 $T2 }. - -(* Reducibility *************************************************************) - -notation "hvbox( 𝐑 [ T ] )" - non associative with precedence 45 - for @{ 'Reducible $T }. - -notation "hvbox( L ⊢ break 𝐑 [ T ] )" - non associative with precedence 45 - for @{ 'Reducible $L $T }. - -notation "hvbox( 𝐈 [ T ] )" - non associative with precedence 45 - for @{ 'NotReducible $T }. - -notation "hvbox( L ⊢ break 𝐈 [ T ] )" - non associative with precedence 45 - for @{ 'NotReducible $L $T }. - -notation "hvbox( 𝐍 [ T ] )" - non associative with precedence 45 - for @{ 'Normal $T }. - -notation "hvbox( L ⊢ break 𝐍 [ T ] )" - non associative with precedence 45 - for @{ 'Normal $L $T }. - -notation "hvbox( 𝐖𝐇𝐑 [ T ] )" - non associative with precedence 45 - for @{ 'WHdReducible $T }. - -notation "hvbox( L ⊢ break 𝐖𝐇𝐑 [ T ] )" - non associative with precedence 45 - for @{ 'WHdReducible $L $T }. - -notation "hvbox( 𝐖𝐇𝐈 [ T ] )" - non associative with precedence 45 - for @{ 'NotWHdReducible $T }. - -notation "hvbox( L ⊢ break 𝐖𝐇𝐈 [ T ] )" - non associative with precedence 45 - for @{ 'NotWHdReducible $L $T }. - -notation "hvbox( 𝐖𝐇𝐍 [ T ] )" - non associative with precedence 45 - for @{ 'WHdNormal $T }. - -notation "hvbox( L ⊢ break 𝐖𝐇𝐍 [ T ] )" - non associative with precedence 45 - for @{ 'WHdNormal $L $T }. - -notation "hvbox( T1 ➡ break T2 )" - non associative with precedence 45 - for @{ 'PRed $T1 $T2 }. - -notation "hvbox( L ⊢ break term 90 T1 ➡ break T2 )" - non associative with precedence 45 - for @{ 'PRed $L $T1 $T2 }. - -notation "hvbox( L1 ⊢ ➡ break L2 )" - non associative with precedence 45 - for @{ 'CPRed $L1 $L2 }. - -(* Computation **************************************************************) - -notation "hvbox( T1 ➡* break T2 )" - non associative with precedence 45 - for @{ 'PRedStar $T1 $T2 }. - -notation "hvbox( L ⊢ break term 90 T1 ➡* break T2 )" - non associative with precedence 45 - for @{ 'PRedStar $L $T1 $T2 }. - -notation "hvbox( L1 ⊢ ➡* break L2 )" - non associative with precedence 45 - for @{ 'CPRedStar $L1 $L2 }. - -notation "hvbox( ⬇ * T )" - non associative with precedence 45 - for @{ 'SN $T }. - -notation "hvbox( L ⊢ ⬇ * T )" - non associative with precedence 45 - for @{ 'SN $L $T }. - -notation "hvbox( L ⊢ ⬇ * * T )" - non associative with precedence 45 - for @{ 'SNStar $L $T }. - -notation "hvbox( ⦃ L, break T ⦄ break [ R ] ϵ break 〚 A 〛 )" - non associative with precedence 45 - for @{ 'InEInt $R $L $T $A }. - -notation "hvbox( T1 break [ R ] ⊑ break T2 )" - non associative with precedence 45 - for @{ 'CrSubEq $T1 $R $T2 }. - -(* Conversion ***************************************************************) - -notation "hvbox( L ⊢ break term 90 T1 ⬌ break T2 )" - non associative with precedence 45 - for @{ 'PConv $L $T1 $T2 }. - -(* Equivalence **************************************************************) - -notation "hvbox( L ⊢ break term 90 T1 ⬌* break T2 )" - non associative with precedence 45 - for @{ 'PConvStar $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cnf.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cnf.ma deleted file mode 100644 index 39ac692cd..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cnf.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/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) - -definition cnf: lenv → predicate term ≝ λL. NF … (cpr L) (eq …). - -interpretation - "context-sensitive normality (term)" - 'Normal L T = (cnf L T). - -(* Basic properties *********************************************************) - -(* Basic_1: was: nf2_sort *) -lemma cnf_sort: ∀L,k. L ⊢ 𝐍[⋆k]. -#L #k #X #H ->(cpr_inv_sort1 … H) // -qed. - -axiom cnf_dec: ∀L,T1. L ⊢ 𝐍[T1] ∨ - ∃∃T2. L ⊢ T1 ➡ T2 & (T1 = T2 → False). - -(* Basic_1: removed theorems 1: nf2_abst_shift *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cnf_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cnf_lift.ma deleted file mode 100644 index b492eb9af..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cnf_lift.ma +++ /dev/null @@ -1,61 +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/reducibility/cpr_lift.ma". -include "Basic_2/reducibility/cnf.ma". - -(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) - -(* Advanced inversion lemmas ************************************************) - -(* Basic_1: was only: nf2_csort_lref *) -lemma cnf_lref_atom: ∀L,i. ⇩[0, i] L ≡ ⋆ → L ⊢ 𝐍[#i]. -#L #i #HLK #X #H -elim (cpr_inv_lref1 … H) // * -#K0 #V0 #V1 #HLK0 #_ #_ #_ -lapply (ldrop_mono … HLK … HLK0) -L #H destruct -qed. - -(* Basic_1: was: nf2_lref_abst *) -lemma cnf_lref_abst: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓛV → L ⊢ 𝐍[#i]. -#L #K #V #i #HLK #X #H -elim (cpr_inv_lref1 … H) // * -#K0 #V0 #V1 #HLK0 #_ #_ #_ -lapply (ldrop_mono … HLK … HLK0) -L #H destruct -qed. - -(* Basic_1: was: nf2_abst *) -lemma cnf_abst: ∀I,L,V,W,T. L ⊢ 𝐍[W] → L. ⓑ{I} V ⊢ 𝐍[T] → L ⊢ 𝐍[ⓛW.T]. -#I #L #V #W #T #HW #HT #X #H -elim (cpr_inv_abst1 … H I V) -H #W0 #T0 #HW0 #HT0 #H destruct ->(HW … HW0) -W >(HT … HT0) -T // -qed. - -(* Basic_1: was only: nf2_appl_lref *) -lemma cnf_appl_simple: ∀L,V,T. L ⊢ 𝐍[V] → L ⊢ 𝐍[T] → 𝐒[T] → L ⊢ 𝐍[ⓐV.T]. -#L #V #T #HV #HT #HS #X #H -elim (cpr_inv_appl1_simple … H ?) -H // #V0 #T0 #HV0 #HT0 #H destruct ->(HV … HV0) -V >(HT … HT0) -T // -qed. - -(* Relocation properties ****************************************************) - -(* Basic_1: was: nf2_lift *) -lemma cnf_lift: ∀L0,L,T,T0,d,e. - L ⊢ 𝐍[T] → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → L0 ⊢ 𝐍[T0]. -#L0 #L #T #T0 #d #e #HLT #HL0 #HT0 #X #H -elim (cpr_inv_lift … HL0 … HT0 … H) -L0 #T1 #HT10 #HT1 ->(HLT … HT1) in HT0; -L #HT0 ->(lift_mono … HT10 … HT0) -T1 -X // -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma deleted file mode 100644 index 0ad09e4b4..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma +++ /dev/null @@ -1,114 +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/grammar/cl_shift.ma". -include "Basic_2/unfold/tpss.ma". -include "Basic_2/reducibility/tpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Basic_1: includes: pr2_delta1 *) -definition cpr: lenv → relation term ≝ - λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T [0, |L|] ▶* T2. - -interpretation - "context-sensitive parallel reduction (term)" - 'PRed L T1 T2 = (cpr L T1 T2). - -(* Basic properties *********************************************************) - -lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 ➡ T2. -/4 width=3/ qed-. - -(* Basic_1: was by definition: pr2_free *) -lemma cpr_pr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. -/2 width=3/ qed. - -lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 ➡ T2. -/3 width=5/ qed. - -lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. -/2 width=1/ qed. - -(* Note: new property *) -(* Basic_1: was only: pr2_thin_dx *) -lemma cpr_flat: ∀I,L,V1,V2,T1,T2. - L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ -qed. - -lemma cpr_cast: ∀L,V,T1,T2. - L ⊢ T1 ➡ T2 → L ⊢ ⓣV. T1 ➡ T2. -#L #V #T1 #T2 * /3 width=3/ -qed. - -(* Note: it does not hold replacing |L1| with |L2| *) -(* Basic_1: was only: pr2_change *) -lemma cpr_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → - ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡ T2. -#L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 -lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 -HL12 /3 width=4/ -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_1: was: pr2_gen_csort *) -lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2. -#T1 #T2 * #T #HT normalize #HT2 -<(tpss_inv_refl_O2 … HT2) -HT2 // -qed-. - -(* Basic_1: was: pr2_gen_sort *) -lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. -#L #T2 #k * #X #H ->(tpr_inv_atom1 … H) -H #H ->(tpss_inv_sort1 … H) -H // -qed-. - -(* Basic_1: was pr2_gen_abbr *) -lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 → - (∃∃V,V2,T2. V1 ➡ V & L ⊢ V [O, |L|] ▶* V2 & - L. ⓓV ⊢ T1 ➡ T2 & - U2 = ⓓV2. T2 - ) ∨ - ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2. -#L #V1 #T1 #Y * #X #H1 #H2 -elim (tpr_inv_abbr1 … H1) -H1 * -[ #V #T0 #T #HV1 #HT10 #HT0 #H destruct - elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct - lapply (tps_lsubs_conf … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0 - lapply (tps_weak_all … HT0) -HT0 #HT0 - lapply (tpss_lsubs_conf … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2 - lapply (tpss_weak_all … HT2) -HT2 #HT2 - lapply (tpss_strap … HT0 HT2) -T /4 width=7/ -| /4 width=5/ -] -qed-. - -(* Basic_1: was: pr2_gen_cast *) -lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓣV1. T1 ➡ U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & - U2 = ⓣV2. T2 - ) ∨ L ⊢ T1 ➡ U2. -#L #V1 #T1 #U2 * #X #H #HU2 -elim (tpr_inv_cast1 … H) -H /3 width=3/ -* #V #T #HV1 #HT1 #H destruct -elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ -qed-. - -(* Basic_1: removed theorems 4: - pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans - Basic_1: removed local theorems 3: - pr2_free_free pr2_free_delta pr2_delta_delta -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_cpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_cpr.ma deleted file mode 100644 index 92db70f0f..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_cpr.ma +++ /dev/null @@ -1,63 +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/reducibility/tpr_tpr.ma". -include "Basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Advanced properties ******************************************************) - -lemma cpr_bind_sn: ∀I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 → - L ⊢ ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12 -@ex2_1_intro [2: @(tpr_delta … HV1 HT12) | skip ] /2 width=3/ (* /3 width=5/ is too slow *) -qed. - -(* Basic_1: was only: pr2_gen_cbind *) -lemma cpr_bind_dx: ∀I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 → - L ⊢ ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2 -elim (tpss_split_up … HT2 1 ? ?) -HT2 // #T0 (tpr_inv_atom1 … H) -H #H -elim (tpss_inv_lref1 … H) -H /2 width=1/ -* /3 width=6/ -qed-. - -(* Basic_1: was: pr2_gen_abst *) -lemma cpr_inv_abst1: ∀L,V1,T1,U2. L ⊢ ⓛV1. T1 ➡ U2 → ∀I,W. - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L. ⓑ{I} W ⊢ T1 ➡ T2 & U2 = ⓛV2. T2. -#L #V1 #T1 #Y * #X #H1 #H2 #I #W -elim (tpr_inv_abst1 … H1) -H1 #V #T #HV1 #HT1 #H destruct -elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct -lapply (tpss_lsubs_conf … HT2 (L. ⓑ{I} W) ?) -HT2 /2 width=1/ /4 width=5/ -qed-. - -(* Basic_1: was pr2_gen_appl *) -lemma cpr_inv_appl1: ∀L,V1,U0,U2. L ⊢ ⓐV1. U0 ➡ U2 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 & - U2 = ⓐV2. T2 - | ∃∃V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 & - U0 = ⓛW. T1 & - U2 = ⓓV2. T2 - | ∃∃V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 & - ⇧[0,1] V2 ≡ V & - U0 = ⓓW1. T1 & - U2 = ⓓW2. ⓐV. T2. -#L #V1 #U0 #Y * #X #H1 #H2 -elim (tpr_inv_appl1 … H1) -H1 * -[ #V #U #HV1 #HU0 #H destruct - elim (tpss_inv_flat1 … H2) -H2 #V2 #U2 #HV2 #HU2 #H destruct /4 width=5/ -| #V #W #T0 #T #HV1 #HT0 #H #H1 destruct - elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct - lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=8/ -| #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct - elim (tpss_inv_bind1 … H2) -H2 #W2 #X #HW02 #HX #HY destruct - elim (tpss_inv_flat1 … HX) -HX #V2 #T2 #HV2 #HT2 #H destruct - elim (tpss_inv_lift1_ge … HV2 … HV0 ?) -V // [3: /2 width=1/ |2: skip ] #V (ltpr_fwd_length … HL12) in HT2; #HT2 -elim (tpr_tpss_ltpr … HL12 … HT2) -L1 /3 width=3/ -qed. - -lemma cpr_ltpr_conf_tpss: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. L1 ⊢ T1 ➡ T2 → - ∀d,e,U1. L1 ⊢ T1 [d, e] ▶* U1 → - ∃∃U2. L2 ⊢ U1 ➡ U2 & L2 ⊢ T2 ➡ U2. -#L1 #L2 #HL12 #T1 #T2 #HT12 #d #e #U1 #HTU1 -elim (cpr_ltpr_conf_eq … HT12 … HL12) -HT12 #T #HT1 #HT2 -elim (cpr_tpss_ltpr … HL12 … HT1 … HTU1) -L1 -HT1 #U2 #HU12 #HTU2 -lapply (tpss_weak_all … HTU2) -HTU2 #HTU2 /3 width=5/ (**) (* /4 width=5/ is too slow *) -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_ltpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_ltpss.ma deleted file mode 100644 index 148a29645..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_ltpss.ma +++ /dev/null @@ -1,27 +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/unfold/ltpss_tpss.ma". -include "Basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Properties concerning partial unfold on local environments ***************) - -lemma ltpss_cpr_trans: ∀L1,L2,d,e. L1 [d, e] ▶* L2 → - ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡ T2. -#L1 #L2 #d #e #HL12 #T1 #T2 * -lapply (ltpss_weak_all … HL12) -<(ltpss_fwd_length … HL12) -HL12 /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr.ma deleted file mode 100644 index 80d14a2e5..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr.ma +++ /dev/null @@ -1,37 +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/unfold/ltpss.ma". -include "Basic_2/reducibility/ltpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS *************) - -definition lcpr: relation lenv ≝ - λL1,L2. ∃∃L. L1 ➡ L & L [0, |L|] ▶* L2 -. - -interpretation - "context-sensitive parallel reduction (environment)" - 'CPRed L1 L2 = (lcpr L1 L2). - -(* Basic properties *********************************************************) - -lemma lcpr_refl: ∀L. L ⊢ ➡ L. -/2 width=3/ qed. - -(* Basic inversion lemmas ***************************************************) - -lemma lcpr_inv_atom1: ∀L2. ⋆ ⊢ ➡ L2 → L2 = ⋆. -#L2 * #L #HL >(ltpr_inv_atom1 … HL) -HL #HL2 >(ltpss_inv_atom1 … HL2) -HL2 // -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr_cpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr_cpr.ma deleted file mode 100644 index 7abe34742..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/lcpr_cpr.ma +++ /dev/null @@ -1,27 +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/unfold/ltpss_ltpss.ma". -include "Basic_2/reducibility/cpr.ma". -include "Basic_2/reducibility/lcpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS *************) - -(* Advanced properties ****************************************************) - -lemma lcpr_pair: ∀L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. L2 ⊢ V1 ➡ V2 → - ∀I. L1. ⓑ{I} V1 ⊢ ➡ L2. ⓑ{I} V2. -#L1 #L2 * #L #HL1 #HL2 #V1 #V2 * -<(ltpss_fwd_length … HL2) /4 width=5/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr.ma deleted file mode 100644 index 8e834b2da..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr.ma +++ /dev/null @@ -1,89 +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/reducibility/tpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -inductive ltpr: relation lenv ≝ -| ltpr_stom: ltpr (⋆) (⋆) -| ltpr_pair: ∀K1,K2,I,V1,V2. - ltpr K1 K2 → V1 ➡ V2 → ltpr (K1. ⓑ{I} V1) (K2. ⓑ{I} V2) (*ⓑ*) -. - -interpretation - "context-free parallel reduction (environment)" - 'PRed L1 L2 = (ltpr L1 L2). - -(* Basic properties *********************************************************) - -lemma ltpr_refl: ∀L:lenv. L ➡ L. -#L elim L -L // /2 width=1/ -qed. - -(* Basic inversion lemmas ***************************************************) - -fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ➡ L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 * -L1 -L2 -[ // -| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct -] -qed. - -(* Basic_1: was: wcpr0_gen_sort *) -lemma ltpr_inv_atom1: ∀L2. ⋆ ➡ L2 → L2 = ⋆. -/2 width=3/ qed-. - -fact ltpr_inv_pair1_aux: ∀L1,L2. L1 ➡ L2 → ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. -#L1 #L2 * -L1 -L2 -[ #K1 #I #V1 #H destruct -| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct /2 width=5/ -] -qed. - -(* Basic_1: was: wcpr0_gen_head *) -lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. ⓑ{I} V1 ➡ L2 → - ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. - -fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ➡ L2 → L2 = ⋆ → L1 = ⋆. -#L1 #L2 * -L1 -L2 -[ // -| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct -] -qed. - -lemma ltpr_inv_atom2: ∀L1. L1 ➡ ⋆ → L1 = ⋆. -/2 width=3/ qed-. - -fact ltpr_inv_pair2_aux: ∀L1,L2. L1 ➡ L2 → ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. -#L1 #L2 * -L1 -L2 -[ #K2 #I #V2 #H destruct -| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct /2 width=5/ -] -qed. - -lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ➡ K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|. -#L1 #L2 #H elim H -L1 -L2 normalize // -qed-. - -(* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr_ldrop.ma deleted file mode 100644 index 090e8cfbf..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr_ldrop.ma +++ /dev/null @@ -1,70 +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/reducibility/tpr_lift.ma". -include "Basic_2/reducibility/ltpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Basic_1: was: wcpr0_drop *) -lemma ltpr_ldrop_conf: ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀L2. L1 ➡ L2 → - ∃∃K2. ⇩[d, e] L2 ≡ K2 & K1 ➡ K2. -#L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ #d #e #X #H >(ltpr_inv_atom1 … H) -H /2 width=3/ -| #K1 #I #V1 #X #H - elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct /3 width=5/ -| #L1 #K1 #I #V1 #e #_ #IHLK1 #X #H - elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct - elim (IHLK1 … HL12) -L1 /3 width=3/ -| #L1 #K1 #I #V1 #W1 #d #e #_ #HWV1 #IHLK1 #X #H - elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct - elim (tpr_inv_lift … HV12 … HWV1) -V1 - elim (IHLK1 … HL12) -L1 /3 width=5/ -] -qed. - -(* Basic_1: was: wcpr0_drop_back *) -lemma ldrop_ltpr_trans: ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 ➡ K2 → - ∃∃L2. ⇩[d, e] L2 ≡ K2 & L1 ➡ L2. -#L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ #d #e #X #H >(ltpr_inv_atom1 … H) -H /2 width=3/ -| #K1 #I #V1 #X #H - elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct /3 width=5/ -| #L1 #K1 #I #V1 #e #_ #IHLK1 #K2 #HK12 - elim (IHLK1 … HK12) -K1 /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d #e #_ #HWV1 #IHLK1 #X #H - elim (ltpr_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (lift_total W2 d e) #V2 #HWV2 - lapply (tpr_lift … HW12 … HWV1 … HWV2) -W1 - elim (IHLK1 … HK12) -K1 /3 width=5/ -] -qed. - -fact ltpr_ldrop_trans_O1_aux: ∀L2,K2,d,e. ⇩[d, e] L2 ≡ K2 → ∀L1. L1 ➡ L2 → - d = 0 → ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 ➡ K2. -#L2 #K2 #d #e #H elim H -L2 -K2 -d -e -[ #d #e #X #H >(ltpr_inv_atom2 … H) -H /2 width=3/ -| #K2 #I #V2 #X #H - elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct /3 width=5/ -| #L2 #K2 #I #V2 #e #_ #IHLK2 #X #H #_ - elim (ltpr_inv_pair2 … H) -H #L1 #V1 #HL12 #HV12 #H destruct - elim (IHLK2 … HL12 ?) -L2 // /3 width=3/ -| #L2 #K2 #I #V2 #W2 #d #e #_ #_ #_ #L1 #_ - >commutative_plus normalize #H destruct -] -qed. - -lemma ltpr_ldrop_trans_O1: ∀L1,L2. L1 ➡ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → - ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 ➡ K2. -/2 width=5/ qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr_tps.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr_tps.ma deleted file mode 100644 index 91b30b014..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr_tps.ma +++ /dev/null @@ -1,55 +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/reducibility/ltpr_ldrop.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properties concerning parallel substitution on terms *********************) - -lemma ltpr_tps_trans: ∀L2,T1,T2,d,e. L2 ⊢ T1 [d, e] ▶ T2 → ∀L1. L1 ➡ L2 → - ∃∃T. L1 ⊢ T1 [d, e] ▶ T & T ➡ T2. -#L2 #T1 #T2 #d #e #H elim H -L2 -T1 -T2 -d -e -[ /2 width=3/ -| #L2 #K2 #V2 #W2 #i #d #e #Hdi #Hide #HLK2 #HVW2 #L1 #HL12 - elim (ltpr_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H - elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct -K2 - elim (lift_total V1 0 (i+1)) #W1 #HVW1 - lapply (tpr_lift … HV12 … HVW1 … HVW2) -V2 /3 width=4/ -| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 - elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 - elim (IHT12 (L1.ⓑ{I}V) ?) /2 width=1/ -L2 /3 width=5/ -| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 - elim (IHV12 … HL12) -IHV12 - elim (IHT12 … HL12) -L2 /3 width=5/ -] -qed. - -lemma ltpr_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 → ∀L2. L1 ➡ L2 → - ∃∃T. L2 ⊢ T1 [d, e] ▶ T & T2 ➡ T. -#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e -[ /2 width=3/ -| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L2 #HL12 - elim (ltpr_ldrop_conf … HLK1 … HL12) -L1 #X #HLK2 #H - elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct -K1 - elim (lift_total V2 0 (i+1)) #W2 #HVW2 - lapply (tpr_lift … HV12 … HVW1 … HVW2) -V1 /3 width=4/ -| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 - elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 - elim (IHT12 (L2.ⓑ{I}V) ?) /2 width=1/ -L1 /3 width=5/ -| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 - elim (IHV12 … HL12) -IHV12 - elim (IHT12 … HL12) -L1 /3 width=5/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tif.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tif.ma deleted file mode 100644 index 5b89755a7..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tif.ma +++ /dev/null @@ -1,54 +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/reducibility/trf.ma". - -(* CONTEXT-FREE IRREDUCIBLE TERMS *******************************************) - -definition tif: predicate term ≝ λT. 𝐑[T] → False. - -interpretation "context-free irreducibility (term)" - 'NotReducible T = (tif T). - -(* Basic inversion lemmas ***************************************************) - -lemma tif_inv_abbr: ∀V,T. 𝐈[ⓓV.T] → False. -/2 width=1/ qed-. - -lemma tif_inv_abst: ∀V,T. 𝐈[ⓛV.T] → 𝐈[V] ∧ 𝐈[T]. -/4 width=1/ qed-. - -lemma tif_inv_appl: ∀V,T. 𝐈[ⓐV.T] → ∧∧ 𝐈[V] & 𝐈[T] & 𝐒[T]. -#V #T #HVT @and3_intro /3 width=1/ -generalize in match HVT; -HVT elim T -T // -* // * #U #T #_ #_ #H elim (H ?) -H /2 width=1/ -qed-. - -lemma tif_inv_cast: ∀V,T. 𝐈[ⓣV.T] → False. -/2 width=1/ qed-. - -(* Basic properties *********************************************************) - -lemma tif_atom: ∀I. 𝐈[⓪{I}]. -/2 width=4/ qed. - -lemma tif_abst: ∀V,T. 𝐈[V] → 𝐈[T] → 𝐈[ⓛV.T]. -#V #T #HV #HT #H -elim (trf_inv_abst … H) -H /2 width=1/ -qed. - -lemma tif_appl: ∀V,T. 𝐈[V] → 𝐈[T] → 𝐒[T] → 𝐈[ⓐV.T]. -#V #T #HV #HT #S #H -elim (trf_inv_appl … H) -H /2 width=1/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf.ma deleted file mode 100644 index 5a7205cc0..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf.ma +++ /dev/null @@ -1,59 +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/reducibility/tpr.ma". - -(* CONTEXT-FREE NORMAL TERMS ************************************************) - -definition tnf: predicate term ≝ NF … tpr (eq …). - -interpretation - "context-free normality (term)" - 'Normal T = (tnf T). - -(* Basic inversion lemmas ***************************************************) - -lemma tnf_inv_abst: ∀V,T. 𝐍[ⓛV.T] → 𝐍[V] ∧ 𝐍[T]. -#V1 #T1 #HVT1 @conj -[ #V2 #HV2 lapply (HVT1 (ⓛV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct // -| #T2 #HT2 lapply (HVT1 (ⓛV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct // -] -qed-. - -lemma tnf_inv_appl: ∀V,T. 𝐍[ⓐV.T] → ∧∧ 𝐍[V] & 𝐍[T] & 𝐒[T]. -#V1 #T1 #HVT1 @and3_intro -[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct // -| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct // -| generalize in match HVT1; -HVT1 elim T1 -T1 * // * #W1 #U1 #_ #_ #H - [ elim (lift_total V1 0 1) #V2 #HV12 - lapply (H (ⓓW1.ⓐV2.U1) ?) -H /2 width=3/ -HV12 #H destruct - | lapply (H (ⓓV1.U1) ?) -H /2 width=1/ #H destruct -] -qed-. - -lemma tnf_inv_abbr: ∀V,T. 𝐍[ⓓV.T] → False. -#V #T #H elim (is_lift_dec T 0 1) -[ * #U #HTU - lapply (H U ?) -H /2 width=3/ #H destruct - elim (lift_inv_pair_xy_y … HTU) -| #HT - elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12 - lapply (H (ⓓV.T2) ?) -H /2 width=3/ -HT2 #H destruct /3 width=2/ -] -qed. - -lemma tnf_inv_cast: ∀V,T. 𝐍[ⓣV.T] → False. -#V #T #H lapply (H T ?) -H /2 width=1/ #H -@(discr_tpair_xy_y … H) -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf_tif.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf_tif.ma deleted file mode 100644 index 2888ff9d6..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf_tif.ma +++ /dev/null @@ -1,74 +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/tps_lift.ma". -include "Basic_2/reducibility/tif.ma". -include "Basic_2/reducibility/tnf.ma". - -(* CONTEXT-FREE NORMAL TERMS ************************************************) - -(* Main properties properties ***********************************************) - -lemma tpr_tif_eq: ∀T1,T2. T1 ➡ T2 → 𝐈[T1] → T1 = T2. -#T1 #T2 #H elim H -T1 -T2 -[ // -| * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H - [ elim (tif_inv_appl … H) -H #HV1 #HT1 #_ - >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // - | elim (tif_inv_cast … H) - ] -| #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H - elim (tif_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -| * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #H - [ -HT2 -IHV1 -IHT1 elim (tif_inv_abbr … H) - | <(tps_inv_refl_SO2 … HT2 ?) -HT2 // - elim (tif_inv_abst … H) -H #HV1 #HT1 - >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // - ] -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H - elim (tif_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -| #V1 #T1 #T2 #T #_ #_ #_ #H - elim (tif_inv_abbr … H) -| #V1 #T1 #T #_ #_ #H - elim (tif_inv_cast … H) -] -qed. - -theorem tif_tnf: ∀T1. 𝐈[T1] → 𝐍[T1]. -/2 width=1/ qed. - -(* Note: this property is unusual *) -lemma tnf_trf_false: ∀T1. 𝐑[T1] → 𝐍[T1] → False. -#T1 #H elim H -T1 -[ #V #T #_ #IHV #H elim (tnf_inv_abst … H) -H /2 width=1/ -| #V #T #_ #IHT #H elim (tnf_inv_abst … H) -H /2 width=1/ -| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/ -| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/ -| #V #T #H elim (tnf_inv_abbr … H) -| #V #T #H elim (tnf_inv_cast … H) -| #V #W #T #H elim (tnf_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -] -qed. - -theorem tnf_tif: ∀T1. 𝐍[T1] → 𝐈[T1]. -/2 width=3/ qed. - -lemma tnf_abst: ∀V,T. 𝐍[V] → 𝐍[T] → 𝐍[ⓛV.T]. -/4 width=1/ qed. - -lemma tnf_appl: ∀V,T. 𝐍[V] → 𝐍[T] → 𝐒[T] → 𝐍[ⓐV.T]. -/4 width=1/ qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr.ma deleted file mode 100644 index 0edd15211..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr.ma +++ /dev/null @@ -1,209 +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/tps.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Basic_1: includes: pr0_delta1 *) -inductive tpr: relation term ≝ -| tpr_atom : ∀I. tpr (⓪{I}) (⓪{I}) -| tpr_flat : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 → - tpr (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) -| tpr_beta : ∀V1,V2,W,T1,T2. - tpr V1 V2 → tpr T1 T2 → tpr (ⓐV1. ⓛW. T1) (ⓓV2. T2) -| tpr_delta: ∀I,V1,V2,T1,T2,T. - tpr V1 V2 → tpr T1 T2 → ⋆. ⓑ{I} V2 ⊢ T2 [0, 1] ▶ T → - tpr (ⓑ{I} V1. T1) (ⓑ{I} V2. T) -| tpr_theta: ∀V,V1,V2,W1,W2,T1,T2. - tpr V1 V2 → ⇧[0,1] V2 ≡ V → tpr W1 W2 → tpr T1 T2 → - tpr (ⓐV1. ⓓW1. T1) (ⓓW2. ⓐV. T2) -| tpr_zeta : ∀V,T,T1,T2. ⇧[0,1] T1 ≡ T → tpr T1 T2 → tpr (ⓓV. T) T2 -| tpr_tau : ∀V,T1,T2. tpr T1 T2 → tpr (ⓣV. T1) T2 -. - -interpretation - "context-free parallel reduction (term)" - 'PRed T1 T2 = (tpr T1 T2). - -(* Basic properties *********************************************************) - -lemma tpr_bind: ∀I,V1,V2,T1,T2. V1 ➡ V2 → T1 ➡ T2 → - ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2. -/2 width=3/ qed. - -(* Basic_1: was by definition: pr0_refl *) -lemma tpr_refl: ∀T. T ➡ T. -#T elim T -T // -#I elim I -I /2 width=1/ -qed. - -(* Basic inversion lemmas ***************************************************) - -fact tpr_inv_atom1_aux: ∀U1,U2. U1 ➡ U2 → ∀I. U1 = ⓪{I} → U2 = ⓪{I}. -#U1 #U2 * -U1 -U2 -[ // -| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #T #_ #_ #_ #k #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct -| #V #T #T1 #T2 #_ #_ #k #H destruct -| #V #T1 #T2 #_ #k #H destruct -] -qed. - -(* Basic_1: was: pr0_gen_sort pr0_gen_lref *) -lemma tpr_inv_atom1: ∀I,U2. ⓪{I} ➡ U2 → U2 = ⓪{I}. -/2 width=3/ qed-. - -fact tpr_inv_bind1_aux: ∀U1,U2. U1 ➡ U2 → ∀I,V1,T1. U1 = ⓑ{I} V1. T1 → - (∃∃V2,T2,T. V1 ➡ V2 & T1 ➡ T2 & - ⋆. ⓑ{I} V2 ⊢ T2 [0, 1] ▶ T & - U2 = ⓑ{I} V2. T - ) ∨ - ∃∃T. ⇧[0,1] T ≡ T1 & T ➡ U2 & I = Abbr. -#U1 #U2 * -U1 -U2 -[ #J #I #V #T #H destruct -| #I1 #V1 #V2 #T1 #T2 #_ #_ #I #V #T #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #I #V #T #H destruct -| #I1 #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #I0 #V0 #T0 #H destruct /3 width=7/ -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #I0 #V0 #T0 #H destruct -| #V #T #T1 #T2 #HT1 #HT12 #I0 #V0 #T0 #H destruct /3 width=3/ -| #V #T1 #T2 #_ #I0 #V0 #T0 #H destruct -] -qed. - -lemma tpr_inv_bind1: ∀V1,T1,U2,I. ⓑ{I} V1. T1 ➡ U2 → - (∃∃V2,T2,T. V1 ➡ V2 & T1 ➡ T2 & - ⋆. ⓑ{I} V2 ⊢ T2 [0, 1] ▶ T & - U2 = ⓑ{I} V2. T - ) ∨ - ∃∃T. ⇧[0,1] T ≡ T1 & T ➡ U2 & I = Abbr. -/2 width=3/ qed-. - -(* Basic_1: was pr0_gen_abbr *) -lemma tpr_inv_abbr1: ∀V1,T1,U2. ⓓV1. T1 ➡ U2 → - (∃∃V2,T2,T. V1 ➡ V2 & T1 ➡ T2 & - ⋆. ⓓV2 ⊢ T2 [0, 1] ▶ T & - U2 = ⓓV2. T - ) ∨ - ∃∃T. ⇧[0,1] T ≡ T1 & T ➡ U2. -#V1 #T1 #U2 #H -elim (tpr_inv_bind1 … H) -H * /3 width=7/ -qed-. - -fact tpr_inv_flat1_aux: ∀U1,U2. U1 ➡ U2 → ∀I,V1,U0. U1 = ⓕ{I} V1. U0 → - ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 & - U2 = ⓕ{I} V2. T2 - | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 & - U0 = ⓛW. T1 & - U2 = ⓓV2. T2 & I = Appl - | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 & - ⇧[0,1] V2 ≡ V & - U0 = ⓓW1. T1 & - U2 = ⓓW2. ⓐV. T2 & - I = Appl - | (U0 ➡ U2 ∧ I = Cast). -#U1 #U2 * -U1 -U2 -[ #I #J #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=5/ -| #V1 #V2 #W #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=8/ -| #I #V1 #V2 #T1 #T2 #T #_ #_ #_ #J #V0 #T0 #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #J #V0 #T0 #H destruct /3 width=12/ -| #V #T #T1 #T2 #_ #_ #J #V0 #T0 #H destruct -| #V #T1 #T2 #HT12 #J #V0 #T0 #H destruct /3 width=1/ -] -qed. - -lemma tpr_inv_flat1: ∀V1,U0,U2,I. ⓕ{I} V1. U0 ➡ U2 → - ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 & - U2 = ⓕ{I} V2. T2 - | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 & - U0 = ⓛW. T1 & - U2 = ⓓV2. T2 & I = Appl - | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 & - ⇧[0,1] V2 ≡ V & - U0 = ⓓW1. T1 & - U2 = ⓓW2. ⓐV. T2 & - I = Appl - | (U0 ➡ U2 ∧ I = Cast). -/2 width=3/ qed-. - -(* Basic_1: was pr0_gen_appl *) -lemma tpr_inv_appl1: ∀V1,U0,U2. ⓐV1. U0 ➡ U2 → - ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 & - U2 = ⓐV2. T2 - | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 & - U0 = ⓛW. T1 & - U2 = ⓓV2. T2 - | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 & - ⇧[0,1] V2 ≡ V & - U0 = ⓓW1. T1 & - U2 = ⓓW2. ⓐV. T2. -#V1 #U0 #U2 #H -elim (tpr_inv_flat1 … H) -H * /3 width=12/ #_ #H destruct -qed-. - -(* Note: the main property of simple terms *) -lemma tpr_inv_appl1_simple: ∀V1,T1,U. ⓐV1. T1 ➡ U → 𝐒[T1] → - ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & - U = ⓐV2. T2. -#V1 #T1 #U #H #HT1 -elim (tpr_inv_appl1 … H) -H * -[ /2 width=5/ -| #V2 #W #W1 #W2 #_ #_ #H #_ destruct - elim (simple_inv_bind … HT1) -| #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct - elim (simple_inv_bind … HT1) -] -qed-. - -(* Basic_1: was: pr0_gen_cast *) -lemma tpr_inv_cast1: ∀V1,T1,U2. ⓣV1. T1 ➡ U2 → - (∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓣV2. T2) - ∨ T1 ➡ U2. -#V1 #T1 #U2 #H -elim (tpr_inv_flat1 … H) -H * /3 width=5/ -[ #V2 #W #W1 #W2 #_ #_ #_ #_ #H destruct -| #V2 #W #W1 #W2 #T2 #U1 #_ #_ #_ #_ #_ #_ #H destruct -] -qed-. - -fact tpr_inv_lref2_aux: ∀T1,T2. T1 ➡ T2 → ∀i. T2 = #i → - ∨∨ T1 = #i - | ∃∃V,T,T0. ⇧[O,1] T0 ≡ T & T0 ➡ #i & - T1 = ⓓV. T - | ∃∃V,T. T ➡ #i & T1 = ⓣV. T. -#T1 #T2 * -T1 -T2 -[ #I #i #H destruct /2 width=1/ -| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct -| #I #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct -| #V #T #T1 #T2 #HT1 #HT12 #i #H destruct /3 width=6/ -| #V #T1 #T2 #HT12 #i #H destruct /3 width=4/ -] -qed. - -lemma tpr_inv_lref2: ∀T1,i. T1 ➡ #i → - ∨∨ T1 = #i - | ∃∃V,T,T0. ⇧[O,1] T0 ≡ T & T0 ➡ #i & - T1 = ⓓV. T - | ∃∃V,T. T ➡ #i & T1 = ⓣV. T. -/2 width=3/ qed-. - -(* Basic_1: removed theorems 3: - pr0_subst0_back pr0_subst0_fwd pr0_subst0 - Basic_1: removed local theorems: 1: pr0_delta_tau -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_lift.ma deleted file mode 100644 index be40639a8..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_lift.ma +++ /dev/null @@ -1,121 +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/tps_lift.ma". -include "Basic_2/reducibility/tpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Relocation properties ****************************************************) - -(* Basic_1: was: pr0_lift *) -lemma tpr_lift: ∀T1,T2. T1 ➡ T2 → - ∀d,e,U1. ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → U1 ➡ U2. -#T1 #T2 #H elim H -T1 -T2 -[ * #i #d #e #U1 #HU1 #U2 #HU2 - lapply (lift_mono … HU1 … HU2) -HU1 #H destruct - [ lapply (lift_inv_sort1 … HU2) -HU2 #H destruct // - | lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct // - | lapply (lift_inv_gref1 … HU2) -HU2 #H destruct // - ] -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct - elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct /3 width=4/ -| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #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 #V3 #T3 #HV23 #HT23 #HX2 destruct /3 width=4/ -| #I #V1 #V2 #T1 #T2 #T0 #HV12 #HT12 #HT2 #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 - elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct - elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct - elim (lift_total T2 (d + 1) e) #U2 #HTU2 - @tpr_delta - [4: @(tps_lift_le … HT2 … HTU2 HTU0 ?) /2 width=1/ |1: skip |2: /2 width=4/ |3: /2 width=4/ ] (**) (*/3. is too slow *) -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #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 - elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct - elim (lift_trans_ge … HV2 … HV3 ?) -V // /3 width=4/ -| #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #X #HX #T0 #HT20 - elim (lift_inv_bind1 … HX) -HX #V3 #T3 #_ #HT3 #HX destruct - elim (lift_trans_ge … HT1 … HT3 ?) -T // /3 width=6/ -| #V #T1 #T2 #_ #IHT12 #d #e #X #HX #T #HT2 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #_ #HT0 #HX destruct /3 width=4/ -] -qed. - -(* Basic_1: was: pr0_gen_lift *) -lemma tpr_inv_lift: ∀T1,T2. T1 ➡ T2 → - ∀d,e,U1. ⇧[d, e] U1 ≡ T1 → - ∃∃U2. ⇧[d, e] U2 ≡ T2 & U1 ➡ U2. -#T1 #T2 #H elim H -T1 -T2 -[ * #i #d #e #U1 #HU1 - [ lapply (lift_inv_sort2 … HU1) -HU1 #H destruct /2 width=3/ - | lapply (lift_inv_lref2 … HU1) -HU1 * * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … HU1) -HU1 #H destruct /2 width=3/ - ] -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct - elim (IHV12 … HV01) -V1 - elim (IHT12 … HT01) -T1 /3 width=5/ -| #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct - elim (IHV12 … HV01) -V1 - elim (IHT12 … HT01) -T1 /3 width=5/ -| #I #V1 #V2 #T1 #T2 #T0 #_ #_ #HT20 #IHV12 #IHT12 #d #e #X #HX - elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct - elim (IHV12 … HWV1) -V1 #W2 #HWV2 #HW12 - elim (IHT12 … HUT1) -T1 #U2 #HUT2 #HU12 - elim (tps_inv_lift1_le … HT20 … HUT2 ?) -T2 // [3: /2 width=5/ |2: skip ] #U0 #HU20 #HUT0 - @ex2_1_intro [2: /2 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /3 width=5/ is slow *) -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct - elim (IHV12 … HV01) -V1 #V3 #HV32 #HV03 - elim (IHW12 … HW01) -W1 #W3 #HW32 #HW03 - elim (IHT12 … HT01) -T1 #T3 #HT32 #HT03 - elim (lift_trans_le … HV32 … HV2 ?) -V2 // #V2 #HV32 #HV2 - @ex2_1_intro [2: /3 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /4 width=5/ is slow *) -| #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #X #HX - elim (lift_inv_bind2 … HX) -HX #V0 #T0 #_ #HT0 #H destruct - elim (lift_div_le … HT1 … HT0 ?) -T // #T #HT0 #HT1 - elim (IHT12 … HT1) -T1 /3 width=5/ -| #V #T1 #T2 #_ #IHT12 #d #e #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct - elim (IHT12 … HT01) -T1 /3 width=3/ -] -qed. - -(* Advanced inversion lemmas ************************************************) - -fact tpr_inv_abst1_aux: ∀U1,U2. U1 ➡ U2 → ∀V1,T1. U1 = ⓛV1. T1 → - ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2. -#U1 #U2 * -U1 -U2 -[ #I #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct -| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct -| #I #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #V0 #T0 #H destruct - <(tps_inv_refl_SO2 … HT2 ? ? ?) -T /2 width=5/ -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct -| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct -| #V #T1 #T2 #_ #V0 #T0 #H destruct -] -qed. - -(* Basic_1: was pr0_gen_abst *) -lemma tpr_inv_abst1: ∀V1,T1,U2. ⓛV1. T1 ➡ U2 → - ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2. -/2 width=3/ qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_tpr.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_tpr.ma deleted file mode 100644 index b1e3f3c1b..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_tpr.ma +++ /dev/null @@ -1,287 +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/reducibility/tpr_tpss.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Confluence lemmas ********************************************************) - -fact tpr_conf_atom_atom: ∀I. ∃∃X. ⓪{I} ➡ X & ⓪{I} ➡ X. -/2 width=3/ qed. - -fact tpr_conf_flat_flat: - ∀I,V0,V1,T0,T1,V2,T2. ( - ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → - ∃∃T0. ⓕ{I} V1. T1 ➡ T0 & ⓕ{I} V2. T2 ➡ T0. -#I #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 -elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=5/ -qed. - -fact tpr_conf_flat_beta: - ∀V0,V1,T1,V2,W0,U0,T2. ( - ∀X0:term. #[X0] < #[V0] + (#[W0] + #[U0] + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → - U0 ➡ T2 → ⓛW0. U0 ➡ T1 → - ∃∃X. ⓐV1. T1 ➡ X & ⓓV2. T2 ➡ X. -#V0 #V1 #T1 #V2 #W0 #U0 #T2 #IH #HV01 #HV02 #HT02 #H -elim (tpr_inv_abst1 … H) -H #W1 #U1 #HW01 #HU01 #H destruct -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 -elim (IH … HT02 … HU01) -HT02 -HU01 -IH /2 width=1/ /3 width=5/ -qed. - -(* basic-1: was: - pr0_cong_upsilon_refl pr0_cong_upsilon_zeta - pr0_cong_upsilon_cong pr0_cong_upsilon_delta -*) -fact tpr_conf_flat_theta: - ∀V0,V1,T1,V2,V,W0,W2,U0,U2. ( - ∀X0:term. #[X0] < #[V0] + (#[W0] + #[U0] + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → ⇧[O,1] V2 ≡ V → - W0 ➡ W2 → U0 ➡ U2 → ⓓW0. U0 ➡ T1 → - ∃∃X. ⓐV1. T1 ➡ X & ⓓW2. ⓐV. U2 ➡ X. -#V0 #V1 #T1 #V2 #V #W0 #W2 #U0 #U2 #IH #HV01 #HV02 #HV2 #HW02 #HU02 #H -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #VV #HVV1 #HVV2 -elim (lift_total VV 0 1) #VVV #HVV -lapply (tpr_lift … HVV2 … HV2 … HVV) #HVVV -elim (tpr_inv_abbr1 … H) -H * -(* case 1: delta *) -[ -HV2 -HVV2 #WW2 #UU2 #UU #HWW2 #HUU02 #HUU2 #H destruct - elim (IH … HW02 … HWW2) -HW02 -HWW2 /2 width=1/ #W #HW02 #HWW2 - elim (IH … HU02 … HUU02) -HU02 -HUU02 -IH /2 width=1/ #U #HU2 #HUUU2 - elim (tpr_tps_bind … HWW2 HUUU2 … HUU2) -UU2 #UUU #HUUU2 #HUUU1 - @ex2_1_intro - [2: @tpr_theta [6: @HVV |7: @HWW2 |8: @HUUU2 |1,2,3,4: skip | // ] - |1:skip - |3: @tpr_delta [3: @tpr_flat |1: skip ] /2 width=5/ - ] (**) (* /5 width=14/ is too slow *) -(* case 3: zeta *) -| -HW02 -HVV -HVVV #UU1 #HUU10 #HUUT1 - elim (tpr_inv_lift … HU02 … HUU10) -HU02 #UU #HUU2 #HUU1 - lapply (tw_lift … HUU10) -HUU10 #HUU10 - elim (IH … HUUT1 … HUU1) -HUUT1 -HUU1 -IH /2 width=1/ -HUU10 #U #HU2 #HUUU2 - @ex2_1_intro - [2: @tpr_flat - |1: skip - |3: @tpr_zeta [2: @lift_flat |1: skip |3: @tpr_flat ] - ] /2 width=5/ (**) (* /5 width=5/ is too slow *) -] -qed. - -fact tpr_conf_flat_cast: - ∀X2,V0,V1,T0,T1. ( - ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → T0 ➡ T1 → T0 ➡ X2 → - ∃∃X. ⓣV1. T1 ➡ X & X2 ➡ X. -#X2 #V0 #V1 #T0 #T1 #IH #_ #HT01 #HT02 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=3/ -qed. - -fact tpr_conf_beta_beta: - ∀W0:term. ∀V0,V1,T0,T1,V2,T2. ( - ∀X0:term. #[X0] < #[V0] + (#[W0] + #[T0] + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → - ∃∃X. ⓓV1. T1 ➡X & ⓓV2. T2 ➡ X. -#W0 #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ -elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ /3 width=5/ -qed. - -(* Basic_1: was: pr0_cong_delta pr0_delta_delta *) -fact tpr_conf_delta_delta: - ∀I1,V0,V1,T0,T1,TT1,V2,T2,TT2. ( - ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → - ⋆. ⓑ{I1} V1 ⊢ T1 [O, 1] ▶ TT1 → - ⋆. ⓑ{I1} V2 ⊢ T2 [O, 1] ▶ TT2 → - ∃∃X. ⓑ{I1} V1. TT1 ➡ X & ⓑ{I1} V2. TT2 ➡ X. -#I1 #V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2 -elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // #T #HT1 #HT2 -elim (tpr_tps_bind … HV1 HT1 … HTT1) -T1 #U1 #TTU1 #HTU1 -elim (tpr_tps_bind … HV2 HT2 … HTT2) -T2 #U2 #TTU2 #HTU2 -elim (tps_conf_eq … HTU1 … HTU2) -T #U #HU1 #HU2 -@ex2_1_intro [2,3: @tpr_delta |1: skip ] /width=10/ (**) (* /3 width=10/ is too slow *) -qed. - -fact tpr_conf_delta_zeta: - ∀X2,V0,V1,T0,T1,TT1,T2. ( - ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → T0 ➡ T1 → ⋆. ⓓV1 ⊢ T1 [O,1] ▶ TT1 → - T2 ➡ X2 → ⇧[O, 1] T2 ≡ T0 → - ∃∃X. ⓓV1. TT1 ➡ X & X2 ➡ X. -#X2 #V0 #V1 #T0 #T1 #TT1 #T2 #IH #_ #HT01 #HTT1 #HTX2 #HTT20 -elim (tpr_inv_lift … HT01 … HTT20) -HT01 #TT2 #HTT21 #HTT2 -lapply (tps_inv_lift1_eq … HTT1 … HTT21) -HTT1 #HTT1 destruct -lapply (tw_lift … HTT20) -HTT20 #HTT20 -elim (IH … HTX2 … HTT2) -HTX2 -HTT2 -IH // /3 width=3/ -qed. - -(* Basic_1: was: pr0_upsilon_upsilon *) -fact tpr_conf_theta_theta: - ∀VV1,V0,V1,W0,W1,T0,T1,V2,VV2,W2,T2. ( - ∀X0:term. #[X0] < #[V0] + (#[W0] + #[T0] + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → W0 ➡ W1 → W0 ➡ W2 → T0 ➡ T1 → T0 ➡ T2 → - ⇧[O, 1] V1 ≡ VV1 → ⇧[O, 1] V2 ≡ VV2 → - ∃∃X. ⓓW1. ⓐVV1. T1 ➡ X & ⓓW2. ⓐVV2. T2 ➡ X. -#VV1 #V0 #V1 #W0 #W1 #T0 #T1 #V2 #VV2 #W2 #T2 #IH #HV01 #HV02 #HW01 #HW02 #HT01 #HT02 #HVV1 #HVV2 -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 -elim (IH … HW01 … HW02) -HW01 -HW02 /2 width=1/ #W #HW1 #HW2 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ #T #HT1 #HT2 -elim (lift_total V 0 1) #VV #HVV -lapply (tpr_lift … HV1 … HVV1 … HVV) -V1 #HVV1 -lapply (tpr_lift … HV2 … HVV2 … HVV) -V2 -HVV #HVV2 -@ex2_1_intro [2,3: @tpr_bind |1:skip ] /2 width=5/ (**) (* /4 width=5/ is too slow *) -qed. - -fact tpr_conf_zeta_zeta: - ∀V0:term. ∀X2,TT0,T0,T1,T2. ( - ∀X0:term. #[X0] < #[V0] + #[TT0] + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - T0 ➡ T1 → T2 ➡ X2 → - ⇧[O, 1] T0 ≡ TT0 → ⇧[O, 1] T2 ≡ TT0 → - ∃∃X. T1 ➡ X & X2 ➡ X. -#V0 #X2 #TT0 #T0 #T1 #T2 #IH #HT01 #HTX2 #HTT0 #HTT20 -lapply (lift_inj … HTT0 … HTT20) -HTT0 #H destruct -lapply (tw_lift … HTT20) -HTT20 #HTT20 -elim (IH … HT01 … HTX2) -HT01 -HTX2 -IH // /2 width=3/ -qed. - -fact tpr_conf_tau_tau: - ∀V0,T0:term. ∀X2,T1. ( - ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - T0 ➡ T1 → T0 ➡ X2 → - ∃∃X. T1 ➡ X & X2 ➡ X. -#V0 #T0 #X2 #T1 #IH #HT01 #HT02 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /2 width=3/ -qed. - -(* Confluence ***************************************************************) - -fact tpr_conf_aux: - ∀Y0:term. ( - ∀X0:term. #[X0] < #[Y0] → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - ∀X0,X1,X2. X0 ➡ X1 → X0 ➡ X2 → X0 = Y0 → - ∃∃X. X1 ➡ X & X2 ➡ X. -#Y0 #IH #X0 #X1 #X2 * -X0 -X1 -[ #I1 #H1 #H2 destruct - lapply (tpr_inv_atom1 … H1) -H1 -(* case 1: atom, atom *) - #H1 destruct // -| #I #V0 #V1 #T0 #T1 #HV01 #HT01 #H1 #H2 destruct - elim (tpr_inv_flat1 … H1) -H1 * -(* case 2: flat, flat *) - [ #V2 #T2 #HV02 #HT02 #H destruct - /3 width=7 by tpr_conf_flat_flat/ (**) (* /3 width=7/ is too slow *) -(* case 3: flat, beta *) - | #V2 #W #U0 #T2 #HV02 #HT02 #H1 #H2 #H3 destruct - /3 width=8 by tpr_conf_flat_beta/ (**) (* /3 width=8/ is too slow *) -(* case 4: flat, theta *) - | #V2 #V #W0 #W2 #U0 #U2 #HV02 #HW02 #HT02 #HV2 #H1 #H2 #H3 destruct - /3 width=11 by tpr_conf_flat_theta/ (**) (* /3 width=11/ is too slow *) -(* case 5: flat, tau *) - | #HT02 #H destruct - /3 width=6 by tpr_conf_flat_cast/ (**) (* /3 width=6/ is too slow *) - ] -| #V0 #V1 #W0 #T0 #T1 #HV01 #HT01 #H1 #H2 destruct - elim (tpr_inv_appl1 … H1) -H1 * -(* case 6: beta, flat (repeated) *) - [ #V2 #T2 #HV02 #HT02 #H destruct - @ex2_1_comm /3 width=8 by tpr_conf_flat_beta/ -(* case 7: beta, beta *) - | #V2 #WW0 #TT0 #T2 #HV02 #HT02 #H1 #H2 destruct - /3 width=8 by tpr_conf_beta_beta/ (**) (* /3 width=8/ is too slow *) -(* case 8, beta, theta (excluded) *) - | #V2 #VV2 #WW0 #W2 #TT0 #T2 #_ #_ #_ #_ #H destruct - ] -| #I1 #V0 #V1 #T0 #T1 #TT1 #HV01 #HT01 #HTT1 #H1 #H2 destruct - elim (tpr_inv_bind1 … H1) -H1 * -(* case 9: delta, delta *) - [ #V2 #T2 #TT2 #HV02 #HT02 #HTT2 #H destruct - /3 width=11 by tpr_conf_delta_delta/ (**) (* /3 width=11/ is too slow *) -(* case 10: delta, zata *) - | #T2 #HT20 #HTX2 #H destruct - /3 width=10 by tpr_conf_delta_zeta/ (**) (* /3 width=10/ is too slow *) - ] -| #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #HV01 #HVV1 #HW01 #HT01 #H1 #H2 destruct - elim (tpr_inv_appl1 … H1) -H1 * -(* case 11: theta, flat (repeated) *) - [ #V2 #T2 #HV02 #HT02 #H destruct - @ex2_1_comm /3 width=11 by tpr_conf_flat_theta/ -(* case 12: theta, beta (repeated) *) - | #V2 #WW0 #TT0 #T2 #_ #_ #H destruct -(* case 13: theta, theta *) - | #V2 #VV2 #WW0 #W2 #TT0 #T2 #V02 #HW02 #HT02 #HVV2 #H1 #H2 destruct - /3 width=14 by tpr_conf_theta_theta/ (**) (* /3 width=14/ is too slow *) - ] -| #V0 #TT0 #T0 #T1 #HTT0 #HT01 #H1 #H2 destruct - elim (tpr_inv_abbr1 … H1) -H1 * -(* case 14: zeta, delta (repeated) *) - [ #V2 #T2 #TT2 #HV02 #HT02 #HTT2 #H destruct - @ex2_1_comm /3 width=10 by tpr_conf_delta_zeta/ -(* case 15: zeta, zeta *) - | #T2 #HTT20 #HTX2 - /3 width=9 by tpr_conf_zeta_zeta/ (**) (* /3 width=9/ is too slow *) - ] -| #V0 #T0 #T1 #HT01 #H1 #H2 destruct - elim (tpr_inv_cast1 … H1) -H1 -(* case 16: tau, flat (repeated) *) - [ * #V2 #T2 #HV02 #HT02 #H destruct - @ex2_1_comm /3 width=6 by tpr_conf_flat_cast/ -(* case 17: tau, tau *) - | #HT02 - /3 width=5 by tpr_conf_tau_tau/ - ] -] -qed. - -(* Basic_1: was: pr0_confluence *) -theorem tpr_conf: ∀T0:term. ∀T1,T2. T0 ➡ T1 → T0 ➡ T2 → - ∃∃T. T1 ➡ T & T2 ➡ T. -#T @(tw_wf_ind … T) -T /3 width=6 by tpr_conf_aux/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_tpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_tpss.ma deleted file mode 100644 index cf52701b5..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/tpr_tpss.ma +++ /dev/null @@ -1,94 +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/unfold/ltpss_ltpss.ma". -include "Basic_2/reducibility/ltpr_ldrop.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Unfold properties ********************************************************) - -(* Basic_1: was: pr0_subst1 *) -lemma tpr_tps_ltpr: ∀T1,T2. T1 ➡ T2 → - ∀L1,d,e,U1. L1 ⊢ T1 [d, e] ▶ U1 → - ∀L2. L1 ➡ L2 → - ∃∃U2. U1 ➡ U2 & L2 ⊢ T2 [d, e] ▶* U2. -#T1 #T2 #H elim H -T1 -T2 -[ #I #L1 #d #e #X #H - elim (tps_inv_atom1 … H) -H - [ #H destruct /2 width=3/ - | * #K1 #V1 #i #Hdi #Hide #HLK1 #HVU1 #H #L2 #HL12 destruct - elim (ltpr_ldrop_conf … HLK1 … HL12) -L1 #X #HLK2 #H - elim (ltpr_inv_pair1 … H) -H #K2 #V2 #_ #HV12 #H destruct - elim (lift_total V2 0 (i+1)) #U2 #HVU2 - lapply (tpr_lift … HV12 … HVU1 … HVU2) -V1 #HU12 - @ex2_1_intro [2: @HU12 | skip | /3 width=4/ ] (**) (* /4 width=6/ is too slow *) - ] -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_flat1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct - elim (IHV12 … HVW1 … HL12) -V1 - elim (IHT12 … HTU1 … HL12) -T1 -HL12 /3 width=5/ -| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct - elim (tps_inv_bind1 … HY) -HY #WW #TT1 #_ #HTT1 #H destruct - elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 - elim (IHT12 … HTT1 (L2. ⓛWW) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 - lapply (tpss_lsubs_conf … HTT2 (L2. ⓓVV2) ?) -HTT2 /3 width=5/ -| #I #V1 #V2 #T1 #T2 #U2 #HV12 #_ #HTU2 #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct - elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 - elim (IHT12 … HTT1 (L2. ⓑ{I} VV2) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 - elim (tpss_strip_neq … HTT2 … HTU2 ?) -T2 /2 width=1/ #T2 #HTT2 #HUT2 - lapply (tps_lsubs_conf … HTT2 (L2. ⓑ{I} V2) ?) -HTT2 /2 width=1/ #HTT2 - elim (ltpss_tps_conf … HTT2 (L2. ⓑ{I} VV2) (d + 1) e ?) -HTT2 /2 width=1/ #W2 #HTTW2 #HTW2 - lapply (tpss_lsubs_conf … HTTW2 (⋆. ⓑ{I} VV2) ?) -HTTW2 /2 width=1/ #HTTW2 - lapply (tpss_tps … HTTW2) -HTTW2 #HTTW2 - lapply (tpss_lsubs_conf … HTW2 (L2. ⓑ{I} VV2) ?) -HTW2 /2 width=1/ #HTW2 - lapply (tpss_trans_eq … HUT2 … HTW2) -T2 /3 width=5/ -| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct - elim (tps_inv_bind1 … HY) -HY #WW1 #TT1 #HWW1 #HTT1 #H destruct - elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 - elim (IHW12 … HWW1 … HL12) -W1 #WW2 #HWW12 #HWW2 - elim (IHT12 … HTT1 (L2. ⓓWW2) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 - elim (lift_total VV2 0 1) #VV #H2VV - lapply (tpss_lift_ge … HVV2 (L2. ⓓWW2) … HV2 … H2VV) -V2 /2 width=1/ #HVV - @ex2_1_intro [2: @tpr_theta |1: skip |3: @tpss_bind [2: @tpss_flat ] ] /width=11/ (**) (* /4 width=11/ is too slow *) -| #V1 #TT1 #T1 #T2 #HTT1 #_ #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_bind1 … H) -H #V2 #TT2 #HV12 #HTT12 #H destruct - elim (tps_inv_lift1_ge … HTT12 L1 … HTT1 ?) -TT1 /2 width=1/ #T2 #HT12 #HTT2 - elim (IHT12 … HT12 … HL12) -T1 -HL12 (aaa_inv_sort … H) -H // -| #I1 #L #K1 #V1 #B #i #HLK1 #_ #IHA1 #A2 #H - elim (aaa_inv_lref … H) -H #I2 #K2 #V2 #HLK2 #HA2 - lapply (ldrop_mono … HLK1 … HLK2) -L #H destruct /2 width=1/ -| #L #V #T #B1 #A1 #_ #_ #_ #IHA1 #A2 #H - elim (aaa_inv_abbr … H) -H /2 width=1/ -| #L #V1 #T1 #B1 #A1 #_ #_ #IHB1 #IHA1 #X #H - elim (aaa_inv_abst … H) -H #B2 #A2 #HB2 #HA2 #H destruct /3 width=1/ -| #L #V1 #T1 #B1 #A1 #_ #_ #_ #IHA1 #A2 #H - elim (aaa_inv_appl … H) -H #B2 #_ #HA2 - lapply (IHA1 … HA2) -L #H destruct // -| #L #V #T #A1 #_ #_ #_ #IHA1 #A2 #H - elim (aaa_inv_cast … H) -H /2 width=1/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/aaa_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/aaa_lift.ma deleted file mode 100644 index edf08147e..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/static/aaa_lift.ma +++ /dev/null @@ -1,47 +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/ldrop_ldrop.ma". -include "Basic_2/static/aaa.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties concerning basic relocation ***********************************) - -lemma aaa_lift: ∀L1,T1,A. L1 ⊢ T1 ÷ A → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → - ∀T2. ⇧[d, e] T1 ≡ T2 → L2 ⊢ T2 ÷ A. -#L1 #T1 #A #H elim H -L1 -T1 -A -[ #L1 #k #L2 #d #e #_ #T2 #H - >(lift_inv_sort1 … H) -H // -| #I #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #d #e #HL21 #T2 #H - elim (lift_inv_lref1 … H) -H * #Hid #H destruct - [ 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 (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ - ] -| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H - elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /4 width=4/ -| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H - elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /4 width=4/ -| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H - elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /3 width=4/ -| #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #d #e #HL21 #X #H - elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - /3 width=4/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/aaa_lifts.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/aaa_lifts.ma deleted file mode 100644 index 156e75c9d..000000000 --- a/matita/matita/contribs/lambda_delta/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/unfold/ldrops.ma". -include "Basic_2/static/aaa_lift.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties concerning generic relocation *********************************) - -lemma aaa_lifts: ∀L1,L2,T2,A,des. ⇩*[des] L2 ≡ L1 → ∀T1. ⇧*[des] T1 ≡ T2 → - L1 ⊢ T1 ÷ A → L2 ⊢ T2 ÷ A. -#L1 #L2 #T2 #A #des #H elim H -L1 -L2 -des -[ #L #T1 #H #HT1 - <(lifts_inv_nil … H) -H // -| #L1 #L #L2 #des #d #e #_ #HL2 #IHL1 #T1 #H #HT1 - elim (lifts_inv_cons … H) -H /3 width=9/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba.ma deleted file mode 100644 index 5d087c93c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba.ma +++ /dev/null @@ -1,92 +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.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) - -inductive lsuba: relation lenv ≝ -| lsuba_atom: lsuba (⋆) (⋆) -| lsuba_pair: ∀I,L1,L2,V. lsuba L1 L2 → lsuba (L1. ⓑ{I} V) (L2. ⓑ{I} V) -| lsuba_abbr: ∀L1,L2,V,W,A. L1 ⊢ V ÷ A → L2 ⊢ W ÷ A → - lsuba L1 L2 → lsuba (L1. ⓓV) (L2. ⓛW) -. - -interpretation - "local environment refinement (atomic arity assigment)" - 'CrSubEqA L1 L2 = (lsuba L1 L2). - -(* Basic inversion lemmas ***************************************************) - -fact lsuba_inv_atom1_aux: ∀L1,L2. L1 ÷⊑ L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 * -L1 -L2 -[ // -| #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V #W #A #_ #_ #_ #H destruct -] -qed. - -lemma lsuba_inv_atom1: ∀L2. ⋆ ÷⊑ L2 → L2 = ⋆. -/2 width=3/ qed-. - -fact lsuba_inv_pair1_aux: ∀L1,L2. L1 ÷⊑ L2 → ∀I,K1,V. L1 = K1. ⓑ{I} V → - (∃∃K2. K1 ÷⊑ K2 & L2 = K2. ⓑ{I} V) ∨ - ∃∃K2,W,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & - L2 = K2. ⓛW & I = Abbr. -#L1 #L2 * -L1 -L2 -[ #I #K1 #V #H destruct -| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3/ -| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K1 #V #H destruct /3 width=7/ -] -qed. - -lemma lsuba_inv_pair1: ∀I,K1,L2,V. K1. ⓑ{I} V ÷⊑ L2 → - (∃∃K2. K1 ÷⊑ K2 & L2 = K2. ⓑ{I} V) ∨ - ∃∃K2,W,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & - L2 = K2. ⓛW & I = Abbr. -/2 width=3/ qed-. - -fact lsuba_inv_atom2_aux: ∀L1,L2. L1 ÷⊑ L2 → L2 = ⋆ → L1 = ⋆. -#L1 #L2 * -L1 -L2 -[ // -| #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V #W #A #_ #_ #_ #H destruct -] -qed. - -lemma lsubc_inv_atom2: ∀L1. L1 ÷⊑ ⋆ → L1 = ⋆. -/2 width=3/ qed-. - -fact lsuba_inv_pair2_aux: ∀L1,L2. L1 ÷⊑ L2 → ∀I,K2,W. L2 = K2. ⓑ{I} W → - (∃∃K1. K1 ÷⊑ K2 & L1 = K1. ⓑ{I} W) ∨ - ∃∃K1,V,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & - L1 = K1. ⓓV & I = Abst. -#L1 #L2 * -L1 -L2 -[ #I #K2 #W #H destruct -| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3/ -| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K2 #W #H destruct /3 width=7/ -] -qed. - -lemma lsuba_inv_pair2: ∀I,L1,K2,W. L1 ÷⊑ K2. ⓑ{I} W → - (∃∃K1. K1 ÷⊑ K2 & L1 = K1. ⓑ{I} W) ∨ - ∃∃K1,V,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & - L1 = K1. ⓓV & I = Abst. -/2 width=3/ qed-. - -(* Basic properties *********************************************************) - -lemma lsuba_refl: ∀L. L ÷⊑ L. -#L elim L -L // /2 width=1/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_aaa.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_aaa.ma deleted file mode 100644 index a6a9c3fae..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_aaa.ma +++ /dev/null @@ -1,54 +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_aaa.ma". -include "Basic_2/static/lsuba_ldrop.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) - -(* Properties concerning atomic arity assignment ****************************) - -lemma lsuba_aaa_conf: ∀L1,V,A. L1 ⊢ V ÷ A → ∀L2. L1 ÷⊑ L2 → L2 ⊢ V ÷ A. -#L1 #V #A #H elim H -L1 -V -A -[ // -| #I #L1 #K1 #V1 #B #i #HLK1 #HV1B #IHV1 #L2 #HL12 - elim (lsuba_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2 - elim (lsuba_inv_pair1 … H) -H * #K2 - [ #HK12 #H destruct /3 width=5/ - | #V2 #A1 #HV1A1 #HV2 #_ #H1 #H2 destruct - >(aaa_mono … HV1B … HV1A1) -B -HV1A1 /2 width=5/ - ] -| /4 width=2/ -| /4 width=1/ -| /3 width=3/ -| /3 width=1/ -] -qed. - -lemma lsuba_aaa_trans: ∀L2,V,A. L2 ⊢ V ÷ A → ∀L1. L1 ÷⊑ L2 → L1 ⊢ V ÷ A. -#L2 #V #A #H elim H -L2 -V -A -[ // -| #I #L2 #K2 #V2 #B #i #HLK2 #HV2B #IHV2 #L1 #HL12 - elim (lsuba_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 - elim (lsuba_inv_pair2 … H) -H * #K1 - [ #HK12 #H destruct /3 width=5/ - | #V1 #A1 #HV1 #HV2A1 #_ #H1 #H2 destruct - >(aaa_mono … HV2B … HV2A1) -B -HV2A1 /2 width=5/ - ] -| /4 width=2/ -| /4 width=1/ -| /3 width=3/ -| /3 width=1/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_ldrop.ma deleted file mode 100644 index 522252ae1..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_ldrop.ma +++ /dev/null @@ -1,63 +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/lsuba.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) - -(* Properties concerning basic local environment slicing ********************) - -(* Note: the constant 0 cannot be generalized *) - -lemma lsuba_ldrop_O1_conf: ∀L1,L2. L1 ÷⊑ L2 → ∀K1,e. ⇩[0, e] L1 ≡ K1 → - ∃∃K2. K1 ÷⊑ K2 & ⇩[0, e] L2 ≡ K2. -#L1 #L2 #H elim H -L1 -L2 -[ /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 - [ destruct - elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ - | elim (IHL12 … HLK1) -L1 /3 width=3/ - ] -| #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #K1 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK1 - [ destruct - elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ - | elim (IHL12 … HLK1) -L1 /3 width=3/ - ] -] -qed-. - -lemma lsuba_ldrop_O1_trans: ∀L1,L2. L1 ÷⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → - ∃∃K1. K1 ÷⊑ K2 & ⇩[0, e] L1 ≡ K1. -#L1 #L2 #H elim H -L1 -L2 -[ /2 width=3/ -| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 - [ destruct - elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ - | elim (IHL12 … HLK2) -L2 /3 width=3/ - ] -| #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #K2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #HLK2 - [ destruct - elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H - <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ - | elim (IHL12 … HLK2) -L2 /3 width=3/ - ] -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_lsuba.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_lsuba.ma deleted file mode 100644 index 24da36ccd..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/static/lsuba_lsuba.ma +++ /dev/null @@ -1,35 +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/lsuba_aaa.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) - -(* Main properties **********************************************************) - -theorem lsuba_trans: ∀L1,L. L1 ÷⊑ L → ∀L2. L ÷⊑ L2 → L1 ÷⊑ L2. -#L1 #L #H elim H -L1 -L -[ #X #H >(lsuba_inv_atom1 … H) -H // -| #I #L1 #L #V #HL1 #IHL1 #X #H - elim (lsuba_inv_pair1 … H) -H * #L2 - [ #HL2 #H destruct /3 width=1/ - | #V #A #HLV #HL2V #HL2 #H1 #H2 destruct /3 width=3/ - ] -| #L1 #L #V1 #W #A1 #HV1 #HW #HL1 #IHL1 #X #H - elim (lsuba_inv_pair1 … H) -H * #L2 - [ #HL2 #H destruct /3 width=5/ - | #V #A2 #_ #_ #_ #_ #H destruct - ] -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/static/sh.ma b/matita/matita/contribs/lambda_delta/Basic_2/static/sh.ma deleted file mode 100644 index 7edbfd2af..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/static/sh.ma +++ /dev/null @@ -1,23 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Ground_2/arith.ma". - -(* SORT HIERARCHY ***********************************************************) - -(* sort hierarchy specifications *) -record sh: Type[0] ≝ { - next: nat → nat; (* next sort in the hierarchy *) - next_lt: ∀k. k < next k (* strict monotonicity condition *) -}. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/gdrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/gdrop.ma deleted file mode 100644 index eb4bd4865..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/gdrop.ma +++ /dev/null @@ -1,80 +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/grammar/genv.ma". - -(* GLOBAL ENVIRONMENT SLICING ***********************************************) - -inductive gdrop (e:nat): relation genv ≝ -| gdrop_gt: ∀G. |G| ≤ e → gdrop e G (⋆) -| gdrop_eq: ∀G. |G| = e + 1 → gdrop e G G -| gdrop_lt: ∀I,G1,G2,V. e < |G1| → gdrop e G1 G2 → gdrop e (G1. ⓑ{I} V) G2 -. - -interpretation "global slicing" - 'RDrop e G1 G2 = (gdrop e G1 G2). - -(* basic inversion lemmas ***************************************************) - -lemma gdrop_inv_gt: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| ≤ e → G2 = ⋆. -#G1 #G2 #e * -G1 -G2 // -[ #G #H >H -H >commutative_plus #H - lapply (le_plus_to_le_r … 0 H) -H #H - lapply (le_n_O_to_eq … H) -H #H destruct -| #I #G1 #G2 #V #H1 #_ #H2 - lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 normalize in ⊢ (? % ? → ?); >commutative_plus #H - lapply (lt_plus_to_lt_l … 0 H) -H #H - elim (lt_zero_false … H) -] -qed-. - -lemma gdrop_inv_eq: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| = e + 1 → G1 = G2. -#G1 #G2 #e * -G1 -G2 // -[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H - lapply (le_plus_to_le_r … 0 H) -H #H - lapply (le_n_O_to_eq … H) -H #H destruct -| #I #G1 #G2 #V #H1 #_ normalize #H2 - <(injective_plus_l … H2) in H1; -H2 #H - elim (lt_refl_false … H) -] -qed-. - -fact gdrop_inv_lt_aux: ∀I,G,G1,G2,V,e. ⇩[e] G ≡ G2 → G = G1. ⓑ{I} V → - e < |G1| → ⇩[e] G1 ≡ G2. -#I #G #G1 #G2 #V #e * -G -G2 -[ #G #H1 #H destruct #H2 - lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H - lapply (lt_plus_to_lt_l … 0 H) -H #H - elim (lt_zero_false … H) -| #G #H1 #H2 destruct >(injective_plus_l … H1) -H1 #H - elim (lt_refl_false … H) -| #J #G #G2 #W #_ #HG2 #H destruct // -] -qed. - -lemma gdrop_inv_lt: ∀I,G1,G2,V,e. - ⇩[e] G1. ⓑ{I} V ≡ G2 → e < |G1| → ⇩[e] G1 ≡ G2. -/2 width=5/ qed-. - -(* Basic properties *********************************************************) - -lemma gdrop_total: ∀e,G1. ∃G2. ⇩[e] G1 ≡ G2. -#e #G1 elim G1 -G1 /3 width=2/ -#I #V #G1 * #G2 #HG12 -elim (lt_or_eq_or_gt e (|G1|)) #He -[ /3 width=2/ -| destruct /3 width=2/ -| @ex_intro [2: @gdrop_gt normalize /2 width=1/ | skip ] (**) (* explicit constructor *) -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/gdrop_gdrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/gdrop_gdrop.ma deleted file mode 100644 index 640c36c78..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/gdrop_gdrop.ma +++ /dev/null @@ -1,40 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/substitution/gdrop.ma". - -(* GLOBAL ENVIRONMENT SLICING ***********************************************) - -(* Main properties **********************************************************) - -theorem gdrop_mono: ∀G,G1,e. ⇩[e] G ≡ G1 → ∀G2. ⇩[e] G ≡ G2 → G1 = G2. -#G #G1 #e #H elim H -G -G1 -[ #G #He #G2 #H - >(gdrop_inv_gt … H He) -H -He // -| #G #He #G2 #H - >(gdrop_inv_eq … H He) -H -He // -| #I #G #G1 #V #He #_ #IHG1 #G2 #H - lapply (gdrop_inv_lt … H He) -H -He /2 width=1/ -] -qed-. - -lemma gdrop_dec: ∀G1,G2,e. Decidable (⇩[e] G1 ≡ G2). -#G1 #G2 #e -elim (gdrop_total e G1) #G #HG1 -elim (genv_eq_dec G G2) #HG2 -[ destruct /2 width=1/ -| @or_intror #HG12 - lapply (gdrop_mono … HG1 … HG12) -HG1 -HG12 /2 width=1/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/ldrop.ma deleted file mode 100644 index c00819a74..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ldrop.ma +++ /dev/null @@ -1,227 +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/grammar/lenv_weight.ma". -include "Basic_2/grammar/lsubs.ma". -include "Basic_2/substitution/lift.ma". - -(* LOCAL ENVIRONMENT SLICING ************************************************) - -(* Basic_1: includes: drop_skip_bind *) -inductive ldrop: nat → nat → relation lenv ≝ -| ldrop_atom : ∀d,e. ldrop d e (⋆) (⋆) -| ldrop_pair : ∀L,I,V. ldrop 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) -| ldrop_ldrop: ∀L1,L2,I,V,e. ldrop 0 e L1 L2 → ldrop 0 (e + 1) (L1. ⓑ{I} V) L2 -| ldrop_skip : ∀L1,L2,I,V1,V2,d,e. - ldrop d e L1 L2 → ⇧[d,e] V2 ≡ V1 → - ldrop (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) -. - -interpretation "local slicing" 'RDrop d e L1 L2 = (ldrop d e L1 L2). - -(* Basic inversion lemmas ***************************************************) - -fact ldrop_inv_refl_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| // -| #L1 #L2 #I #V #e #_ #_ >commutative_plus normalize #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -(* Basic_1: was: drop_gen_refl *) -lemma ldrop_inv_refl: ∀L1,L2. ⇩[0, 0] L1 ≡ L2 → L1 = L2. -/2 width=5/ qed-. - -fact ldrop_inv_atom1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → L1 = ⋆ → - L2 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V #e #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -(* Basic_1: was: drop_gen_sort *) -lemma ldrop_inv_atom1: ∀d,e,L2. ⇩[d, e] ⋆ ≡ L2 → L2 = ⋆. -/2 width=5/ qed-. - -fact ldrop_inv_O1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → - ∀K,I,V. L1 = K. ⓑ{I} V → - (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ - (0 < e ∧ ⇩[d, e - 1] K ≡ L2). -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #K #I #V #H destruct -| #L #I #V #_ #K #J #W #HX destruct /3 width=1/ -| #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct /3 width=1/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ldrop_inv_O1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → - (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ - (0 < e ∧ ⇩[0, e - 1] K ≡ L2). -/2 width=3/ qed-. - -(* Basic_1: was: drop_gen_drop *) -lemma ldrop_inv_ldrop1: ∀e,K,I,V,L2. - ⇩[0, e] K. ⓑ{I} V ≡ L2 → 0 < e → ⇩[0, e - 1] K ≡ L2. -#e #K #I #V #L2 #H #He -elim (ldrop_inv_O1 … H) -H * // #H destruct -elim (lt_refl_false … He) -qed-. - -fact ldrop_inv_skip1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → - ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & - ⇧[d - 1, e] V2 ≡ V1 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K #V #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) -| #X #L2 #Y #Z #V2 #d #e #HL12 #HV12 #_ #I #L1 #V1 #H destruct /2 width=5/ -] -qed. - -(* Basic_1: was: drop_gen_skip_l *) -lemma ldrop_inv_skip1: ∀d,e,I,K1,V1,L2. ⇩[d, e] K1. ⓑ{I} V1 ≡ L2 → 0 < d → - ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & - ⇧[d - 1, e] V2 ≡ V1 & - L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. - -fact ldrop_inv_skip2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → - ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & - ⇧[d - 1, e] V2 ≡ V1 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K #V #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) -| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct /2 width=5/ -] -qed. - -(* Basic_1: was: drop_gen_skip_r *) -lemma ldrop_inv_skip2: ∀d,e,I,L1,K2,V2. ⇩[d, e] L1 ≡ K2. ⓑ{I} V2 → 0 < d → - ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & ⇧[d - 1, e] V2 ≡ V1 & - L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was by definition: drop_refl *) -lemma ldrop_refl: ∀L. ⇩[0, 0] L ≡ L. -#L elim L -L // -qed. - -lemma ldrop_ldrop_lt: ∀L1,L2,I,V,e. - ⇩[0, e - 1] L1 ≡ L2 → 0 < e → ⇩[0, e] L1. ⓑ{I} V ≡ L2. -#L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) // /2 width=1/ -qed. - -lemma ldrop_lsubs_ldrop1_abbr: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → - ∀K1,V,i. ⇩[0, i] L1 ≡ K1. ⓓV → - d ≤ i → i < d + e → - ∃∃K2. K1 [0, d + e - i - 1] ≼ K2 & - ⇩[0, i] L2 ≡ K2. ⓓV. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e -[ #d #e #K1 #V #i #H - lapply (ldrop_inv_atom1 … H) -H #H destruct -| #L1 #L2 #K1 #V #i #_ #_ #H - elim (lt_zero_false … H) -| #L1 #L2 #V #e #HL12 #IHL12 #K1 #W #i #H #_ #Hie - elim (ldrop_inv_O1 … H) -H * #Hi #HLK1 - [ -IHL12 -Hie destruct - minus_minus_comm >arith_b1 // /4 width=3/ - ] -| #L1 #L2 #I #V1 #V2 #e #_ #IHL12 #K1 #W #i #H #_ #Hie - elim (ldrop_inv_O1 … H) -H * #Hi #HLK1 - [ -IHL12 -Hie -Hi destruct - | elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hie >minus_minus_comm >arith_b1 // /3 width=3/ - ] -| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide - elim (le_inv_plus_l … Hdi) #Hdim #Hi - lapply (ldrop_inv_ldrop1 … H ?) -H // #HLK1 - elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hdi -Hide >minus_minus_comm >arith_b1 // /3 width=3/ -] -qed. - -(* Basic forvard lemmas *****************************************************) - -(* Basic_1: was: drop_S *) -lemma ldrop_fwd_ldrop2: ∀L1,I2,K2,V2,e. ⇩[O, e] L1 ≡ K2. ⓑ{I2} V2 → - ⇩[O, e + 1] L1 ≡ K2. -#L1 elim L1 -L1 -[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct -| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #H - [ -IHL1 destruct /2 width=1/ - | @ldrop_ldrop >(plus_minus_m_m e 1) // /2 width=3/ - ] -] -qed-. - -lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → #[L2] ≤ #[L1]. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize -[ /2 width=3/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 - >(tw_lift … HV21) -HV21 /2 width=1/ -] -qed-. - -lemma ldrop_fwd_ldrop2_length: ∀L1,I2,K2,V2,e. - ⇩[0, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|. -#L1 elim L1 -L1 -[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct -| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #H - [ -IHL1 destruct // - | lapply (IHL1 … H) -IHL1 -H #HeK1 whd in ⊢ (? ? %); /2 width=1/ - ] -] -qed-. - -lemma ldrop_fwd_O1_length: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → |L2| = |L1| - e. -#L1 elim L1 -L1 -[ #L2 #e #H >(ldrop_inv_atom1 … H) -H // -| #K1 #I1 #V1 #IHL1 #L2 #e #H - elim (ldrop_inv_O1 … H) -H * #He #H - [ -IHL1 destruct // - | lapply (IHL1 … H) -IHL1 -H #H >H -H normalize - >minus_le_minus_minus_comm // - ] -] -qed-. - -(* Basic_1: removed theorems 49: - drop_skip_flat - cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf - drop_clear drop_clear_O drop_clear_S - clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r - clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle - getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans - getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt - getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev - drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge - getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O - getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le - getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ldrop_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/ldrop_ldrop.ma deleted file mode 100644 index 90f724ad3..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ldrop_ldrop.ma +++ /dev/null @@ -1,126 +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/lift_lift.ma". -include "Basic_2/substitution/ldrop.ma". - -(* DROPPING *****************************************************************) - -(* Main properties **********************************************************) - -(* Basic_1: was: drop_mono *) -theorem ldrop_mono: ∀d,e,L,L1. ⇩[d, e] L ≡ L1 → - ∀L2. ⇩[d, e] L ≡ L2 → L1 = L2. -#d #e #L #L1 #H elim H -d -e -L -L1 -[ #d #e #L2 #H - >(ldrop_inv_atom1 … H) -L2 // -| #K #I #V #L2 #HL12 - <(ldrop_inv_refl … HL12) -L2 // -| #L #K #I #V #e #_ #IHLK #L2 #H - lapply (ldrop_inv_ldrop1 … H ?) -H // /2 width=1/ -| #L #K1 #I #T #V1 #d #e #_ #HVT1 #IHLK1 #X #H - elim (ldrop_inv_skip1 … H ?) -H // (lift_inj … HVT1 … HVT2) -HVT1 -HVT2 - >(IHLK1 … HLK2) -IHLK1 -HLK2 // -] -qed-. - -(* Basic_1: was: drop_conf_ge *) -theorem ldrop_conf_ge: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 → - ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 + e1 ≤ e2 → - ⇩[0, e2 - e1] L1 ≡ L2. -#d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 -[ #d #e #e2 #L2 #H - >(ldrop_inv_atom1 … H) -L2 // -| // -| #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H /2 width=2/ #HL2 - minus_minus_comm /3 width=1/ -| #L #K #I #V1 #V2 #d #e #_ #_ #IHLK #e2 #L2 #H #Hdee2 - lapply (transitive_le 1 … Hdee2) // #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // -He2 #HL2 - lapply (transitive_le (1 + e) … Hdee2) // #Hee2 - @ldrop_ldrop_lt >minus_minus_comm /3 width=1/ (**) (* explicit constructor *) -] -qed. - -(* Basic_1: was: drop_conf_lt *) -theorem ldrop_conf_lt: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 → - ∀e2,K2,I,V2. ⇩[0, e2] L ≡ K2. ⓑ{I} V2 → - e2 < d1 → let d ≝ d1 - e2 - 1 in - ∃∃K1,V1. ⇩[0, e2] L1 ≡ K1. ⓑ{I} V1 & - ⇩[d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2. -#d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 -[ #d #e #e2 #K2 #I #V2 #H - lapply (ldrop_inv_atom1 … H) -H #H destruct -| #L #I #V #e2 #K2 #J #V2 #_ #H - elim (lt_zero_false … H) -| #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H - elim (lt_zero_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d - elim (ldrop_inv_O1 … H) -H * - [ -IHL12 -He2d #H1 #H2 destruct /2 width=5/ - | -HL12 -HV12 #He #HLK - elim (IHL12 … HLK ?) -IHL12 -HLK [ (ldrop_inv_atom1 … H) -L2 /2 width=3/ -| #K #I #V #e2 #L2 #HL2 #H - lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ -| #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H - lapply (le_n_O_to_eq … H) -H #H destruct - elim (IHL12 … HL2 ?) -IHL12 -HL2 // #L0 #H #HL0 - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d - elim (ldrop_inv_O1 … H) -H * - [ -He2d -IHL12 #H1 #H2 destruct /3 width=5/ - | -HL12 -HV12 #He2 #HL2 - elim (IHL12 … HL2 ?) -L2 [ >minus_le_minus_minus_comm // /3 width=3/ | /2 width=1/ ] - ] -] -qed. - -(* Basic_1: was: drop_trans_ge *) -theorem ldrop_trans_ge: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L → - ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 ≤ e2 → ⇩[0, e1 + e2] L1 ≡ L2. -#d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L -[ #d #e #e2 #L2 #H - >(ldrop_inv_atom1 … H) -H -L2 // -| // -| /3 width=1/ -| #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2 - lapply (lt_to_le_to_lt 0 … Hde2) // #He2 - lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HL2 - @ldrop_ldrop_lt // >le_plus_minus // @IHL12 /2 width=1/ (**) (* explicit constructor *) -] -qed. - -theorem ldrop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L. - ⇩[d1, e1] L1 ≡ L → ⇩[0, e2] L ≡ L2 → d1 ≤ e2 → - ⇩[0, e2 + e1] L1 ≡ L2. -#e1 #e1 #e2 >commutative_plus /2 width=5/ -qed. - -(* Basic_1: was: drop_conf_rev *) -axiom ldrop_div: ∀e1,L1,L. ⇩[0, e1] L1 ≡ L → ∀e2,L2. ⇩[0, e2] L2 ≡ L → - ∃∃L0. ⇩[0, e1] L0 ≡ L2 & ⇩[e1, e2] L0 ≡ L1. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift.ma deleted file mode 100644 index cb8aac3a0..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift.ma +++ /dev/null @@ -1,376 +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/grammar/term_weight.ma". -include "Basic_2/grammar/term_simple.ma". - -(* BASIC TERM RELOCATION ****************************************************) - -(* Basic_1: includes: - lift_sort lift_lref_lt lift_lref_ge lift_bind lift_flat -*) -inductive lift: nat → nat → relation term ≝ -| lift_sort : ∀k,d,e. lift d e (⋆k) (⋆k) -| lift_lref_lt: ∀i,d,e. i < d → lift d e (#i) (#i) -| lift_lref_ge: ∀i,d,e. d ≤ i → lift d e (#i) (#(i + e)) -| lift_gref : ∀p,d,e. lift d e (§p) (§p) -| lift_bind : ∀I,V1,V2,T1,T2,d,e. - lift d e V1 V2 → lift (d + 1) e T1 T2 → - lift d e (ⓑ{I} V1. T1) (ⓑ{I} V2. T2) -| lift_flat : ∀I,V1,V2,T1,T2,d,e. - lift d e V1 V2 → lift d e T1 T2 → - lift d e (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) -. - -interpretation "relocation" 'RLift d e T1 T2 = (lift d e T1 T2). - -(* Basic inversion lemmas ***************************************************) - -fact lift_inv_refl_O2_aux: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → e = 0 → T1 = T2. -#d #e #T1 #T2 #H elim H -d -e -T1 -T2 // /3 width=1/ -qed. - -lemma lift_inv_refl_O2: ∀d,T1,T2. ⇧[d, 0] T1 ≡ T2 → T1 = T2. -/2 width=4/ qed-. - -fact lift_inv_sort1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k. -#d #e #T1 #T2 * -d -e -T1 -T2 // -[ #i #d #e #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -] -qed. - -lemma lift_inv_sort1: ∀d,e,T2,k. ⇧[d,e] ⋆k ≡ T2 → T2 = ⋆k. -/2 width=5/ qed-. - -fact lift_inv_lref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T1 = #i → - (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). -#d #e #T1 #T2 * -d -e -T1 -T2 -[ #k #d #e #i #H destruct -| #j #d #e #Hj #i #Hi destruct /3 width=1/ -| #j #d #e #Hj #i #Hi destruct /3 width=1/ -| #p #d #e #i #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct -] -qed. - -lemma lift_inv_lref1: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → - (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). -/2 width=3/ qed-. - -lemma lift_inv_lref1_lt: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → i < d → T2 = #i. -#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // -#Hdi #_ #Hid lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd -elim (lt_refl_false … Hdd) -qed-. - -lemma lift_inv_lref1_ge: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → d ≤ i → T2 = #(i + e). -#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // -#Hid #_ #Hdi lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd -elim (lt_refl_false … Hdd) -qed-. - -fact lift_inv_gref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀p. T1 = §p → T2 = §p. -#d #e #T1 #T2 * -d -e -T1 -T2 // -[ #i #d #e #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -] -qed. - -lemma lift_inv_gref1: ∀d,e,T2,p. ⇧[d,e] §p ≡ T2 → T2 = §p. -/2 width=5/ qed-. - -fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → - ∀I,V1,U1. T1 = ⓑ{I} V1.U1 → - ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 & - T2 = ⓑ{I} V2. U2. -#d #e #T1 #T2 * -d -e -T1 -T2 -[ #k #d #e #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #p #d #e #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct -] -qed. - -lemma lift_inv_bind1: ∀d,e,T2,I,V1,U1. ⇧[d,e] ⓑ{I} V1. U1 ≡ T2 → - ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 & - T2 = ⓑ{I} V2. U2. -/2 width=3/ qed-. - -fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → - ∀I,V1,U1. T1 = ⓕ{I} V1.U1 → - ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d,e] U1 ≡ U2 & - T2 = ⓕ{I} V2. U2. -#d #e #T1 #T2 * -d -e -T1 -T2 -[ #k #d #e #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #i #d #e #_ #I #V1 #U1 #H destruct -| #p #d #e #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ -] -qed. - -lemma lift_inv_flat1: ∀d,e,T2,I,V1,U1. ⇧[d,e] ⓕ{I} V1. U1 ≡ T2 → - ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d,e] U1 ≡ U2 & - T2 = ⓕ{I} V2. U2. -/2 width=3/ qed-. - -fact lift_inv_sort2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k. -#d #e #T1 #T2 * -d -e -T1 -T2 // -[ #i #d #e #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct -] -qed. - -(* Basic_1: was: lift_gen_sort *) -lemma lift_inv_sort2: ∀d,e,T1,k. ⇧[d,e] T1 ≡ ⋆k → T1 = ⋆k. -/2 width=5/ qed-. - -fact lift_inv_lref2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T2 = #i → - (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)). -#d #e #T1 #T2 * -d -e -T1 -T2 -[ #k #d #e #i #H destruct -| #j #d #e #Hj #i #Hi destruct /3 width=1/ -| #j #d #e #Hj #i #Hi destruct (plus_minus_m_m i e) in ⊢ (? ? ? ? %); /2 width=2/ /3 width=2/ -qed. - -lemma lift_lref_ge_minus_eq: ∀d,e,i,j. d + e ≤ i → j = i - e → ⇧[d, e] #j ≡ #i. -/2 width=1/ qed-. - -(* Basic_1: was: lift_r *) -lemma lift_refl: ∀T,d. ⇧[d, 0] T ≡ T. -#T elim T -T -[ * #i // #d elim (lt_or_ge i d) /2 width=1/ -| * /2 width=1/ -] -qed. - -lemma lift_total: ∀T1,d,e. ∃T2. ⇧[d,e] T1 ≡ T2. -#T1 elim T1 -T1 -[ * #i /2 width=2/ #d #e elim (lt_or_ge i d) /3 width=2/ -| * #I #V1 #T1 #IHV1 #IHT1 #d #e - elim (IHV1 d e) -IHV1 #V2 #HV12 - [ elim (IHT1 (d+1) e) -IHT1 /3 width=2/ - | elim (IHT1 d e) -IHT1 /3 width=2/ - ] -] -qed. - -(* Basic_1: was: lift_free (right to left) *) -lemma lift_split: ∀d1,e2,T1,T2. ⇧[d1, e2] T1 ≡ T2 → - ∀d2,e1. d1 ≤ d2 → d2 ≤ d1 + e1 → e1 ≤ e2 → - ∃∃T. ⇧[d1, e1] T1 ≡ T & ⇧[d2, e2 - e1] T ≡ T2. -#d1 #e2 #T1 #T2 #H elim H -d1 -e2 -T1 -T2 -[ /3 width=3/ -| #i #d1 #e2 #Hid1 #d2 #e1 #Hd12 #_ #_ - lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4 width=3/ -| #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12 - lapply (transitive_le … (i+e1) Hd21 ?) /2 width=1/ -Hd21 #Hd21 - >(plus_minus_m_m e2 e1 ?) // /3 width=3/ -| /3 width=3/ -| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 - elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b - elim (IHT (d2+1) … ? ? He12) /2 width=1/ /3 width=5/ -| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 - elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b - elim (IHT d2 … ? ? He12) // /3 width=5/ -] -qed. - -(* Basic_1: was only: dnf_dec2 dnf_dec *) -lemma is_lift_dec: ∀T2,d,e. Decidable (∃T1. ⇧[d,e] T1 ≡ T2). -#T1 elim T1 -T1 -[ * [1,3: /3 width=2/ ] #i #d #e - elim (lt_dec i d) #Hid - [ /4 width=2/ - | lapply (false_lt_to_le … Hid) -Hid #Hid - elim (lt_dec i (d + e)) #Hide - [ @or_intror * #T1 #H - elim (lift_inv_lref2_be … H Hid Hide) - | lapply (false_lt_to_le … Hide) -Hide /4 width=2/ - ] - ] -| * #I #V2 #T2 #IHV2 #IHT2 #d #e - [ elim (IHV2 d e) -IHV2 - [ * #V1 #HV12 elim (IHT2 (d+1) e) -IHT2 - [ * #T1 #HT12 @or_introl /3 width=2/ - | -V1 #HT2 @or_intror * #X #H - elim (lift_inv_bind2 … H) -H /3 width=2/ - ] - | -IHT2 #HV2 @or_intror * #X #H - elim (lift_inv_bind2 … H) -H /3 width=2/ - ] - | elim (IHV2 d e) -IHV2 - [ * #V1 #HV12 elim (IHT2 d e) -IHT2 - [ * #T1 #HT12 /4 width=2/ - | -V1 #HT2 @or_intror * #X #H - elim (lift_inv_flat2 … H) -H /3 width=2/ - ] - | -IHT2 #HV2 @or_intror * #X #H - elim (lift_inv_flat2 … H) -H /3 width=2/ - ] - ] -] -qed. - -(* Basic_1: removed theorems 7: - lift_head lift_gen_head - lift_weight_map lift_weight lift_weight_add lift_weight_add_O - lift_tlt_dx -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_lift.ma deleted file mode 100644 index 30bf8886e..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_lift.ma +++ /dev/null @@ -1,200 +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/lift.ma". - -(* BASIC TERM RELOCATION ****************************************************) - -(* Main properies ***********************************************************) - -(* Basic_1: was: lift_inj *) -theorem lift_inj: ∀d,e,T1,U. ⇧[d,e] T1 ≡ U → ∀T2. ⇧[d,e] T2 ≡ U → T1 = T2. -#d #e #T1 #U #H elim H -d -e -T1 -U -[ #k #d #e #X #HX - lapply (lift_inv_sort2 … HX) -HX // -| #i #d #e #Hid #X #HX - lapply (lift_inv_lref2_lt … HX ?) -HX // -| #i #d #e #Hdi #X #HX - lapply (lift_inv_lref2_ge … HX ?) -HX // /2 width=1/ -| #p #d #e #X #HX - lapply (lift_inv_gref2 … HX) -HX // -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ -] -qed-. - -(* Basic_1: was: lift_gen_lift *) -theorem lift_div_le: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ⇧[d2 + e1, e2] T2 ≡ T → - d1 ≤ d2 → - ∃∃T0. ⇧[d1, e1] T0 ≡ T2 & ⇧[d2, e2] T0 ≡ T1. -#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T -[ #k #d1 #e1 #d2 #e2 #T2 #Hk #Hd12 - lapply (lift_inv_sort2 … Hk) -Hk #Hk destruct /3 width=3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12 - lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 - lapply (lift_inv_lref2_lt … Hi ?) -Hi /2 width=3/ /3 width=3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12 - elim (lift_inv_lref2 … Hi) -Hi * #Hid2 #H destruct - [ -Hd12 lapply (lt_plus_to_lt_l … Hid2) -Hid2 #Hid2 /3 width=3/ - | -Hid1 >plus_plus_comm_23 in Hid2; #H lapply (le_plus_to_le_r … H) -H #H - elim (le_inv_plus_l … H) -H #Hide2 #He2i - lapply (transitive_le … Hd12 Hide2) -Hd12 #Hd12 - >le_plus_minus_comm // >(plus_minus_m_m i e2) in ⊢ (? ? ? %); // -He2i - /4 width=3/ - ] -| #p #d1 #e1 #d2 #e2 #T2 #Hk #Hd12 - lapply (lift_inv_gref2 … Hk) -Hk #Hk destruct /3 width=3/ -| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 - lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct - elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1 - >plus_plus_comm_23 in HU2; #HU2 elim (IHU … HU2 ?) /2 width=1/ /3 width=5/ -| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 - lapply (lift_inv_flat2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct - elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1 - elim (IHU … HU2 ?) // /3 width=5/ -] -qed. - -(* Note: apparently this was missing in Basic_1 *) -theorem lift_div_be: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → - ∀e,e2,T2. ⇧[d1 + e, e2] T2 ≡ T → - e ≤ e1 → e1 ≤ e + e2 → - ∃∃T0. ⇧[d1, e] T0 ≡ T2 & ⇧[d1, e + e2 - e1] T0 ≡ T1. -#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T -[ #k #d1 #e1 #e #e2 #T2 #H >(lift_inv_sort2 … H) -H /2 width=3/ -| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2 - >(lift_inv_lref2_lt … H) -H [ /3 width=3/ | /2 width=3/ ] -| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2 - elim (lt_or_ge (i+e1) (d1+e+e2)) #Hie1d1e2 - [ elim (lift_inv_lref2_be … H ? ?) -H // /2 width=1/ - | >(lift_inv_lref2_ge … H ?) -H // - lapply (le_plus_to_minus … Hie1d1e2) #Hd1e21i - elim (le_inv_plus_l … Hie1d1e2) -Hie1d1e2 #Hd1e12 #He2ie1 - @ex2_1_intro [2: /2 width=1/ | skip ] -Hd1e12 - @lift_lref_ge_minus_eq [ >plus_minus_commutative // | /2 width=1/ ] - ] -| #p #d1 #e1 #e #e2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3/ -| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2 - elim (lift_inv_bind2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct - elim (IHV1 … HV2 ? ?) -V // >plus_plus_comm_23 in HT2; #HT2 - elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/ -| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2 - elim (lift_inv_flat2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct - elim (IHV1 … HV2 ? ?) -V // - elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/ -] -qed. - -theorem lift_mono: ∀d,e,T,U1. ⇧[d,e] T ≡ U1 → ∀U2. ⇧[d,e] T ≡ U2 → U1 = U2. -#d #e #T #U1 #H elim H -d -e -T -U1 -[ #k #d #e #X #HX - lapply (lift_inv_sort1 … HX) -HX // -| #i #d #e #Hid #X #HX - lapply (lift_inv_lref1_lt … HX ?) -HX // -| #i #d #e #Hdi #X #HX - lapply (lift_inv_lref1_ge … HX ?) -HX // -| #p #d #e #X #HX - lapply (lift_inv_gref1 … HX) -HX // -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ -| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ -] -qed-. - -(* Basic_1: was: lift_free (left to right) *) -theorem lift_trans_be: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ⇧[d2, e2] T ≡ T2 → - d1 ≤ d2 → d2 ≤ d1 + e1 → ⇧[d1, e1 + e2] T1 ≡ T2. -#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T -[ #k #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_ - >(lift_inv_sort1 … HT2) -HT2 // -| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #Hd12 #_ - lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 - lapply (lift_inv_lref1_lt … HT2 Hid2) /2 width=1/ -| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #_ #Hd21 - lapply (lift_inv_lref1_ge … HT2 ?) -HT2 - [ @(transitive_le … Hd21 ?) -Hd21 /2 width=1/ - | -Hd21 /2 width=1/ - ] -| #p #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_ - >(lift_inv_gref1 … HT2) -HT2 // -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10 - lapply (IHT12 … HT20 ? ?) /2 width=1/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10 - lapply (IHT12 … HT20 ? ?) // /2 width=1/ -] -qed. - -(* Basic_1: was: lift_d (right to left) *) -theorem lift_trans_le: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ⇧[d2, e2] T ≡ T2 → d2 ≤ d1 → - ∃∃T0. ⇧[d2, e2] T1 ≡ T0 & ⇧[d1 + e2, e1] T0 ≡ T2. -#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T -[ #k #d1 #e1 #d2 #e2 #X #HX #_ - >(lift_inv_sort1 … HX) -HX /2 width=3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ - lapply (lt_to_le_to_lt … (d1+e2) Hid1 ?) // #Hie2 - elim (lift_inv_lref1 … HX) -HX * #Hid2 #HX destruct /3 width=3/ /4 width=3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hd21 - lapply (transitive_le … Hd21 Hid1) -Hd21 #Hid2 - lapply (lift_inv_lref1_ge … HX ?) -HX /2 width=3/ #HX destruct - >plus_plus_comm_23 /4 width=3/ -| #p #d1 #e1 #d2 #e2 #X #HX #_ - >(lift_inv_gref1 … HX) -HX /2 width=3/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20 ?) -IHV12 -HV20 // - elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ /3 width=5/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20 ?) -IHV12 -HV20 // - elim (IHT12 … HT20 ?) -IHT12 -HT20 // /3 width=5/ -] -qed. - -(* Basic_1: was: lift_d (left to right) *) -theorem lift_trans_ge: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → - ∀d2,e2,T2. ⇧[d2, e2] T ≡ T2 → d1 + e1 ≤ d2 → - ∃∃T0. ⇧[d2 - e1, e2] T1 ≡ T0 & ⇧[d1, e1] T0 ≡ T2. -#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T -[ #k #d1 #e1 #d2 #e2 #X #HX #_ - >(lift_inv_sort1 … HX) -HX /2 width=3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hded - lapply (lt_to_le_to_lt … (d1+e1) Hid1 ?) // #Hid1e - lapply (lt_to_le_to_lt … (d2-e1) Hid1 ?) /2 width=1/ #Hid2e - lapply (lt_to_le_to_lt … Hid1e Hded) -Hid1e -Hded #Hid2 - lapply (lift_inv_lref1_lt … HX ?) -HX // #HX destruct /3 width=3/ -| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ - elim (lift_inv_lref1 … HX) -HX * #Hied #HX destruct /4 width=3/ -| #p #d1 #e1 #d2 #e2 #X #HX #_ - >(lift_inv_gref1 … HX) -HX /2 width=3/ -| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20 ?) -IHV12 -HV20 // - elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ #T - (liftv_inv_nil1 … H) -H // -| #Ts #U1s #T #U1 #HTU1 #_ #IHTU1s #X #H destruct - elim (liftv_inv_cons1 … H) -H #U2 #U2s #HTU2 #HTU2s #H destruct - >(lift_mono … HTU1 … HTU2) -T /3 width=1/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_vector.ma deleted file mode 100644 index aa0a30f95..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_vector.ma +++ /dev/null @@ -1,62 +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/grammar/term_vector.ma". -include "Basic_2/substitution/lift.ma". - -(* BASIC TERM VECTOR RELOCATION *********************************************) - -inductive liftv (d,e:nat) : relation (list term) ≝ -| liftv_nil : liftv d e ◊ ◊ -| liftv_cons: ∀T1s,T2s,T1,T2. - ⇧[d, e] T1 ≡ T2 → liftv d e T1s T2s → - liftv d e (T1 :: T1s) (T2 :: T2s) -. - -interpretation "relocation (vector)" 'RLift d e T1s T2s = (liftv d e T1s T2s). - -(* Basic inversion lemmas ***************************************************) - -fact liftv_inv_nil1_aux: ∀T1s,T2s,d,e. ⇧[d, e] T1s ≡ T2s → T1s = ◊ → T2s = ◊. -#T1s #T2s #d #e * -T1s -T2s // -#T1s #T2s #T1 #T2 #_ #_ #H destruct -qed. - -lemma liftv_inv_nil1: ∀T2s,d,e. ⇧[d, e] ◊ ≡ T2s → T2s = ◊. -/2 width=5/ qed-. - -fact liftv_inv_cons1_aux: ∀T1s,T2s,d,e. ⇧[d, e] T1s ≡ T2s → - ∀U1,U1s. T1s = U1 :: U1s → - ∃∃U2,U2s. ⇧[d, e] U1 ≡ U2 & ⇧[d, e] U1s ≡ U2s & - T2s = U2 :: U2s. -#T1s #T2s #d #e * -T1s -T2s -[ #U1 #U1s #H destruct -| #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5/ -] -qed. - -lemma liftv_inv_cons1: ∀U1,U1s,T2s,d,e. ⇧[d, e] U1 :: U1s ≡ T2s → - ∃∃U2,U2s. ⇧[d, e] U1 ≡ U2 & ⇧[d, e] U1s ≡ U2s & - T2s = U2 :: U2s. -/2 width=3/ qed-. - -(* Basic properties *********************************************************) - -lemma liftv_total: ∀d,e. ∀T1s:list term. ∃T2s. ⇧[d, e] T1s ≡ T2s. -#d #e #T1s elim T1s -T1s -[ /2 width=2/ -| #T1 #T1s * #T2s #HT12s - elim (lift_total T1 d e) /3 width=2/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ltps.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/ltps.ma deleted file mode 100644 index 4f30025d3..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ltps.ma +++ /dev/null @@ -1,191 +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/tps.ma". - -(* PARALLEL SUBSTITUTION ON LOCAL ENVIRONMENTS ******************************) - -(* Basic_1: includes: csubst1_bind *) -inductive ltps: nat → nat → relation lenv ≝ -| ltps_atom: ∀d,e. ltps d e (⋆) (⋆) -| ltps_pair: ∀L,I,V. ltps 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) -| ltps_tps2: ∀L1,L2,I,V1,V2,e. - ltps 0 e L1 L2 → L2 ⊢ V1 [0, e] ▶ V2 → - ltps 0 (e + 1) (L1. ⓑ{I} V1) L2. ⓑ{I} V2 -| ltps_tps1: ∀L1,L2,I,V1,V2,d,e. - ltps d e L1 L2 → L2 ⊢ V1 [d, e] ▶ V2 → - ltps (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) -. - -interpretation "parallel substritution (local environment)" - 'PSubst L1 d e L2 = (ltps d e L1 L2). - -(* Basic properties *********************************************************) - -lemma ltps_tps2_lt: ∀L1,L2,I,V1,V2,e. - L1 [0, e - 1] ▶ L2 → L2 ⊢ V1 [0, e - 1] ▶ V2 → - 0 < e → L1. ⓑ{I} V1 [0, e] ▶ L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) /2 width=1/ -qed. - -lemma ltps_tps1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 [d - 1, e] ▶ L2 → L2 ⊢ V1 [d - 1, e] ▶ V2 → - 0 < d → L1. ⓑ{I} V1 [d, e] ▶ L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) /2 width=1/ -qed. - -(* Basic_1: was by definition: csubst1_refl *) -lemma ltps_refl: ∀L,d,e. L [d, e] ▶ L. -#L elim L -L // -#L #I #V #IHL * /2 width=1/ * /2 width=1/ -qed. - -lemma ltps_weak_all: ∀L1,L2,d,e. L1 [d, e] ▶ L2 → L1 [0, |L2|] ▶ L2. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e -// /3 width=2/ /3 width=3/ -qed. - -(* Basic forward lemmas *****************************************************) - -lemma ltps_fwd_length: ∀L1,L2,d,e. L1 [d, e] ▶ L2 → |L1| = |L2|. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e -normalize // -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact ltps_inv_refl_O2_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → e = 0 → L1 = L2. -#d #e #L1 #L2 #H elim H -d -e -L1 -L2 // -[ #L1 #L2 #I #V1 #V2 #e #_ #_ #_ >commutative_plus normalize #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV12 #IHL12 #He destruct - >(IHL12 ?) -IHL12 // >(tps_inv_refl_O2 … HV12) // -] -qed. - -lemma ltps_inv_refl_O2: ∀d,L1,L2. L1 [d, 0] ▶ L2 → L1 = L2. -/2 width=4/ qed-. - -fact ltps_inv_atom1_aux: ∀d,e,L1,L2. - L1 [d, e] ▶ L2 → L1 = ⋆ → L2 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma ltps_inv_atom1: ∀d,e,L2. ⋆ [d, e] ▶ L2 → L2 = ⋆. -/2 width=5/ qed-. - -fact ltps_inv_tps21_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → d = 0 → 0 < e → - ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 [0, e - 1] ▶ K2 & - K2 ⊢ V1 [0, e - 1] ▶ V2 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K1 #J #W1 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ltps_inv_tps21: ∀e,K1,I,V1,L2. K1. ⓑ{I} V1 [0, e] ▶ L2 → 0 < e → - ∃∃K2,V2. K1 [0, e - 1] ▶ K2 & K2 ⊢ V1 [0, e - 1] ▶ V2 & - L2 = K2. ⓑ{I} V2. -/2 width=5/ qed-. - -fact ltps_inv_tps11_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → 0 < d → - ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 [d - 1, e] ▶ K2 & - K2 ⊢ V1 [d - 1, e] ▶ V2 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K1 #V1 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K1 #W1 #H destruct /2 width=5/ -] -qed. - -lemma ltps_inv_tps11: ∀d,e,I,K1,V1,L2. K1. ⓑ{I} V1 [d, e] ▶ L2 → 0 < d → - ∃∃K2,V2. K1 [d - 1, e] ▶ K2 & - K2 ⊢ V1 [d - 1, e] ▶ V2 & - L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. - -fact ltps_inv_atom2_aux: ∀d,e,L1,L2. - L1 [d, e] ▶ L2 → L2 = ⋆ → L1 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma ltps_inv_atom2: ∀d,e,L1. L1 [d, e] ▶ ⋆ → L1 = ⋆. -/2 width=5/ qed-. - -fact ltps_inv_tps22_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → d = 0 → 0 < e → - ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & - K2 ⊢ V1 [0, e - 1] ▶ V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ltps_inv_tps22: ∀e,L1,K2,I,V2. L1 [0, e] ▶ K2. ⓑ{I} V2 → 0 < e → - ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & K2 ⊢ V1 [0, e - 1] ▶ V2 & - L1 = K1. ⓑ{I} V1. -/2 width=5/ qed-. - -fact ltps_inv_tps12_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → 0 < d → - ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & - K2 ⊢ V1 [d - 1, e] ▶ V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K2 #V2 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ -] -qed. - -lemma ltps_inv_tps12: ∀L1,K2,I,V2,d,e. L1 [d, e] ▶ K2. ⓑ{I} V2 → 0 < d → - ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & - K2 ⊢ V1 [d - 1, e] ▶ V2 & - L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. - -(* Basic_1: removed theorems 27: - csubst0_clear_O csubst0_drop_lt csubst0_drop_gt csubst0_drop_eq - csubst0_clear_O_back csubst0_clear_S csubst0_clear_trans - csubst0_drop_gt_back csubst0_drop_eq_back csubst0_drop_lt_back - csubst0_gen_sort csubst0_gen_head csubst0_getl_ge csubst0_getl_lt - csubst0_gen_S_bind_2 csubst0_getl_ge_back csubst0_getl_lt_back - csubst0_snd_bind csubst0_fst_bind csubst0_both_bind - csubst1_head csubst1_flat csubst1_gen_head - csubst1_getl_ge csubst1_getl_lt csubst1_getl_ge_back getl_csubst1 - -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ltps_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/ltps_ldrop.ma deleted file mode 100644 index bf4e1f711..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/ltps_ldrop.ma +++ /dev/null @@ -1,131 +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/ltps.ma". - -(* PARALLEL SUBSTITUTION ON LOCAL ENVIRONMENTS ******************************) - -lemma ltps_ldrop_conf_ge: ∀L0,L1,d1,e1. L0 [d1, e1] ▶ L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. -#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // -| // -| normalize #K0 #K1 #I #V0 #V1 #e1 #_ #_ #IHK01 #L2 #e2 #H #He12 - elim (le_inv_plus_l … He12) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ -| #K0 #K1 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ -] -qed. - -lemma ltps_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶ L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. -#L1 #L0 #d1 #e1 #H elim H -L1 -L0 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // -| // -| normalize #K1 #K0 #I #V1 #V0 #e1 #_ #_ #IHK10 #L2 #e2 #H #He12 - elim (le_inv_plus_l … He12) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK10 … HK0L2 ?) -K0 /2 width=1/ -| #K0 #K1 #I #V1 #V0 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK10 … HK0L2 ?) -IHK10 -HK0L2 /2 width=1/ -] -qed. - -lemma ltps_ldrop_conf_be: ∀L0,L1,d1,e1. L0 [d1, e1] ▶ L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → - ∃∃L. L2 [0, d1 + e1 - e2] ▶ L & ⇩[0, e2] L1 ≡ L. -#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ -| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 - lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ -| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK01 -He21 destruct plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 #He2de1 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - (ldrop_inv_atom1 … H) -H /2 width=3/ -| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 - lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ -| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK10 -He21 destruct plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 #He2de1 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - (ldrop_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #_ #L2 #e2 #H #He2 - lapply (le_n_O_to_eq … He2) -He2 #He2 destruct - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ -| #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #He2d1 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK01 -He2d1 destruct (ldrop_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #_ #L2 #e2 #H #He2 - lapply (le_n_O_to_eq … He2) -He2 #He2 destruct - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ -| #K1 #K0 #I #V1 #V0 #d1 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #He2d1 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK10 -He2d1 destruct minus_plus minus_plus >commutative_plus /2 width=1/ - | lapply (ltps_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /3 width=4/ - ] - ] -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 - elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2 - elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 /3 width=5/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 - elim (IHVW2 … HL01) -IHVW2 - elim (IHTU2 … HL01) -IHTU2 -HL01 /3 width=5/ -] -qed. - -lemma ltps_tps_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 → - ∀L1,d1,e1. L1 [d1, e1] ▶ L0 → d1 + e1 ≤ d2 → - L1 ⊢ T2 [d2, e2] ▶ U2. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ // -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 #Hde1d2 - lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2 - lapply (ltps_ldrop_trans_ge … HL10 … HLK0 ?) -L0 // /2 width=4/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 #Hde1d2 - @tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltps_tps1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *) -| /3 width=4/ -] -qed. - -(* Basic_1: was: subst1_subst1 *) -lemma ltps_tps_trans: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 → - ∀L1,d1,e1. L1 [d1, e1] ▶ L0 → - ∃∃T. L1 ⊢ T2 [d2, e2] ▶ T & - L0 ⊢ T [d1, e1] ▶ U2. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ /2 width=3/ -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 - elim (lt_or_ge i2 d1) #Hi2d1 - [ elim (ltps_ldrop_trans_le … HL10 … HLK0 ?) -HL10 /2 width=2/ #X #H #HLK1 - elim (ltps_inv_tps12 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct - lapply (ldrop_fwd_ldrop2 … HLK0) -HLK0 #H - elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 - lapply (tps_lift_ge … HV01 … H HVW1 HVW0 ?) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ - | lapply (ltps_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/ - ] - ] -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 - elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2 - elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 /3 width=5/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 - elim (IHVW2 … HL10) -IHVW2 - elim (IHTU2 … HL10) -IHTU2 -HL10 /3 width=5/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/substitution/tps.ma b/matita/matita/contribs/lambda_delta/Basic_2/substitution/tps.ma deleted file mode 100644 index 11edc90b1..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/substitution/tps.ma +++ /dev/null @@ -1,239 +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/grammar/cl_weight.ma". -include "Basic_2/substitution/ldrop.ma". - -(* PARALLEL SUBSTITUTION ON TERMS *******************************************) - -inductive tps: nat → nat → lenv → relation term ≝ -| tps_atom : ∀L,I,d,e. tps d e L (⓪{I}) (⓪{I}) -| tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e → - ⇩[0, i] L ≡ K. ⓓV → ⇧[0, i + 1] V ≡ W → tps d e L (#i) W -| tps_bind : ∀L,I,V1,V2,T1,T2,d,e. - tps d e L V1 V2 → tps (d + 1) e (L. ⓑ{I} V2) T1 T2 → - tps d e L (ⓑ{I} V1. T1) (ⓑ{I} V2. T2) -| tps_flat : ∀L,I,V1,V2,T1,T2,d,e. - tps d e L V1 V2 → tps d e L T1 T2 → - tps d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) -. - -interpretation "parallel substritution (term)" - 'PSubst L T1 d e T2 = (tps d e L T1 T2). - -(* Basic properties *********************************************************) - -lemma tps_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 → - ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ▶ T2. -#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e -[ // -| #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 - elim (ldrop_lsubs_ldrop1_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /2 width=4/ -| /4 width=1/ -| /3 width=1/ -] -qed. - -lemma tps_refl: ∀T,L,d,e. L ⊢ T [d, e] ▶ T. -#T elim T -T // -#I elim I -I /2 width=1/ -qed. - -(* Basic_1: was: subst1_ex *) -lemma tps_full: ∀K,V,T1,L,d. ⇩[0, d] L ≡ (K. ⓓV) → - ∃∃T2,T. L ⊢ T1 [d, 1] ▶ T2 & ⇧[d, 1] T ≡ T2. -#K #V #T1 elim T1 -T1 -[ * #i #L #d #HLK /2 width=4/ - elim (lt_or_eq_or_gt i d) #Hid /3 width=4/ - destruct - elim (lift_total V 0 (i+1)) #W #HVW - elim (lift_split … HVW i i ? ? ?) // /3 width=4/ -| * #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK - elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 - [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=8/ - | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8/ - ] -] -qed. - -lemma tps_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ▶ T2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → - L ⊢ T1 [d2, e2] ▶ T2. -#L #T1 #T2 #d1 #e1 #H elim H -L -T1 -T2 -d1 -e1 -[ // -| #L #K #V #W #i #d1 #e1 #Hid1 #Hide1 #HLK #HVW #d2 #e2 #Hd12 #Hde12 - lapply (transitive_le … Hd12 … Hid1) -Hd12 -Hid1 #Hid2 - lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 /2 width=4/ -| /4 width=3/ -| /4 width=1/ -] -qed. - -lemma tps_weak_top: ∀L,T1,T2,d,e. - L ⊢ T1 [d, e] ▶ T2 → L ⊢ T1 [d, |L| - d] ▶ T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ // -| #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW - lapply (ldrop_fwd_ldrop2_length … HLK) #Hi - lapply (le_to_lt_to_lt … Hdi Hi) /3 width=4/ -| normalize /2 width=1/ -| /2 width=1/ -] -qed. - -lemma tps_weak_all: ∀L,T1,T2,d,e. - L ⊢ T1 [d, e] ▶ T2 → L ⊢ T1 [0, |L|] ▶ T2. -#L #T1 #T2 #d #e #HT12 -lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 -lapply (tps_weak_top … HT12) // -qed. - -lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶ T2 → ∀i. d ≤ i → i ≤ d + e → - ∃∃T. L ⊢ T1 [d, i - d] ▶ T & L ⊢ T [i, d + e - i] ▶ T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ /2 width=3/ -| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde - elim (lt_or_ge i j) - [ -Hide -Hjde - >(plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=4/ - | -Hdi -Hdj #Hid - generalize in match Hide; -Hide (**) (* rewriting in the premises, rewrites in the goal too *) - >(plus_minus_m_m … Hjde) in ⊢ (% → ?); -Hjde /4 width=4/ - ] -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 - elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ - -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // - -Hdi -Hide /3 width=5/ -] -qed. - -(* Basic inversion lemmas ***************************************************) - -fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶ T2 → ∀I. T1 = ⓪{I} → - T2 = ⓪{I} ∨ - ∃∃K,V,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2 & - I = LRef i. -#L #T1 #T2 #d #e * -L -T1 -T2 -d -e -[ #L #I #d #e #J #H destruct /2 width=1/ -| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -] -qed. - -lemma tps_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} [d, e] ▶ T2 → - T2 = ⓪{I} ∨ - ∃∃K,V,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2 & - I = LRef i. -/2 width=3/ qed-. - - -(* Basic_1: was: subst1_gen_sort *) -lemma tps_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k [d, e] ▶ T2 → T2 = ⋆k. -#L #T2 #k #d #e #H -elim (tps_inv_atom1 … H) -H // -* #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -(* Basic_1: was: subst1_gen_lref *) -lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i [d, e] ▶ T2 → - T2 = #i ∨ - ∃∃K,V. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2. -#L #T2 #i #d #e #H -elim (tps_inv_atom1 … H) -H /2 width=1/ -* #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=4/ -qed-. - -lemma tps_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p [d, e] ▶ T2 → T2 = §p. -#L #T2 #p #d #e #H -elim (tps_inv_atom1 … H) -H // -* #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ▶ U2 → - ∀I,V1,T1. U1 = ⓑ{I} V1. T1 → - ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & - L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶ T2 & - U2 = ⓑ{I} V2. T2. -#d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ -| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct -] -qed. - -lemma tps_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 [d, e] ▶ U2 → - ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & - L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶ T2 & - U2 = ⓑ{I} V2. T2. -/2 width=3/ qed-. - -fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ▶ U2 → - ∀I,V1,T1. U1 = ⓕ{I} V1. T1 → - ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & L ⊢ T1 [d, e] ▶ T2 & - U2 = ⓕ{I} V2. T2. -#d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ -] -qed. - -lemma tps_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 [d, e] ▶ U2 → - ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & L ⊢ T1 [d, e] ▶ T2 & - U2 = ⓕ{I} V2. T2. -/2 width=3/ qed-. - -fact tps_inv_refl_O2_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶ T2 → e = 0 → T1 = T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ // -| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct - lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi -Hide (le_to_le_to_eq … Hdi ?) /2 width=1/ -d #K #V #HLK - lapply (ldrop_mono … HLK0 … HLK) #H destruct -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK - >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 K V) -IHT12 // /2 width=1/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK - >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 … HLK) -IHT12 // -] -qed. - -lemma tps_inv_refl_SO2: ∀L,T1,T2,d. L ⊢ T1 [d, 1] ▶ T2 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → T1 = T2. -/2 width=8/ qed-. - -(* Relocation properties ****************************************************) - -(* Basic_1: was: subst1_lift_lt *) -lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶ T2 → - ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt + et ≤ d → - L ⊢ U1 [dt, et] ▶ U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdetd - lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid - lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct - elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=4/ -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=6/ | @IHT12 /2 width=6/ ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ -] -qed. - -lemma tps_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶ T2 → - ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt ≤ d → d ≤ dt + et → - L ⊢ U1 [dt, et + e] ▶ U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdtd #_ - elim (lift_inv_lref1 … H) -H * #Hid #H destruct - [ -Hdtd - lapply (lt_to_le_to_lt … (dt+et+e) Hidet ?) // -Hidet #Hidete - elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=4/ - | -Hdti - lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti - lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ - ] -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=6/ | @IHT12 [3,4: // | skip |5,6: /2 width=1/ | /2 width=1/ ] - ] (**) (* /3 width=6/ is too slow, simplification like tps_lift_le is too slow *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ -] -qed. - -(* Basic_1: was: subst1_lift_ge *) -lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶ T2 → - ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - d ≤ dt → - L ⊢ U1 [dt + e, et] ▶ U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt - lapply (transitive_le … Hddt … Hdti) -Hddt #Hid - lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct - lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=5/ -] -qed. - -(* Basic_1: was: subst1_gen_lift_lt *) -lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶ U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. K ⊢ T1 [dt, et] ▶ T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et -[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ - ] -| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd - lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid - lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct - elim (ldrop_conf_lt … HLK … HLKV ?) -L // #L #U #HKL #_ #HUV - elim (lift_trans_le … HUV … HVW ?) -V // >minus_plus minus_plus plus_minus // commutative_plus >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) - ] -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet - elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (IHV12 … HLK … HWV1 ? ?) -V1 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ? ?) -U1 [5: @ldrop_skip // |2: skip |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1/ |4: /2 width=1/ ] (**) (* 29s without the rewrites *) - /3 width=5/ -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet - elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (IHV12 … HLK … HWV1 ? ?) -V1 // - elim (IHU12 … HLK … HTU1 ? ?) -U1 -HLK // /3 width=5/ -] -qed. - -(* Basic_1: was: subst1_gen_lift_ge *) -lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶ U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d + e ≤ dt → - ∃∃T2. K ⊢ T1 [dt - e, et] ▶ T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et -[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ - ] -| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt - lapply (transitive_le … Hdedt … Hdti) #Hdei - elim (le_inv_plus_l … Hdedt) -Hdedt #_ #Hedt - elim (le_inv_plus_l … Hdei) #Hdie #Hei - lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct - lapply (ldrop_conf_ge … HLK … HLKV ?) -L // #HKV - elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW [4: // |3: /2 width=1/ |2: /3 width=1/ ] -Hdei -Hdie - #V0 #HV10 >plus_minus // plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd - elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (le_inv_plus_l … Hdetd) #_ #Hedt - elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ?) -U1 [4: @ldrop_skip // |2: skip |3: /2 width=1/ ] - IHV12 // >IHT12 // -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct - >IHV12 // >IHT12 // -] -qed. -(* - Theorem subst0_gen_lift_rev_ge: (t1,v,u2,i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le (plus d h) i) -> - (EX u1 | (subst0 (minus i h) v u1 u2) & - t1 = (lift h d u1) - ). - - - Theorem subst0_gen_lift_rev_lelt: (t1,v,u2,i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le d i) -> (lt i (plus d h)) -> - (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). -*) -lemma tps_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶ U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ▶ T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (tps_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 -lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -U -L // (lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3/ - ] -| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - lapply (tps_lsubs_conf … HT02 (L. ⓑ{I} V1) ?) -HT02 /2 width=1/ #HT02 - elim (IHV01 … HV02) -V0 #V #HV1 #HV2 - elim (IHT01 … HT02) -T0 #T #HT1 #HT2 - lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ - lapply (tps_lsubs_conf … HT2 (L. ⓑ{I} V) ?) -HT2 /3 width=5/ -| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02) -V0 - elim (IHT01 … HT02) -T0 /3 width=5/ -] -qed. - -(* Basic_1: was: subst1_confluence_neq *) -theorem tps_conf_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 [d1, e1] ▶ T1 → - ∀L2,T2,d2,e2. L2 ⊢ T0 [d2, e2] ▶ T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. L2 ⊢ T1 [d2, e2] ▶ T & L1 ⊢ T2 [d1, e1] ▶ T. -#L1 #T0 #T1 #d1 #e1 #H elim H -L1 -T0 -T1 -d1 -e1 -[ /2 width=3/ -| #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 - elim (tps_inv_lref1 … H1) -H1 - [ #H destruct /4 width=4/ - | -HLK1 -HVT1 * #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded - [ -Hd1 -Hde2 - lapply (transitive_le … Hded Hd2) -Hded -Hd2 #H - lapply (lt_to_le_to_lt … Hde1 H) -Hde1 -H #H - elim (lt_refl_false … H) - | -Hd2 -Hde1 - lapply (transitive_le … Hded Hd1) -Hded -Hd1 #H - lapply (lt_to_le_to_lt … Hde2 H) -Hde2 -H #H - elim (lt_refl_false … H) - ] - ] -| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -V0 #V #HV1 #HV2 - elim (IHT01 … HT02 ?) -T0 - [ -H #T #HT1 #HT2 - lapply (tps_lsubs_conf … HT1 (L2. ⓑ{I} V) ?) -HT1 /2 width=1/ - lapply (tps_lsubs_conf … HT2 (L1. ⓑ{I} V) ?) -HT2 /2 width=1/ /3 width=5/ - | -HV1 -HV2 >plus_plus_comm_23 >plus_plus_comm_23 in ⊢ (? ? %); elim H -H #H - [ @or_introl | @or_intror ] /2 by monotonic_le_plus_l/ (**) (* /3 / is too slow *) - ] -| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -V0 - elim (IHT01 … HT02 H) -T0 -H /3 width=5/ -] -qed. - -(* Note: the constant 1 comes from tps_subst *) -(* Basic_1: was: subst1_trans *) -theorem tps_trans_ge: ∀L,T1,T0,d,e. L ⊢ T1 [d, e] ▶ T0 → - ∀T2. L ⊢ T0 [d, 1] ▶ T2 → 1 ≤ e → - L ⊢ T1 [d, e] ▶ T2. -#L #T1 #T0 #d #e #H elim H -L -T1 -T0 -d -e -[ #L #I #d #e #T2 #H #He - elim (tps_inv_atom1 … H) -H - [ #H destruct // - | * #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct - lapply (lt_to_le_to_lt … (d + e) Hide2 ?) /2 width=4/ - ] -| #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He - lapply (tps_weak … HVT2 0 (i +1) ? ?) -HVT2 /2 width=1/ #HVT2 - <(tps_inv_lift1_eq … HVT2 … HVW) -HVT2 /2 width=4/ -| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He - elim (tps_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct - lapply (tps_lsubs_conf … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 - lapply (IHT10 … HT02 He) -T0 #HT12 - lapply (tps_lsubs_conf … HT12 (L. ⓑ{I} V2) ?) -HT12 /2 width=1/ /3 width=1/ -| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He - elim (tps_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1/ -] -qed. - -theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ▶ T0 → - ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶ T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 [d2, e2] ▶ T & L ⊢ T [d1, e1] ▶ T2. -#L #T1 #T0 #d1 #e1 #H elim H -L -T1 -T0 -d1 -e1 -[ /2 width=3/ -| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 - lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1 - lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2 normalize /2 width=1/ -Hde2i1 #HWT2 - <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /4 width=4/ -| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - lapply (tps_lsubs_conf … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 - elim (IHV10 … HV02 ?) -IHV10 -HV02 // #V - elim (IHT10 … HT02 ?) -T0 /2 width=1/ #T #HT1 #HT2 - lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ - lapply (tps_lsubs_conf … HT2 (L. ⓑ{I} V2) ?) -HT2 /2 width=1/ /3 width=6/ -| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV10 … HV02 ?) -V0 // - elim (IHT10 … HT02 ?) -T0 // /3 width=6/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/delift.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/delift.ma deleted file mode 100644 index ec2d6c373..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/delift.ma +++ /dev/null @@ -1,84 +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/unfold/tpss.ma". - -(* DELIFT ON TERMS **********************************************************) - -definition delift: nat → nat → lenv → relation term ≝ - λd,e,L,T1,T2. ∃∃T. L ⊢ T1 [d, e] ▶* T & ⇧[d, e] T2 ≡ T. - -interpretation "delift (term)" - 'TSubst L T1 d e T2 = (delift d e L T1 T2). - -(* Basic properties *********************************************************) - -lemma delift_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≡ T2 → - ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ≡ T2. -#L1 #T1 #T2 #d #e * /3 width=3/ -qed. - -lemma delift_bind: ∀I,L,V1,V2,T1,T2,d,e. - L ⊢ V1 [d, e] ≡ V2 → L. ⓑ{I} V2 ⊢ T1 [d+1, e] ≡ T2 → - L ⊢ ⓑ{I} V1. T1 [d, e] ≡ ⓑ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * #T #HT1 #HT2 -lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ /3 width=5/ -qed. - -lemma delift_flat: ∀I,L,V1,V2,T1,T2,d,e. - L ⊢ V1 [d, e] ≡ V2 → L ⊢ T1 [d, e] ≡ T2 → - L ⊢ ⓕ{I} V1. T1 [d, e] ≡ ⓕ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * /3 width=5/ -qed. - -(* Basic forward lemmas *****************************************************) - -lemma delift_fwd_sort1: ∀L,U2,d,e,k. L ⊢ ⋆k [d, e] ≡ U2 → U2 = ⋆k. -#L #U2 #d #e #k * #U #HU ->(tpss_inv_sort1 … HU) -HU #HU2 ->(lift_inv_sort2 … HU2) -HU2 // -qed-. - -lemma delift_fwd_gref1: ∀L,U2,d,e,p. L ⊢ §p [d, e] ≡ U2 → U2 = §p. -#L #U #d #e #p * #U #HU ->(tpss_inv_gref1 … HU) -HU #HU2 ->(lift_inv_gref2 … HU2) -HU2 // -qed-. - -lemma delift_fwd_bind1: ∀I,L,V1,T1,U2,d,e. L ⊢ ⓑ{I} V1. T1 [d, e] ≡ U2 → - ∃∃V2,T2. L ⊢ V1 [d, e] ≡ V2 & - L. ⓑ{I} V2 ⊢ T1 [d+1, e] ≡ T2 & - U2 = ⓑ{I} V2. T2. -#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 -elim (tpss_inv_bind1 … HU) -HU #V #T #HV1 #HT1 #X destruct -elim (lift_inv_bind2 … HU2) -HU2 #V2 #T2 #HV2 #HT2 -lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ -qed-. - -lemma delift_fwd_flat1: ∀I,L,V1,T1,U2,d,e. L ⊢ ⓕ{I} V1. T1 [d, e] ≡ U2 → - ∃∃V2,T2. L ⊢ V1 [d, e] ≡ V2 & - L ⊢ T1 [d, e] ≡ T2 & - U2 = ⓕ{I} V2. T2. -#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 -elim (tpss_inv_flat1 … HU) -HU #V #T #HV1 #HT1 #X destruct -elim (lift_inv_flat2 … HU2) -HU2 /3 width=5/ -qed-. - -(* Basic Inversion lemmas ***************************************************) - -lemma delift_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ≡ T2 → T1 = T2. -#L #T1 #T2 #d * #T #HT1 ->(tpss_inv_refl_O2 … HT1) -HT1 #HT2 ->(lift_inv_refl_O2 … HT2) -HT2 // -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/delift_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/delift_lift.ma deleted file mode 100644 index a02d3db2c..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/delift_lift.ma +++ /dev/null @@ -1,54 +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/unfold/tpss_lift.ma". -include "Basic_2/unfold/delift.ma". - -(* DELIFT ON TERMS **********************************************************) - -(* Advanced forward lemmas **************************************************) - -lemma delift_fwd_lref1_lt: ∀L,U2,i,d,e. L ⊢ #i [d, e] ≡ U2 → i < d → U2 = #i. -#L #U2 #i #d #e * #U #HU #HU2 #Hid -elim (tpss_inv_lref1 … HU) -HU -[ #H destruct >(lift_inv_lref2_lt … HU2) // -| * #K #V1 #V2 #Hdi - lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi - elim (lt_refl_false … Hi) -] -qed-. - -lemma delift_fwd_lref1_be: ∀L,U2,d,e,i. L ⊢ #i [d, e] ≡ U2 → - d ≤ i → i < d + e → - ∃∃K,V1,V2. ⇩[0, i] L ≡ K. ⓓV1 & - K ⊢ V1 [0, d + e - i - 1] ≡ V2 & - ⇧[0, d] V2 ≡ U2. -#L #U2 #d #e #i * #U #HU #HU2 #Hdi #Hide -elim (tpss_inv_lref1 … HU) -HU -[ #H destruct elim (lift_inv_lref2_be … HU2 ? ?) // -| * #K #V1 #V #_ #_ #HLK #HV1 #HVU - elim (lift_div_be … HVU … HU2 ? ?) -U // /2 width=1/ /3 width=6/ -] -qed-. - -lemma delift_fwd_lref1_ge: ∀L,U2,i,d,e. L ⊢ #i [d, e] ≡ U2 → - d + e ≤ i → U2 = #(i - e). -#L #U2 #i #d #e * #U #HU #HU2 #Hdei -elim (tpss_inv_lref1 … HU) -HU -[ #H destruct >(lift_inv_lref2_ge … HU2) // -| * #K #V1 #V2 #_ #Hide - lapply (lt_to_le_to_lt … Hide Hdei) -Hide -Hdei #Hi - elim (lt_refl_false … Hi) -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2.ma deleted file mode 100644 index 0896b6ba3..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2.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/grammar/term_vector.ma". - -(* GENERIC RELOCATION WITH PAIRS ********************************************) - -inductive at: list2 nat nat → relation nat ≝ -| at_nil: ∀i. at ⟠ i i -| at_lt : ∀des,d,e,i1,i2. i1 < d → - at des i1 i2 → at ({d, e} :: des) i1 i2 -| at_ge : ∀des,d,e,i1,i2. d ≤ i1 → - at des (i1 + e) i2 → at ({d, e} :: des) i1 i2 -. - -interpretation "application (generic relocation with pairs)" - 'RAt i1 des i2 = (at des i1 i2). - -(* Basic inversion lemmas ***************************************************) - -fact at_inv_nil_aux: ∀des,i1,i2. @[i1] des ≡ i2 → des = ⟠ → i1 = i2. -#des #i1 #i2 * -des -i1 -i2 -[ // -| #des #d #e #i1 #i2 #_ #_ #H destruct -| #des #d #e #i1 #i2 #_ #_ #H destruct -] -qed. - -lemma at_inv_nil: ∀i1,i2. @[i1] ⟠ ≡ i2 → i1 = i2. -/2 width=3/ qed-. - -fact at_inv_cons_aux: ∀des,i1,i2. @[i1] des ≡ i2 → - ∀d,e,des0. des = {d, e} :: des0 → - i1 < d ∧ @[i1] des0 ≡ i2 ∨ - d ≤ i1 ∧ @[i1 + e] des0 ≡ i2. -#des #i1 #i2 * -des -i1 -i2 -[ #i #d #e #des #H destruct -| #des1 #d1 #e1 #i1 #i2 #Hid1 #Hi12 #d2 #e2 #des2 #H destruct /3 width=1/ -| #des1 #d1 #e1 #i1 #i2 #Hdi1 #Hi12 #d2 #e2 #des2 #H destruct /3 width=1/ -] -qed. - -lemma at_inv_cons: ∀des,d,e,i1,i2. @[i1] {d, e} :: des ≡ i2 → - i1 < d ∧ @[i1] des ≡ i2 ∨ - d ≤ i1 ∧ @[i1 + e] des ≡ i2. -/2 width=3/ qed-. - -lemma at_inv_cons_lt: ∀des,d,e,i1,i2. @[i1] {d, e} :: des ≡ i2 → - i1 < d → @[i1] des ≡ i2. -#des #d #e #i1 #e2 #H -elim (at_inv_cons … H) -H * // #Hdi1 #_ #Hi1d -lapply (le_to_lt_to_lt … Hdi1 Hi1d) -Hdi1 -Hi1d #Hd -elim (lt_refl_false … Hd) -qed-. - -lemma at_inv_cons_ge: ∀des,d,e,i1,i2. @[i1] {d, e} :: des ≡ i2 → - d ≤ i1 → @[i1 + e] des ≡ i2. -#des #d #e #i1 #e2 #H -elim (at_inv_cons … H) -H * // #Hi1d #_ #Hdi1 -lapply (le_to_lt_to_lt … Hdi1 Hi1d) -Hdi1 -Hi1d #Hd -elim (lt_refl_false … Hd) -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_gr2.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_gr2.ma deleted file mode 100644 index 438605e85..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_gr2.ma +++ /dev/null @@ -1,29 +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/unfold/gr2.ma". - -(* GENERIC RELOCATION WITH PAIRS ********************************************) - -(* Main properties **********************************************************) - -theorem at_mono: ∀des,i,i1. @[i] des ≡ i1 → ∀i2. @[i] des ≡ i2 → i1 = i2. -#des #i #i1 #H elim H -des -i -i1 -[ #i #x #H <(at_inv_nil … H) -x // -| #des #d #e #i #i1 #Hid #_ #IHi1 #x #H - lapply (at_inv_cons_lt … H Hid) -H -Hid /2 width=1/ -| #des #d #e #i #i1 #Hdi #_ #IHi1 #x #H - lapply (at_inv_cons_ge … H Hdi) -H -Hdi /2 width=1/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_minus.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_minus.ma deleted file mode 100644 index 5e3144c93..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_minus.ma +++ /dev/null @@ -1,76 +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/unfold/gr2.ma". - -(* GENERIC RELOCATION WITH PAIRS ********************************************) - -inductive minuss: nat → relation (list2 nat nat) ≝ -| minuss_nil: ∀i. minuss i ⟠ ⟠ -| minuss_lt : ∀des1,des2,d,e,i. i < d → minuss i des1 des2 → - minuss i ({d, e} :: des1) ({d - i, e} :: des2) -| minuss_ge : ∀des1,des2,d,e,i. d ≤ i → minuss (e + i) des1 des2 → - minuss i ({d, e} :: des1) des2 -. - -interpretation "minus (generic relocation with pairs)" - 'RMinus des1 i des2 = (minuss i des1 des2). - -(* Basic inversion lemmas ***************************************************) - -fact minuss_inv_nil1_aux: ∀des1,des2,i. des1 ▭ i ≡ des2 → des1 = ⟠ → des2 = ⟠. -#des1 #des2 #i * -des1 -des2 -i -[ // -| #des1 #des2 #d #e #i #_ #_ #H destruct -| #des1 #des2 #d #e #i #_ #_ #H destruct -] -qed. - -lemma minuss_inv_nil1: ∀des2,i. ⟠ ▭ i ≡ des2 → des2 = ⟠. -/2 width=4/ qed-. - -fact minuss_inv_cons1_aux: ∀des1,des2,i. des1 ▭ i ≡ des2 → - ∀d,e,des. des1 = {d, e} :: des → - d ≤ i ∧ des ▭ e + i ≡ des2 ∨ - ∃∃des0. i < d & des ▭ i ≡ des0 & - des2 = {d - i, e} :: des0. -#des1 #des2 #i * -des1 -des2 -i -[ #i #d #e #des #H destruct -| #des1 #des #d1 #e1 #i1 #Hid1 #Hdes #d2 #e2 #des2 #H destruct /3 width=3/ -| #des1 #des #d1 #e1 #i1 #Hdi1 #Hdes #d2 #e2 #des2 #H destruct /3 width=1/ -] -qed. - -lemma minuss_inv_cons1: ∀des1,des2,d,e,i. {d, e} :: des1 ▭ i ≡ des2 → - d ≤ i ∧ des1 ▭ e + i ≡ des2 ∨ - ∃∃des. i < d & des1 ▭ i ≡ des & - des2 = {d - i, e} :: des. -/2 width=3/ qed-. - -lemma minuss_inv_cons1_ge: ∀des1,des2,d,e,i. {d, e} :: des1 ▭ i ≡ des2 → - d ≤ i → des1 ▭ e + i ≡ des2. -#des1 #des2 #d #e #i #H -elim (minuss_inv_cons1 … H) -H * // #des #Hid #_ #_ #Hdi -lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi -elim (lt_refl_false … Hi) -qed-. - -lemma minuss_inv_cons1_lt: ∀des1,des2,d,e,i. {d, e} :: des1 ▭ i ≡ des2 → - i < d → - ∃∃des. des1 ▭ i ≡ des & des2 = {d - i, e} :: des. -#des1 #des2 #d #e #i #H -elim (minuss_inv_cons1 … H) -H * /2 width=3/ #Hdi #_ #Hid -lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi -elim (lt_refl_false … Hi) -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_plus.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_plus.ma deleted file mode 100644 index 938f955b9..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/gr2_plus.ma +++ /dev/null @@ -1,40 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/unfold/gr2.ma". - -(* GENERIC RELOCATION WITH PAIRS ********************************************) - -let rec pluss (des:list2 nat nat) (i:nat) on des ≝ match des with -[ nil2 ⇒ ⟠ -| cons2 d e des ⇒ {d + i, e} :: pluss des i -]. - -interpretation "plus (generic relocation with pairs)" - 'plus x y = (pluss x y). - -(* Basic inversion lemmas ***************************************************) - -lemma pluss_inv_nil2: ∀i,des. des + i = ⟠ → des = ⟠. -#i * // normalize -#d #e #des #H destruct -qed. - -lemma pluss_inv_cons2: ∀i,d,e,des2,des. des + i = {d, e} :: des2 → - ∃∃des1. des1 + i = des2 & des = {d - i, e} :: des1. -#i #d #e #des2 * normalize -[ #H destruct -| #d1 #e1 #des1 #H destruct /2 width=3/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops.ma deleted file mode 100644 index 28b9a8c7e..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops.ma +++ /dev/null @@ -1,90 +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/ldrop.ma". -include "Basic_2/unfold/gr2_minus.ma". -include "Basic_2/unfold/lifts.ma". - -(* GENERIC LOCAL ENVIRONMENT SLICING ****************************************) - -inductive ldrops: list2 nat nat → relation lenv ≝ -| ldrops_nil : ∀L. ldrops ⟠ L L -| ldrops_cons: ∀L1,L,L2,des,d,e. - ldrops des L1 L → ⇩[d,e] L ≡ L2 → ldrops ({d, e} :: des) L1 L2 -. - -interpretation "generic local environment slicing" - 'RDropStar des T1 T2 = (ldrops des T1 T2). - -(* Basic inversion lemmas ***************************************************) - -fact ldrops_inv_nil_aux: ∀L1,L2,des. ⇩*[des] L1 ≡ L2 → des = ⟠ → L1 = L2. -#L1 #L2 #des * -L1 -L2 -des // -#L1 #L #L2 #d #e #des #_ #_ #H destruct -qed. - -(* Basic_1: was: drop1_gen_pnil *) -lemma ldrops_inv_nil: ∀L1,L2. ⇩*[⟠] L1 ≡ L2 → L1 = L2. -/2 width=3/ qed-. - -fact ldrops_inv_cons_aux: ∀L1,L2,des. ⇩*[des] L1 ≡ L2 → - ∀d,e,tl. des = {d, e} :: tl → - ∃∃L. ⇩*[tl] L1 ≡ L & ⇩[d, e] L ≡ L2. -#L1 #L2 #des * -L1 -L2 -des -[ #L #d #e #tl #H destruct -| #L1 #L #L2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct - /2 width=3/ -qed. - -(* Basic_1: was: drop1_gen_pcons *) -lemma ldrops_inv_cons: ∀L1,L2,d,e,des. ⇩*[{d, e} :: des] L1 ≡ L2 → - ∃∃L. ⇩*[des] L1 ≡ L & ⇩[d, e] L ≡ L2. -/2 width=3/ qed-. - -lemma ldrops_inv_skip2: ∀I,des,i,des2. des ▭ i ≡ des2 → - ∀L1,K2,V2. ⇩*[des2] L1 ≡ K2. ⓑ{I} V2 → - ∃∃K1,V1,des1. des + 1 ▭ i + 1 ≡ des1 + 1 & - ⇩*[des1] K1 ≡ K2 & - ⇧*[des1] V2 ≡ V1 & - L1 = K1. ⓑ{I} V1. -#I #des #i #des2 #H elim H -des -i -des2 -[ #i #L1 #K2 #V2 #H - >(ldrops_inv_nil … H) -L1 /2 width=7/ -| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H - elim (ldrops_inv_cons … H) -H #L #HL1 #H - elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ #K #V >minus_plus #HK2 #HV2 #H destruct - elim (IHdes2 … HL1) -IHdes2 -HL1 #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct - @(ex4_3_intro … K1 V1 … ) // [3,4: /2 width=7/ | skip ] - normalize >plus_minus // @minuss_lt // /2 width=1/ (**) (* explicit constructors, /3 width=1/ is a bit slow *) -| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H - elim (IHdes2 … H) -IHdes2 -H #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct - /4 width=7/ -] -qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: drop1_skip_bind *) -lemma ldrops_skip: ∀L1,L2,des. ⇩*[des] L1 ≡ L2 → ∀V1,V2. ⇧*[des] V2 ≡ V1 → - ∀I. ⇩*[des + 1] L1. ⓑ{I} V1 ≡ L2. ⓑ{I} V2. -#L1 #L2 #des #H elim H -L1 -L2 -des -[ #L #V1 #V2 #HV12 #I - >(lifts_inv_nil … HV12) -HV12 // -| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #V1 #V2 #H #I - elim (lifts_inv_cons … H) -H /3 width=5/ -]. -qed. - -(* Basic_1: removed theorems 1: drop1_getl_trans -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops_ldrop.ma deleted file mode 100644 index de8a1fef6..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops_ldrop.ma +++ /dev/null @@ -1,35 +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/ldrop_ldrop.ma". -include "Basic_2/unfold/ldrops.ma". - -(* GENERIC LOCAL ENVIRONMENT SLICING ****************************************) - -(* Properties concerning basic local environment slicing ********************) - -lemma ldrops_ldrop_trans: ∀L1,L,des. ⇩*[des] L1 ≡ L → ∀L2,i. ⇩[0, i] L ≡ L2 → - ∃∃L0,des0,i0. ⇩[0, i0] L1 ≡ L0 & ⇩*[des0] L0 ≡ L2 & - @[i] des ≡ i0 & des ▭ i ≡ des0. -#L1 #L #des #H elim H -L1 -L -des -[ /2 width=7/ -| #L1 #L3 #L #des3 #d #e #_ #HL3 #IHL13 #L2 #i #HL2 - elim (lt_or_ge i d) #Hid - [ elim (ldrop_trans_le … HL3 … HL2 ?) -L /2 width=2/ #L #HL3 #HL2 - elim (IHL13 … HL3) -L3 /3 width=7/ - | lapply (ldrop_trans_ge … HL3 … HL2 ?) -L // #HL32 - elim (IHL13 … HL32) -L3 /3 width=7/ - ] -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops_ldrops.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops_ldrops.ma deleted file mode 100644 index 1bb40cb5d..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ldrops_ldrops.ma +++ /dev/null @@ -1,25 +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/unfold/ldrops_ldrop.ma". - -(* GENERIC LOCAL ENVIRONMENT SLICING ****************************************) - -(* Main properties **********************************************************) - -(* Basic_1: was: drop1_trans *) -theorem ldrops_trans: ∀L,L2,des2. ⇩*[des2] L ≡ L2 → ∀L1,des1. ⇩*[des1] L1 ≡ L → - ⇩*[des2 @ des1] L1 ≡ L2. -#L #L2 #des2 #H elim H -L -L2 -des2 // /3 width=3/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts.ma deleted file mode 100644 index 6a3c647a0..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts.ma +++ /dev/null @@ -1,150 +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/lift.ma". -include "Basic_2/unfold/gr2_plus.ma". - -(* GENERIC TERM RELOCATION **************************************************) - -inductive lifts: list2 nat nat → relation term ≝ -| lifts_nil : ∀T. lifts ⟠ T T -| lifts_cons: ∀T1,T,T2,des,d,e. - ⇧[d,e] T1 ≡ T → lifts des T T2 → lifts ({d, e} :: des) T1 T2 -. - -interpretation "generic relocation (term)" - 'RLiftStar des T1 T2 = (lifts des T1 T2). - -(* Basic inversion lemmas ***************************************************) - -fact lifts_inv_nil_aux: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → des = ⟠ → T1 = T2. -#T1 #T2 #des * -T1 -T2 -des // -#T1 #T #T2 #d #e #des #_ #_ #H destruct -qed. - -lemma lifts_inv_nil: ∀T1,T2. ⇧*[⟠] T1 ≡ T2 → T1 = T2. -/2 width=3/ qed-. - -fact lifts_inv_cons_aux: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → - ∀d,e,tl. des = {d, e} :: tl → - ∃∃T. ⇧[d, e] T1 ≡ T & ⇧*[tl] T ≡ T2. -#T1 #T2 #des * -T1 -T2 -des -[ #T #d #e #tl #H destruct -| #T1 #T #T2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct - /2 width=3/ -qed. - -lemma lifts_inv_cons: ∀T1,T2,d,e,des. ⇧*[{d, e} :: des] T1 ≡ T2 → - ∃∃T. ⇧[d, e] T1 ≡ T & ⇧*[des] T ≡ T2. -/2 width=3/ qed-. - -(* Basic_1: was: lift1_sort *) -lemma lifts_inv_sort1: ∀T2,k,des. ⇧*[des] ⋆k ≡ T2 → T2 = ⋆k. -#T2 #k #des elim des -des -[ #H <(lifts_inv_nil … H) -H // -| #d #e #des #IH #H - elim (lifts_inv_cons … H) -H #X #H - >(lift_inv_sort1 … H) -H /2 width=1/ -] -qed-. - -(* Basic_1: was: lift1_lref *) -lemma lifts_inv_lref1: ∀T2,des,i1. ⇧*[des] #i1 ≡ T2 → - ∃∃i2. @[i1] des ≡ i2 & T2 = #i2. -#T2 #des elim des -des -[ #i1 #H <(lifts_inv_nil … H) -H /2 width=3/ -| #d #e #des #IH #i1 #H - elim (lifts_inv_cons … H) -H #X #H1 #H2 - elim (lift_inv_lref1 … H1) -H1 * #Hdi1 #H destruct - elim (IH … H2) -IH -H2 /3 width=3/ -] -qed-. - -lemma lifts_inv_gref1: ∀T2,p,des. ⇧*[des] §p ≡ T2 → T2 = §p. -#T2 #p #des elim des -des -[ #H <(lifts_inv_nil … H) -H // -| #d #e #des #IH #H - elim (lifts_inv_cons … H) -H #X #H - >(lift_inv_gref1 … H) -H /2 width=1/ -] -qed-. - -(* Basic_1: was: lift1_bind *) -lemma lifts_inv_bind1: ∀I,T2,des,V1,U1. ⇧*[des] ⓑ{I} V1. U1 ≡ T2 → - ∃∃V2,U2. ⇧*[des] V1 ≡ V2 & ⇧*[des + 1] U1 ≡ U2 & - T2 = ⓑ{I} V2. U2. -#I #T2 #des elim des -des -[ #V1 #U1 #H - <(lifts_inv_nil … H) -H /2 width=5/ -| #d #e #des #IHdes #V1 #U1 #H - elim (lifts_inv_cons … H) -H #X #H #HT2 - elim (lift_inv_bind1 … H) -H #V #U #HV1 #HU1 #H destruct - elim (IHdes … HT2) -IHdes -HT2 #V2 #U2 #HV2 #HU2 #H destruct - /3 width=5/ -] -qed-. - -(* Basic_1: was: lift1_flat *) -lemma lifts_inv_flat1: ∀I,T2,des,V1,U1. ⇧*[des] ⓕ{I} V1. U1 ≡ T2 → - ∃∃V2,U2. ⇧*[des] V1 ≡ V2 & ⇧*[des] U1 ≡ U2 & - T2 = ⓕ{I} V2. U2. -#I #T2 #des elim des -des -[ #V1 #U1 #H - <(lifts_inv_nil … H) -H /2 width=5/ -| #d #e #des #IHdes #V1 #U1 #H - elim (lifts_inv_cons … H) -H #X #H #HT2 - elim (lift_inv_flat1 … H) -H #V #U #HV1 #HU1 #H destruct - elim (IHdes … HT2) -IHdes -HT2 #V2 #U2 #HV2 #HU2 #H destruct - /3 width=5/ -] -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lifts_simple_dx: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → 𝐒[T1] → 𝐒[T2]. -#T1 #T2 #des #H elim H -T1 -T2 -des // /3 width=5 by lift_simple_dx/ -qed-. - -lemma lifts_simple_sn: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → 𝐒[T2] → 𝐒[T1]. -#T1 #T2 #des #H elim H -T1 -T2 -des // /3 width=5 by lift_simple_sn/ -qed-. - -(* Basic properties *********************************************************) - -lemma lifts_bind: ∀I,T2,V1,V2,des. ⇧*[des] V1 ≡ V2 → - ∀T1. ⇧*[des + 1] T1 ≡ T2 → - ⇧*[des] ⓑ{I} V1. T1 ≡ ⓑ{I} V2. T2. -#I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des -[ #V #T1 #H >(lifts_inv_nil … H) -H // -| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H - elim (lifts_inv_cons … H) -H /3 width=3/ -] -qed. - -lemma lifts_flat: ∀I,T2,V1,V2,des. ⇧*[des] V1 ≡ V2 → - ∀T1. ⇧*[des] T1 ≡ T2 → - ⇧*[des] ⓕ{I} V1. T1 ≡ ⓕ{I} V2. T2. -#I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des -[ #V #T1 #H >(lifts_inv_nil … H) -H // -| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H - elim (lifts_inv_cons … H) -H /3 width=3/ -] -qed. - -lemma lifts_total: ∀des,T1. ∃T2. ⇧*[des] T1 ≡ T2. -#des elim des -des /2 width=2/ -#d #e #des #IH #T1 -elim (lift_total T1 d e) #T #HT1 -elim (IH T) -IH /3 width=4/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_lift.ma deleted file mode 100644 index 3f63d2eb8..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_lift.ma +++ /dev/null @@ -1,59 +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/lift_lift.ma". -include "Basic_2/unfold/gr2_minus.ma". -include "Basic_2/unfold/lifts.ma". - -(* GENERIC TERM RELOCATION **************************************************) - -(* Properties concerning basic term relocation ******************************) - -(* Basic_1: was: lift1_xhg (right to left) *) -lemma lifts_lift_trans_le: ∀T1,T,des. ⇧*[des] T1 ≡ T → ∀T2. ⇧[0, 1] T ≡ T2 → - ∃∃T0. ⇧[0, 1] T1 ≡ T0 & ⇧*[des + 1] T0 ≡ T2. -#T1 #T #des #H elim H -T1 -T -des -[ /2 width=3/ -| #T1 #T3 #T #des #d #e #HT13 #_ #IHT13 #T2 #HT2 - elim (IHT13 … HT2) -T #T #HT3 #HT2 - elim (lift_trans_le … HT13 … HT3 ?) -T3 // /3 width=5/ -] -qed-. - -(* Basic_1: was: lift1_free (right to left) *) -lemma lifts_lift_trans: ∀des,i,i0. @[i] des ≡ i0 → - ∀des0. des + 1 ▭ i + 1 ≡ des0 + 1 → - ∀T1,T0. ⇧*[des0] T1 ≡ T0 → - ∀T2. ⇧[O, i0 + 1] T0 ≡ T2 → - ∃∃T. ⇧[0, i + 1] T1 ≡ T & ⇧*[des] T ≡ T2. -#des elim des -des normalize -[ #i #x #H1 #des0 #H2 #T1 #T0 #HT10 #T2 - <(at_inv_nil … H1) -x #HT02 - lapply (minuss_inv_nil1 … H2) -H2 #H - >(pluss_inv_nil2 … H) in HT10; -des0 #H - >(lifts_inv_nil … H) -T1 /2 width=3/ -| #d #e #des #IHdes #i #i0 #H1 #des0 #H2 #T1 #T0 #HT10 #T2 #HT02 - elim (at_inv_cons … H1) -H1 * #Hid #Hi0 - [ elim (minuss_inv_cons1_lt … H2 ?) -H2 [2: /2 width=1/ ] #des1 #Hdes1 minus_plus #HT1 #HT0 - elim (IHdes … Hi0 … Hdes1 … HT0 … HT02) -IHdes -Hi0 -Hdes1 -T0 #T0 #HT0 #HT02 - elim (lift_trans_le … HT1 … HT0 ?) -T /2 width=1/ #T #HT1 commutative_plus in Hi0; #Hi0 - lapply (minuss_inv_cons1_ge … H2 ?) -H2 [ /2 width=1/ ] (liftv_inv_nil1 … H) -T1s /2 width=3/ -| #T1s #Ts #T1 #T #HT1 #_ #IHT1s #X #H - elim (liftv_inv_cons1 … H) -H #T2 #T2s #HT2 #HT2s #H destruct - elim (IHT1s … HT2s) -Ts #Ts #HT1s #HT2s - elim (lifts_lift_trans_le … HT1 … HT2) -T /3 width=5/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_lifts.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_lifts.ma deleted file mode 100644 index 6062f89b1..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_lifts.ma +++ /dev/null @@ -1,25 +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/unfold/lifts_lift.ma". - -(* GENERIC RELOCATION *******************************************************) - -(* Main properties **********************************************************) - -(* Basic_1: was: lift1_lift1 (left to right) *) -theorem lifts_trans: ∀T1,T,des1. ⇧*[des1] T1 ≡ T → ∀T2:term. ∀des2. ⇧*[des2] T ≡ T2 → - ⇧*[des1 @ des2] T1 ≡ T2. -#T1 #T #des1 #H elim H -T1 -T -des1 // /3 width=3/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_vector.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_vector.ma deleted file mode 100644 index 2bd579d01..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/lifts_vector.ma +++ /dev/null @@ -1,51 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "Basic_2/substitution/lift_vector.ma". -include "Basic_2/unfold/lifts.ma". - -(* GENERIC TERM VECTOR RELOCATION *******************************************) - -inductive liftsv (des:list2 nat nat) : relation (list term) ≝ -| liftsv_nil : liftsv des ◊ ◊ -| liftsv_cons: ∀T1s,T2s,T1,T2. - ⇧*[des] T1 ≡ T2 → liftsv des T1s T2s → - liftsv des (T1 :: T1s) (T2 :: T2s) -. - -interpretation "generic relocation (vector)" - 'RLiftStar des T1s T2s = (liftsv des T1s T2s). - -(* Basic inversion lemmas ***************************************************) - -lemma lifts_inv_applv1: ∀V1s,U1,T2,des. ⇧*[des] Ⓐ V1s. U1 ≡ T2 → - ∃∃V2s,U2. ⇧*[des] V1s ≡ V2s & ⇧*[des] U1 ≡ U2 & - T2 = Ⓐ V2s. U2. -#V1s elim V1s -V1s normalize -[ #T1 #T2 #des #HT12 - @(ex3_2_intro) [3,4: // |1,2: skip | // ] (**) (* explicit constructor *) -| #V1 #V1s #IHV1s #T1 #X #des #H - elim (lifts_inv_flat1 … H) -H #V2 #Y #HV12 #HY #H destruct - elim (IHV1s … HY) -IHV1s -HY #V2s #T2 #HV12s #HT12 #H destruct - @(ex3_2_intro) [4: // |3: /2 width=2/ |1,2: skip | // ] (**) (* explicit constructor *) -] -qed-. - -(* Basic properties *********************************************************) - -lemma lifts_applv: ∀V1s,V2s,des. ⇧*[des] V1s ≡ V2s → - ∀T1,T2. ⇧*[des] T1 ≡ T2 → - ⇧*[des] Ⓐ V1s. T1 ≡ Ⓐ V2s. T2. -#V1s #V2s #des #H elim H -V1s -V2s // /3 width=1/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss.ma deleted file mode 100644 index 207bfc60f..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss.ma +++ /dev/null @@ -1,115 +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/ltps.ma". -include "Basic_2/unfold/tpss.ma". - -(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) - -definition ltpss: nat → nat → relation lenv ≝ - λd,e. TC … (ltps d e). - -interpretation "partial unfold (local environment)" - 'PSubstStar L1 d e L2 = (ltpss d e L1 L2). - -(* Basic eliminators ********************************************************) - -lemma ltpss_ind: ∀d,e,L1. ∀R:predicate lenv. R L1 → - (∀L,L2. L1 [d, e] ▶* L → L [d, e] ▶ L2 → R L → R L2) → - ∀L2. L1 [d, e] ▶* L2 → R L2. -#d #e #L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // -qed-. - -(* Basic properties *********************************************************) - -lemma ltpss_strap: ∀L1,L,L2,d,e. - L1 [d, e] ▶ L → L [d, e] ▶* L2 → L1 [d, e] ▶* L2. -/2 width=3/ qed. - -lemma ltpss_refl: ∀L,d,e. L [d, e] ▶* L. -/2 width=1/ qed. - -lemma ltpss_weak_all: ∀L1,L2,d,e. L1 [d, e] ▶* L2 → L1 [0, |L2|] ▶* L2. -#L1 #L2 #d #e #H @(ltpss_ind … H) -L2 // -#L #L2 #_ #HL2 ->(ltps_fwd_length … HL2) /3 width=5/ -qed. - -(* Basic forward lemmas *****************************************************) - -lemma ltpss_fwd_length: ∀L1,L2,d,e. L1 [d, e] ▶* L2 → |L1| = |L2|. -#L1 #L2 #d #e #H @(ltpss_ind … H) -L2 // -#L #L2 #_ #HL2 #IHL12 >IHL12 -IHL12 -/2 width=3 by ltps_fwd_length/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma ltpss_inv_refl_O2: ∀d,L1,L2. L1 [d, 0] ▶* L2 → L1 = L2. -#d #L1 #L2 #H @(ltpss_ind … H) -L2 // -#L #L2 #_ #HL2 #IHL <(ltps_inv_refl_O2 … HL2) -HL2 // -qed-. - -lemma ltpss_inv_atom1: ∀d,e,L2. ⋆ [d, e] ▶* L2 → L2 = ⋆. -#d #e #L2 #H @(ltpss_ind … H) -L2 // -#L #L2 #_ #HL2 #IHL destruct ->(ltps_inv_atom1 … HL2) -HL2 // -qed-. - -fact ltpss_inv_atom2_aux: ∀d,e,L1,L2. L1 [d, e] ▶* L2 → L2 = ⋆ → L1 = ⋆. -#d #e #L1 #L2 #H @(ltpss_ind … H) -L2 // -#L2 #L #_ #HL2 #IHL2 #H destruct -lapply (ltps_inv_atom2 … HL2) -HL2 /2 width=1/ -qed. - -lemma ltpss_inv_atom2: ∀d,e,L1. L1 [d, e] ▶* ⋆ → L1 = ⋆. -/2 width=5/ qed-. -(* -fact ltps_inv_tps22_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → d = 0 → 0 < e → - ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & - K2 ⊢ V1 [0, e - 1] ▶ V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d e L1 L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H elim (plus_S_eq_O_false … H) -] -qed. - -lemma ltps_inv_tps22: ∀e,L1,K2,I,V2. L1 [0, e] ▶ K2. ⓑ{I} V2 → 0 < e → - ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & K2 ⊢ V1 [0, e - 1] ▶ V2 & - L1 = K1. ⓑ{I} V1. -/2 width=5/ qed. - -fact ltps_inv_tps12_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → 0 < d → - ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & - K2 ⊢ V1 [d - 1, e] ▶ V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d e L1 L2 -[ #d #e #_ #I #K2 #V2 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ -] -qed. - -lemma ltps_inv_tps12: ∀L1,K2,I,V2,d,e. L1 [d, e] ▶ K2. ⓑ{I} V2 → 0 < d → - ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & - K2 ⊢ V1 [d - 1, e] ▶ V2 & - L1 = K1. ⓑ{I} V1. -/2 width=1/ qed. -*) diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ldrop.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ldrop.ma deleted file mode 100644 index b52cedae7..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ldrop.ma +++ /dev/null @@ -1,74 +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/ltps_ldrop.ma". -include "Basic_2/unfold/ltpss.ma". - -(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) - -lemma ltpss_ldrop_conf_ge: ∀L0,L1,d1,e1. L0 [d1, e1] ▶* L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. -#L0 #L1 #d1 #e1 #H @(ltpss_ind … H) -L1 // /3 width=6/ -qed. - -lemma ltpss_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. -#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 // /3 width=6/ -qed. - -lemma ltpss_ldrop_conf_be: ∀L0,L1,d1,e1. L0 [d1, e1] ▶* L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → - ∃∃L. L2 [0, d1 + e1 - e2] ▶* L & ⇩[0, e2] L1 ≡ L. -#L0 #L1 #d1 #e1 #H @(ltpss_ind … H) -L1 -[ /2 width=3/ -| #L #L1 #_ #HL1 #IHL #L2 #e2 #HL02 #Hd1e2 #He2de1 - elim (IHL … HL02 Hd1e2 He2de1) -L0 #L0 #HL20 #HL0 - elim (ltps_ldrop_conf_be … HL1 … HL0 Hd1e2 He2de1) -L /3 width=3/ -] -qed. - -lemma ltpss_ldrop_trans_be: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → - ∃∃L. L [0, d1 + e1 - e2] ▶* L2 & ⇩[0, e2] L1 ≡ L. -#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 -[ /2 width=3/ -| #L #L0 #_ #HL0 #IHL #L2 #e2 #HL02 #Hd1e2 #He2de1 - elim (ltps_ldrop_trans_be … HL0 … HL02 Hd1e2 He2de1) -L0 #L0 #HL02 #HL0 - elim (IHL … HL0 Hd1e2 He2de1) -L /3 width=3/ -] -qed. - -lemma ltpss_ldrop_conf_le: ∀L0,L1,d1,e1. L0 [d1, e1] ▶* L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 → - ∃∃L. L2 [d1 - e2, e1] ▶* L & ⇩[0, e2] L1 ≡ L. -#L0 #L1 #d1 #e1 #H @(ltpss_ind … H) -L1 -[ /2 width=3/ -| #L #L1 #_ #HL1 #IHL #L2 #e2 #HL02 #He2d1 - elim (IHL … HL02 He2d1) -L0 #L0 #HL20 #HL0 - elim (ltps_ldrop_conf_le … HL1 … HL0 He2d1) -L /3 width=3/ -] -qed. - -lemma ltpss_ldrop_trans_le: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 → - ∃∃L. L [d1 - e2, e1] ▶* L2 & ⇩[0, e2] L1 ≡ L. -#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 -[ /2 width=3/ -| #L #L0 #_ #HL0 #IHL #L2 #e2 #HL02 #He2d1 - elim (ltps_ldrop_trans_le … HL0 … HL02 He2d1) -L0 #L0 #HL02 #HL0 - elim (IHL … HL0 He2d1) -L /3 width=3/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma deleted file mode 100644 index 7091e0803..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_ltpss.ma +++ /dev/null @@ -1,55 +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/unfold/ltpss_tpss.ma". - -(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) - -(* Main properties **********************************************************) - -theorem ltpss_trans_eq: ∀L1,L,L2,d,e. - L1 [d, e] ▶* L → L [d, e] ▶* L2 → L1 [d, e] ▶* L2. -/2 width=3/ qed. - -lemma ltpss_tpss2: ∀L1,L2,I,V1,V2,e. - L1 [0, e] ▶* L2 → L2 ⊢ V1 [0, e] ▶* V2 → - L1. ⓑ{I} V1 [0, e + 1] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #H @(tpss_ind … H) -V2 -[ /2 width=1/ -| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2 width=1/ -] -qed. - -lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e. - L1 [0, e - 1] ▶* L2 → L2 ⊢ V1 [0, e - 1] ▶* V2 → - 0 < e → L1. ⓑ{I} V1 [0, e] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) // /2 width=1/ -qed. - -lemma ltpss_tpss1: ∀L1,L2,I,V1,V2,d,e. - L1 [d, e] ▶* L2 → L2 ⊢ V1 [d, e] ▶* V2 → - L1. ⓑ{I} V1 [d + 1, e] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #H @(tpss_ind … H) -V2 -[ /2 width=1/ -| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2 width=1/ -] -qed. - -lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 [d - 1, e] ▶* L2 → L2 ⊢ V1 [d - 1, e] ▶* V2 → - 0 < d → L1. ⓑ{I} V1 [d, e] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) // /2 width=1/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_tpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_tpss.ma deleted file mode 100644 index 24f1a595e..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/ltpss_tpss.ma +++ /dev/null @@ -1,169 +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/unfold/tpss_ltps.ma". -include "Basic_2/unfold/ltpss.ma". - -(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) - -(* Properties concerning partial unfold on terms ****************************) - -lemma ltpss_tpss_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → - ∀T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶* U2 → - d1 + e1 ≤ d2 → L1 ⊢ T2 [d2, e2] ▶* U2. -#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 // -#L #L0 #_ #HL0 #IHL #T2 #U2 #d2 #e2 #HTU2 #Hde1d2 -lapply (ltps_tpss_trans_ge … HL0 HTU2) -HL0 -HTU2 /2 width=1/ -qed. - -lemma ltpss_tps_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → - ∀T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 → - d1 + e1 ≤ d2 → L1 ⊢ T2 [d2, e2] ▶* U2. -#L1 #L0 #d1 #e1 #HL10 #T2 #U2 #d2 #e2 #HTU2 #Hde1d2 -@(ltpss_tpss_trans_ge … HL10 … Hde1d2) /2 width=1/ (**) (* /3 width=6/ is too slow *) -qed. - -lemma ltpss_tpss_trans_eq: ∀L0,L1,d,e. L0 [d, e] ▶* L1 → - ∀T2,U2. L1 ⊢ T2 [d, e] ▶* U2 → L0 ⊢ T2 [d, e] ▶* U2. -#L0 #L1 #d #e #H @(ltpss_ind … H) -L1 -[ /2 width=1/ -| #L #L1 #_ #HL1 #IHL #T2 #U2 #HTU2 - lapply (ltps_tpss_trans_eq … HL1 HTU2) -HL1 -HTU2 /2 width=1/ -] -qed. - -lemma ltpss_tps_trans_eq: ∀L0,L1,d,e. L0 [d, e] ▶* L1 → - ∀T2,U2. L1 ⊢ T2 [d, e] ▶ U2 → L0 ⊢ T2 [d, e] ▶* U2. -/3 width=3/ qed. - -lemma ltpss_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. d1 + e1 ≤ d2 → - L0 ⊢ T2 [d2, e2] ▶* U2 → L0 [d1, e1] ▶* L1 → - L1 ⊢ T2 [d2, e2] ▶* U2. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HTU2 #H @(ltpss_ind … H) -L1 -[ // -| -HTU2 #L #L1 #_ #HL1 #IHL - lapply (ltps_tpss_conf_ge … HL1 IHL) -HL1 -IHL // -] -qed. - -lemma ltpss_tps_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. d1 + e1 ≤ d2 → - L0 ⊢ T2 [d2, e2] ▶ U2 → L0 [d1, e1] ▶* L1 → - L1 ⊢ T2 [d2, e2] ▶* U2. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HTU2 #HL01 -@(ltpss_tpss_conf_ge … Hde1d2 … HL01) /2 width=1/ (**) (* /3 width=6/ is too slow *) -qed. - -lemma ltpss_tpss_conf_eq: ∀L0,L1,T2,U2,d,e. - L0 ⊢ T2 [d, e] ▶* U2 → L0 [d, e] ▶* L1 → - ∃∃T. L1 ⊢ T2 [d, e] ▶* T & L1 ⊢ U2 [d, e] ▶* T. -#L0 #L1 #T2 #U2 #d #e #HTU2 #H @(ltpss_ind … H) -L1 -[ /2 width=3/ -| -HTU2 #L #L1 #_ #HL1 * #W2 #HTW2 #HUW2 - elim (ltps_tpss_conf … HL1 HTW2) -HTW2 #T #HT2 #HW2T - elim (ltps_tpss_conf … HL1 HUW2) -HL1 -HUW2 #U #HU2 #HW2U - elim (tpss_conf_eq … HW2T … HW2U) -HW2T -HW2U #V #HTV #HUV - lapply (tpss_trans_eq … HT2 … HTV) -T - lapply (tpss_trans_eq … HU2 … HUV) -U /2 width=3/ -] -qed. - -lemma ltpss_tps_conf_eq: ∀L0,L1,T2,U2,d,e. - L0 ⊢ T2 [d, e] ▶ U2 → L0 [d, e] ▶* L1 → - ∃∃T. L1 ⊢ T2 [d, e] ▶* T & L1 ⊢ U2 [d, e] ▶* T. -/3 width=3/ qed. - -lemma ltpss_tpss_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶* T2 → - ∀L2,ds,es. L1 [ds, es] ▶* L2 → - ∃∃T. L2 ⊢ T1 [d, e] ▶* T & L1 ⊢ T2 [ds, es] ▶* T. -#L1 #T1 #T2 #d #e #HT12 #L2 #ds #es #H @(ltpss_ind … H) -L2 -[ /3 width=3/ -| #L #L2 #HL1 #HL2 * #T #HT1 #HT2 - elim (ltps_tpss_conf … HL2 HT1) -HT1 #T0 #HT10 #HT0 - lapply (ltps_tpss_trans_eq … HL2 … HT0) -HL2 -HT0 #HT0 - lapply (ltpss_tpss_trans_eq … HL1 … HT0) -HL1 -HT0 #HT0 - lapply (tpss_trans_eq … HT2 … HT0) -T /2 width=3/ -] -qed. - -lemma ltpss_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 → - ∀L2,ds,es. L1 [ds, es] ▶* L2 → - ∃∃T. L2 ⊢ T1 [d, e] ▶* T & L1 ⊢ T2 [ds, es] ▶* T. -/3 width=1/ qed. - -(* Advanced properties ******************************************************) - -lemma ltpss_tps2: ∀L1,L2,I,e. - L1 [0, e] ▶* L2 → ∀V1,V2. L2 ⊢ V1 [0, e] ▶ V2 → - L1. ⓑ{I} V1 [0, e + 1] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #e #H @(ltpss_ind … H) -L2 -[ /3 width=1/ -| #L #L2 #_ #HL2 #IHL #V1 #V2 #HV12 - elim (ltps_tps_trans … HV12 … HL2) -HV12 #V #HV1 #HV2 - lapply (IHL … HV1) -IHL -HV1 #HL1 - @step /2 width=5/ (**) (* /3 width=5/ is too slow *) -] -qed. - -lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e. - L1 [0, e - 1] ▶* L2 → L2 ⊢ V1 [0, e - 1] ▶ V2 → - 0 < e → L1. ⓑ{I} V1 [0, e] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) // /2 width=1/ -qed. - -lemma ltpss_tps1: ∀L1,L2,I,d,e. L1 [d, e] ▶* L2 → - ∀V1,V2. L2 ⊢ V1 [d, e] ▶ V2 → - L1. ⓑ{I} V1 [d + 1, e] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #d #e #H @(ltpss_ind … H) -L2 -[ /3 width=1/ -| #L #L2 #_ #HL2 #IHL #V1 #V2 #HV12 - elim (ltps_tps_trans … HV12 … HL2) -HV12 #V #HV1 #HV2 - lapply (IHL … HV1) -IHL -HV1 #HL1 - @step /2 width=5/ (**) (* /3 width=5/ is too slow *) -] -qed. - -lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 [d - 1, e] ▶* L2 → L2 ⊢ V1 [d - 1, e] ▶ V2 → - 0 < d → L1. ⓑ{I} V1 [d, e] ▶* L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) // /2 width=1/ -qed. - -(* Advanced forward lemmas **************************************************) - -lemma ltpss_fwd_tpss21: ∀e,K1,I,V1,L2. 0 < e → K1. ⓑ{I} V1 [0, e] ▶* L2 → - ∃∃K2,V2. K1 [0, e - 1] ▶* K2 & K1 ⊢ V1 [0, e - 1] ▶* V2 & - L2 = K2. ⓑ{I} V2. -#e #K1 #I #V1 #L2 #He #H @(ltpss_ind … H) -L2 -[ /2 width=5/ -| #L #L2 #_ #HL2 * #K #V #HK1 #HV1 #H destruct - elim (ltps_inv_tps21 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H - lapply (ltps_tps_trans_eq … HV2 … HK2) -HV2 #HV2 - lapply (ltpss_tpss_trans_eq … HK1 … HV2) -HV2 #HV2 /3 width=5/ -] -qed-. - -lemma ltpss_fwd_tpss11: ∀d,e,I,K1,V1,L2. 0 < d → K1. ⓑ{I} V1 [d, e] ▶* L2 → - ∃∃K2,V2. K1 [d - 1, e] ▶* K2 & - K1 ⊢ V1 [d - 1, e] ▶* V2 & - L2 = K2. ⓑ{I} V2. -#d #e #K1 #I #V1 #L2 #Hd #H @(ltpss_ind … H) -L2 -[ /2 width=5/ -| #L #L2 #_ #HL2 * #K #V #HK1 #HV1 #H destruct - elim (ltps_inv_tps11 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H - lapply (ltps_tps_trans_eq … HV2 … HK2) -HV2 #HV2 - lapply (ltpss_tpss_trans_eq … HK1 … HV2) -HV2 #HV2 /3 width=5/ -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss.ma deleted file mode 100644 index 16a0d3e4d..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss.ma +++ /dev/null @@ -1,146 +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/tps.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -definition tpss: nat → nat → lenv → relation term ≝ - λd,e,L. TC … (tps d e L). - -interpretation "partial unfold (term)" - 'PSubstStar L T1 d e T2 = (tpss d e L T1 T2). - -(* Basic eliminators ********************************************************) - -lemma tpss_ind: ∀d,e,L,T1. ∀R:predicate term. R T1 → - (∀T,T2. L ⊢ T1 [d, e] ▶* T → L ⊢ T [d, e] ▶ T2 → R T → R T2) → - ∀T2. L ⊢ T1 [d, e] ▶* T2 → R T2. -#d #e #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // -qed-. - -(* Basic properties *********************************************************) - -lemma tpss_strap: ∀L,T1,T,T2,d,e. - L ⊢ T1 [d, e] ▶ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 [d, e] ▶* T2. -/2 width=3/ qed. - -lemma tpss_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶* T2 → - ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ▶* T2. -/3 width=3/ qed. - -lemma tpss_refl: ∀d,e,L,T. L ⊢ T [d, e] ▶* T. -/2 width=1/ qed. - -lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 [d, e] ▶* V2 → - ∀I,T1,T2. L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶* T2 → - L ⊢ ⓑ{I} V1. T1 [d, e] ▶* ⓑ{I} V2. T2. -#L #V1 #V2 #d #e #HV12 elim HV12 -V2 -[ #V2 #HV12 #I #T1 #T2 #HT12 elim HT12 -T2 - [ /3 width=5/ - | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) - ] -| #V #V2 #_ #HV12 #IHV #I #T1 #T2 #HT12 - lapply (tpss_lsubs_conf … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12 - lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) -] -qed. - -lemma tpss_flat: ∀L,I,V1,V2,T1,T2,d,e. - L ⊢ V1 [d, e] ▶ * V2 → L ⊢ T1 [d, e] ▶* T2 → - L ⊢ ⓕ{I} V1. T1 [d, e] ▶* ⓕ{I} V2. T2. -#L #I #V1 #V2 #T1 #T2 #d #e #HV12 elim HV12 -V2 -[ #V2 #HV12 #HT12 elim HT12 -T2 - [ /3 width=1/ - | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) - ] -| #V #V2 #_ #HV12 #IHV #HT12 - lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) -] -qed. - -lemma tpss_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ▶* T2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → - L ⊢ T1 [d2, e2] ▶* T2. -#L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT12 #IHT - lapply (tps_weak … HT12 … Hd21 Hde12) -HT12 -Hd21 -Hde12 /2 width=3/ -] -qed. - -lemma tpss_weak_top: ∀L,T1,T2,d,e. - L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 [d, |L| - d] ▶* T2. -#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT12 #IHT - lapply (tps_weak_top … HT12) -HT12 /2 width=3/ -] -qed. - -lemma tpss_weak_all: ∀L,T1,T2,d,e. - L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 [0, |L|] ▶* T2. -#L #T1 #T2 #d #e #HT12 -lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 -lapply (tpss_weak_top … HT12) // -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Note: this can be derived from tpss_inv_atom1 *) -lemma tpss_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k [d, e] ▶* T2 → T2 = ⋆k. -#L #T2 #k #d #e #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT2 #IHT destruct - >(tps_inv_sort1 … HT2) -HT2 // -] -qed-. - -(* Note: this can be derived from tpss_inv_atom1 *) -lemma tpss_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p [d, e] ▶* T2 → T2 = §p. -#L #T2 #p #d #e #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT2 #IHT destruct - >(tps_inv_gref1 … HT2) -HT2 // -] -qed-. - -lemma tpss_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 [d, e] ▶* U2 → - ∃∃V2,T2. L ⊢ V1 [d, e] ▶* V2 & - L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶* T2 & - U2 = ⓑ{I} V2. T2. -#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 -[ /2 width=5/ -| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct - elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H - lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ -] -qed-. - -lemma tpss_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 [d, e] ▶* U2 → - ∃∃V2,T2. L ⊢ V1 [d, e] ▶* V2 & L ⊢ T1 [d, e] ▶* T2 & - U2 = ⓕ{I} V2. T2. -#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 -[ /2 width=5/ -| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct - elim (tps_inv_flat1 … HU2) -HU2 /3 width=5/ -] -qed-. - -lemma tpss_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ▶* T2 → T1 = T2. -#L #T1 #T2 #d #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT2 #IHT <(tps_inv_refl_O2 … HT2) -HT2 // -] -qed-. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_lift.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_lift.ma deleted file mode 100644 index 748efd827..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_lift.ma +++ /dev/null @@ -1,166 +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/tps_lift.ma". -include "Basic_2/unfold/tpss.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -(* Advanced properties ******************************************************) - -lemma tpss_subst: ∀L,K,V,U1,i,d,e. - d ≤ i → i < d + e → - ⇩[0, i] L ≡ K. ⓓV → K ⊢ V [0, d + e - i - 1] ▶* U1 → - ∀U2. ⇧[0, i + 1] U1 ≡ U2 → L ⊢ #i [d, e] ▶* U2. -#L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(tpss_ind … H) -U1 -[ /3 width=4/ -| #U #U1 #_ #HU1 #IHU #U2 #HU12 - elim (lift_total U 0 (i+1)) #U0 #HU0 - lapply (IHU … HU0) -IHU #H - lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK - lapply (tps_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // normalize #HU02 - lapply (tps_weak … HU02 d e ? ?) -HU02 [ >minus_plus >commutative_plus /2 width=1/ | /2 width=1/ | /2 width=3/ ] -] -qed. - -(* Advanced inverion lemmas *************************************************) - -lemma tpss_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} [d, e] ▶* T2 → - T2 = ⓪{I} ∨ - ∃∃K,V1,V2,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV1 & - K ⊢ V1 [0, d + e - i - 1] ▶* V2 & - ⇧[O, i + 1] V2 ≡ T2 & - I = LRef i. -#L #T2 #I #d #e #H @(tpss_ind … H) -T2 -[ /2 width=1/ -| #T #T2 #_ #HT2 * - [ #H destruct - elim (tps_inv_atom1 … HT2) -HT2 [ /2 width=1/ | * /3 width=10/ ] - | * #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI - lapply (ldrop_fwd_ldrop2 … HLK) #H - elim (tps_inv_lift1_up … HT2 … H … HVT ? ? ?) normalize -HT2 -H -HVT [2,3,4: /2 width=1/ ] #V2 (lift_mono … HTU1 … H) -H // -| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (IHT … HTU) -IHT #HU1 - lapply (tps_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ -] -qed. - -lemma tpss_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶* T2 → - ∀L,U1,d,e. dt ≤ d → d ≤ dt + et → - ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → - ∀U2. ⇧[d, e] T2 ≡ U2 → L ⊢ U1 [dt, et + e] ▶* U2. -#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hdtd #Hddet #HLK #HTU1 @(tpss_ind … H) -T2 -[ #U2 #H >(lift_mono … HTU1 … H) -H // -| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (IHT … HTU) -IHT #HU1 - lapply (tps_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ -] -qed. - -lemma tpss_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶* T2 → - ∀L,U1,d,e. d ≤ dt → ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → - L ⊢ U1 [dt + e, et] ▶* U2. -#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hddt #HLK #HTU1 @(tpss_ind … H) -T2 -[ #U2 #H >(lift_mono … HTU1 … H) -H // -| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (IHT … HTU) -IHT #HU1 - lapply (tps_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ -] -qed. - -lemma tpss_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. K ⊢ T1 [dt, et] ▶* T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdetd @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_le … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_be: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 [dt, et - e] ▶* T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_be … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d + e ≤ dt → - ∃∃T2. K ⊢ T1 [dt - e, et] ▶* T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdedt @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_ge … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_eq: ∀L,U1,U2,d,e. - L ⊢ U1 [d, e] ▶* U2 → ∀T1. ⇧[d, e] T1 ≡ U1 → U1 = U2. -#L #U1 #U2 #d #e #H #T1 #HTU1 @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU destruct -<(tps_inv_lift1_eq … HU2 … HTU1) -HU2 -HTU1 // -qed. - -lemma tpss_inv_lift1_be_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → dt + et ≤ d + e → - ∃∃T2. K ⊢ T1 [dt, d - dt] ▶* T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_ltps.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_ltps.ma deleted file mode 100644 index 58b63cd01..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_ltps.ma +++ /dev/null @@ -1,95 +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/ltps_tps.ma". -include "Basic_2/unfold/tpss_tpss.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -(* Properties concerning parallel substitution on local environments ********) - -lemma ltps_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. - d1 + e1 ≤ d2 → L0 [d1, e1] ▶ L1 → - L0 ⊢ T2 [d2, e2] ▶* U2 → L1 ⊢ T2 [d2, e2] ▶* U2. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL01 #H @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU -lapply (ltps_tps_conf_ge … HU2 … HL01 ?) -HU2 -HL01 // /2 width=3/ -qed. - -lemma ltps_tpss_conf: ∀L0,L1,T2,U2,d1,e1,d2,e2. - L0 [d1, e1] ▶ L1 → L0 ⊢ T2 [d2, e2] ▶* U2 → - ∃∃T. L1 ⊢ T2 [d2, e2] ▶* T & L1 ⊢ U2 [d1, e1] ▶* T. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #HL01 #H @(tpss_ind … H) -U2 -[ /3 width=3/ -| #U #U2 #_ #HU2 * #T #HT2 #HUT - elim (ltps_tps_conf … HU2 … HL01) -HU2 -HL01 #W #HUW #HU2W - elim (tpss_strip_eq … HUT … HUW) -U - /3 width=5 by ex2_1_intro, step, tpss_strap/ (**) (* just /3 width=5/ is too slow *) -] -qed. - -lemma ltps_tpss_trans_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. - d1 + e1 ≤ d2 → L1 [d1, e1] ▶ L0 → - L0 ⊢ T2 [d2, e2] ▶* U2 → L1 ⊢ T2 [d2, e2] ▶* U2. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL10 #H @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU -lapply (ltps_tps_trans_ge … HU2 … HL10 ?) -HU2 -HL10 // /2 width=3/ -qed. - -lemma ltps_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 → - L1 [d1, e1] ▶ L0 → L0 ⊢ T2 [d2, e2] ▶* U2 → - ∃∃T. L1 ⊢ T2 [d2, e2] ▶* T & L0 ⊢ T [d1, e1] ▶* U2. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde2d1 #HL10 #H @(tpss_ind … H) -U2 -[ /3 width=3/ -| #U #U2 #_ #HU2 * #T #HT2 #HTU - elim (tpss_strap1_down … HTU … HU2 ?) -U // #U #HTU #HU2 - elim (ltps_tps_trans … HTU … HL10) -HTU -HL10 #W #HTW #HWU - @(ex2_1_intro … W) /2 width=3/ (**) (* /3 width=5/ does not work as in ltps_tpss_conf *) -] -qed. - -fact ltps_tps_trans_eq_aux: ∀Y1,X2,L1,T2,U2,d,e. - L1 ⊢ T2 [d, e] ▶ U2 → ∀L0. L0 [d, e] ▶ L1 → - Y1 = L1 → X2 = T2 → L0 ⊢ T2 [d, e] ▶* U2. -#Y1 #X2 @(cw_wf_ind … Y1 X2) -Y1 -X2 #Y1 #X2 #IH -#L1 #T2 #U2 #d #e * -L1 -T2 -U2 -d -e -[ // -| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L0 #HL10 #H1 #H2 destruct - lapply (ldrop_fwd_lw … HLK1) normalize #H1 - elim (ltps_ldrop_trans_be … HL10 … HLK1 ? ?) -HL10 -HLK1 // /2 width=2/ #X #H #HLK0 - elim (ltps_inv_tps22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct - lapply (tps_fwd_tw … HV01) #H2 - lapply (transitive_le (#[K1] + #[V0]) … H1) -H1 /2 width=1/ -H2 #H - lapply (IH … HV01 … HK01 ? ?) -IH -HV01 -HK01 - [1,3: // |2,4: skip | normalize /2 width=1/ | /3 width=6/ ] -| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct - lapply (tps_lsubs_conf … HT12 (L. ⓑ{I} V1) ?) -HT12 /2 width=1/ #HT12 - lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=2/ ] #HV12 - lapply (IH … HT12 (L0. ⓑ{I} V1) ? ? ?) -IH -HT12 [1,3,5: /2 width=2/ |2,4: skip | normalize // ] -HL0 #HT12 - lapply (tpss_lsubs_conf … HT12 (L0. ⓑ{I} V2) ?) -HT12 /2 width=1/ -| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct - lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=3/ ] - lapply (IH … HT12 … HL0 ? ?) -IH -HT12 [1,3,5: normalize // |2,4: skip ] -HL0 /2 width=1/ -] -qed. - -lemma ltps_tps_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 [d, e] ▶ U2 → - ∀L0. L0 [d, e] ▶ L1 → L0 ⊢ T2 [d, e] ▶* U2. -/2 width=5/ qed. - -lemma ltps_tpss_trans_eq: ∀L0,L1,T2,U2,d,e. L0 [d, e] ▶ L1 → - L1 ⊢ T2 [d, e] ▶* U2 → L0 ⊢ T2 [d, e] ▶* U2. -#L0 #L1 #T2 #U2 #d #e #HL01 #H @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU @(tpss_trans_eq … IHU) /2 width=3/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_tpss.ma b/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_tpss.ma deleted file mode 100644 index 981402a24..000000000 --- a/matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_tpss.ma +++ /dev/null @@ -1,93 +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/tps_tps.ma". -include "Basic_2/unfold/tpss_lift.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -(* Advanced properties ******************************************************) - -lemma tpss_tps: ∀L,T1,T2,d. L ⊢ T1 [d, 1] ▶* T2 → L ⊢ T1 [d, 1] ▶ T2. -#L #T1 #T2 #d #H @(tpss_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 -lapply (tps_trans_ge … IHT1 … HT2 ?) // -qed. - -lemma tpss_strip_eq: ∀L,T0,T1,d1,e1. L ⊢ T0 [d1, e1] ▶* T1 → - ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶ T2 → - ∃∃T. L ⊢ T1 [d2, e2] ▶ T & L ⊢ T2 [d1, e1] ▶* T. -/3 width=3/ qed. - -lemma tpss_strip_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 [d1, e1] ▶* T1 → - ∀L2,T2,d2,e2. L2 ⊢ T0 [d2, e2] ▶ T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. L2 ⊢ T1 [d2, e2] ▶ T & L1 ⊢ T2 [d1, e1] ▶* T. -/3 width=3/ qed. - -lemma tpss_strap1_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ▶* T0 → - ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶ T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 [d2, e2] ▶ T & L ⊢ T [d1, e1] ▶* T2. -/3 width=3/ qed. - -lemma tpss_strap2_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ▶ T0 → - ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶* T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 [d2, e2] ▶* T & L ⊢ T [d1, e1] ▶ T2. -/3 width=3/ qed. - -lemma tpss_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶* T2 → - ∀i. d ≤ i → i ≤ d + e → - ∃∃T. L ⊢ T1 [d, i - d] ▶* T & L ⊢ T [i, d + e - i] ▶* T2. -#L #T1 #T2 #d #e #H #i #Hdi #Hide @(tpss_ind … H) -T2 -[ /2 width=3/ -| #T #T2 #_ #HT12 * #T3 #HT13 #HT3 - elim (tps_split_up … HT12 … Hdi Hide) -HT12 -Hide #T0 #HT0 #HT02 - elim (tpss_strap1_down … HT3 … HT0 ?) -T [2: >commutative_plus /2 width=1/ ] - /3 width=7 by ex2_1_intro, step/ (**) (* just /3 width=7/ is too slow *) -] -qed. - -lemma tpss_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ▶* T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (tpss_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (tpss_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 -lapply (tpss_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -elim (tpss_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 -HLK -HTU1 // minus_minus_comm >minus_le_minus_minus_comm // -qed. - -lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b. -#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1/ -qed. - -lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. -/3 by monotonic_le_minus_l, le_to_le_to_eq, le_n/ qed. - -lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b → - a1 - c1 + a2 - (b - c1) = a1 + a2 - b. -#a1 #a2 #b #c1 #H1 #H2 >plus_minus // /2 width=1/ -qed. - -(* inversion & forward lemmas ***********************************************) - -axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). - -axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). - -lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m. -#m #n elim (lt_or_ge m n) /2 width=1/ -#H elim H -m /2 width=1/ -#m #Hm * #H /2 width=1/ /3 width=1/ -qed-. - -lemma lt_refl_false: ∀n. n < n → False. -#n #H elim (lt_to_not_eq … H) -H /2 width=1/ -qed-. - -lemma lt_zero_false: ∀n. n < 0 → False. -#n #H elim (lt_to_not_le … H) -H /2 width=1/ -qed-. - -lemma false_lt_to_le: ∀x,y. (x < y → False) → y ≤ x. -#x #y #H elim (decidable_lt x y) /2 width=1/ -#Hxy elim (H Hxy) -qed-. - -(* -lemma pippo: ∀x,y,z. x < z → y < z - x → x + y < z. -/3 width=2/ - -lemma le_or_ge: ∀m,n. m ≤ n ∨ n ≤ m. -#m #n elim (lt_or_ge m n) /2 width=1/ /3 width=2/ -qed-. -*) diff --git a/matita/matita/contribs/lambda_delta/Ground_2/list.ma b/matita/matita/contribs/lambda_delta/Ground_2/list.ma deleted file mode 100644 index 0a6e69bbe..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/list.ma +++ /dev/null @@ -1,55 +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/arith.ma". - -(* LISTS ********************************************************************) - -inductive list (A:Type[0]) : Type[0] := - | nil : list A - | cons: A → list A → list A. - -interpretation "nil (list)" 'Nil = (nil ?). - -interpretation "cons (list)" 'Cons hd tl = (cons ? hd tl). - -let rec all A (R:predicate A) (l:list A) on l ≝ - match l with - [ nil ⇒ True - | cons hd tl ⇒ R hd ∧ all A R tl - ]. - -inductive list2 (A1,A2:Type[0]) : Type[0] := - | nil2 : list2 A1 A2 - | cons2: A1 → A2 → list2 A1 A2 → list2 A1 A2. - -interpretation "nil (list of pairs)" 'Nil2 = (nil2 ? ?). - -interpretation "cons (list of pairs)" 'Cons hd1 hd2 tl = (cons2 ? ? hd1 hd2 tl). - -let rec append2 (A1,A2:Type[0]) (l1,l2:list2 A1 A2) on l1 ≝ match l1 with -[ nil2 ⇒ l2 -| cons2 a1 a2 tl ⇒ {a1, a2} :: append2 A1 A2 tl l2 -]. - -interpretation "append (list of pairs)" - 'Append l1 l2 = (append2 ? ? l1 l2). - -let rec length2 (A1,A2:Type[0]) (l:list2 A1 A2) on l ≝ match l with -[ nil2 ⇒ 0 -| cons2 _ _ l ⇒ length2 A1 A2 l + 1 -]. - -interpretation "length (list of pairs)" - 'card l = (length2 ? ? l). diff --git a/matita/matita/contribs/lambda_delta/Ground_2/notation.ma b/matita/matita/contribs/lambda_delta/Ground_2/notation.ma deleted file mode 100644 index 71fbbe0c2..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/notation.ma +++ /dev/null @@ -1,37 +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 *) -(* *) -(**************************************************************************) - -(* GENERAL NOTATION USED BY THE FORMAL SYSTEM λδ ****************************) - -(* Lists ********************************************************************) - -notation "hvbox( ◊ )" - non associative with precedence 90 - for @{'Nil}. - -notation "hvbox( hd :: break tl )" - right associative with precedence 47 - for @{'Cons $hd $tl}. - -notation "hvbox( l1 @ break l2 )" - right associative with precedence 47 - for @{'Append $l1 $l2 }. - -notation "hvbox( ⟠ )" - non associative with precedence 90 - for @{'Nil2}. - -notation "hvbox( { hd1 , break hd2 } :: break tl )" - non associative with precedence 47 - for @{'Cons $hd1 $hd2 $tl}. diff --git a/matita/matita/contribs/lambda_delta/Ground_2/star.ma b/matita/matita/contribs/lambda_delta/Ground_2/star.ma deleted file mode 100644 index ed3580642..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/star.ma +++ /dev/null @@ -1,112 +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 "basics/star.ma". -include "Ground_2/xoa_props.ma". -include "Ground_2/notation.ma". - -(* PROPERTIES OF RELATIONS **************************************************) - -definition Decidable: Prop → Prop ≝ λR. R ∨ (R → False). - -definition confluent2: ∀A. ∀R1,R2: relation A. Prop ≝ λA,R1,R2. - ∀a0,a1. R1 a0 a1 → ∀a2. R2 a0 a2 → - ∃∃a. R2 a1 a & R1 a2 a. - -definition transitive2: ∀A. ∀R1,R2: relation A. Prop ≝ λA,R1,R2. - ∀a1,a0. R1 a1 a0 → ∀a2. R2 a0 a2 → - ∃∃a. R2 a1 a & R1 a a2. - -lemma TC_strip1: ∀A,R1,R2. confluent2 A R1 R2 → - ∀a0,a1. TC … R1 a0 a1 → ∀a2. R2 a0 a2 → - ∃∃a. R2 a1 a & TC … R1 a2 a. -#A #R1 #R2 #HR12 #a0 #a1 #H elim H -a1 -[ #a1 #Ha01 #a2 #Ha02 - elim (HR12 … Ha01 … Ha02) -HR12 -a0 /3 width=3/ -| #a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 - elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20 - elim (HR12 … Ha1 … Ha0) -HR12 -a /4 width=3/ -] -qed. - -lemma TC_strip2: ∀A,R1,R2. confluent2 A R1 R2 → - ∀a0,a2. TC … R2 a0 a2 → ∀a1. R1 a0 a1 → - ∃∃a. TC … R2 a1 a & R1 a2 a. -#A #R1 #R2 #HR12 #a0 #a2 #H elim H -a2 -[ #a2 #Ha02 #a1 #Ha01 - elim (HR12 … Ha01 … Ha02) -HR12 -a0 /3 width=3/ -| #a #a2 #_ #Ha2 #IHa0 #a1 #Ha01 - elim (IHa0 … Ha01) -a0 #a0 #Ha10 #Ha0 - elim (HR12 … Ha0 … Ha2) -HR12 -a /4 width=3/ -] -qed. - -lemma TC_confluent2: ∀A,R1,R2. - confluent2 A R1 R2 → confluent2 A (TC … R1) (TC … R2). -#A #R1 #R2 #HR12 #a0 #a1 #H elim H -a1 -[ #a1 #Ha01 #a2 #Ha02 - elim (TC_strip2 … HR12 … Ha02 … Ha01) -HR12 -a0 /3 width=3/ -| #a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 - elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20 - elim (TC_strip2 … HR12 … Ha0 … Ha1) -HR12 -a /4 width=3/ -] -qed. - -lemma TC_strap1: ∀A,R1,R2. transitive2 A R1 R2 → - ∀a1,a0. TC … R1 a1 a0 → ∀a2. R2 a0 a2 → - ∃∃a. R2 a1 a & TC … R1 a a2. -#A #R1 #R2 #HR12 #a1 #a0 #H elim H -a0 -[ #a0 #Ha10 #a2 #Ha02 - elim (HR12 … Ha10 … Ha02) -HR12 -a0 /3 width=3/ -| #a #a0 #_ #Ha0 #IHa #a2 #Ha02 - elim (HR12 … Ha0 … Ha02) -HR12 -a0 #a0 #Ha0 #Ha02 - elim (IHa … Ha0) -a /4 width=3/ -] -qed. - -lemma TC_strap2: ∀A,R1,R2. transitive2 A R1 R2 → - ∀a0,a2. TC … R2 a0 a2 → ∀a1. R1 a1 a0 → - ∃∃a. TC … R2 a1 a & R1 a a2. -#A #R1 #R2 #HR12 #a0 #a2 #H elim H -a2 -[ #a2 #Ha02 #a1 #Ha10 - elim (HR12 … Ha10 … Ha02) -HR12 -a0 /3 width=3/ -| #a #a2 #_ #Ha02 #IHa #a1 #Ha10 - elim (IHa … Ha10) -a0 #a0 #Ha10 #Ha0 - elim (HR12 … Ha0 … Ha02) -HR12 -a /4 width=3/ -] -qed. - -lemma TC_transitive2: ∀A,R1,R2. - transitive2 A R1 R2 → transitive2 A (TC … R1) (TC … R2). -#A #R1 #R2 #HR12 #a1 #a0 #H elim H -a0 -[ #a0 #Ha10 #a2 #Ha02 - elim (TC_strap2 … HR12 … Ha02 … Ha10) -HR12 -a0 /3 width=3/ -| #a #a0 #_ #Ha0 #IHa #a2 #Ha02 - elim (TC_strap2 … HR12 … Ha02 … Ha0) -HR12 -a0 #a0 #Ha0 #Ha02 - elim (IHa … Ha0) -a /4 width=3/ -] -qed. - -definition NF: ∀A. relation A → relation A → predicate A ≝ - λA,R,S,a1. ∀a2. R a1 a2 → S a1 a2. - -inductive SN (A) (R,S:relation A): predicate A ≝ -| SN_intro: ∀a1. (∀a2. R a1 a2 → (S a1 a2 → False) → SN A R S a2) → SN A R S a1 -. - -lemma NF_to_SN: ∀A,R,S,a. NF A R S a → SN A R S a. -#A #R #S #a1 #Ha1 -@SN_intro #a2 #HRa12 #HSa12 -elim (HSa12 ?) -HSa12 /2 width=1/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Ground_2/tri.ma b/matita/matita/contribs/lambda_delta/Ground_2/tri.ma deleted file mode 100644 index 005806c11..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/tri.ma +++ /dev/null @@ -1,44 +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/arith.ma". - -(* TRICOTOMY FUNCTION *******************************************************) - -let rec tri (A:Type[0]) m n a b c on m : A ≝ - match m with - [ O ⇒ match n with [ O ⇒ b | S n ⇒ a ] - | S m ⇒ match n with [ O ⇒ c | S n ⇒ tri A m n a b c ] - ]. - -(* Basic properties *********************************************************) - -lemma tri_lt: ∀A,a,b,c,n,m. m < n → tri A m n a b c = a. -#A #a #b #c #n elim n -n -[ #m #H elim (lt_zero_false … H) -| #n #IH #m elim m -m // /3 width=1/ -] -qed. - -lemma tri_eq: ∀A,a,b,c,m. tri A m m a b c = b. -#A #a #b #c #m elim m -m normalize // -qed. - -lemma tri_gt: ∀A,a,b,c,m,n. n < m → tri A m n a b c = c. -#A #a #b #c #m elim m -m -[ #n #H elim (lt_zero_false … H) -| #m #IH #n elim n -n // /3 width=1/ -] -qed. - diff --git a/matita/matita/contribs/lambda_delta/Ground_2/xoa.conf.xml b/matita/matita/contribs/lambda_delta/Ground_2/xoa.conf.xml deleted file mode 100644 index 2a2da2a5c..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/xoa.conf.xml +++ /dev/null @@ -1,37 +0,0 @@ - - -
- $(MATITA_RT_BASE_DIR) - -
-
- contribs/lambda_delta/Ground_2/ - xoa - xoa_notation - basics/pts.ma - 1 2 - 2 1 - 2 2 - 2 3 - 3 1 - 3 2 - 3 3 - 4 2 - 4 3 - 4 4 - 5 2 - 5 3 - 5 4 - 6 4 - 6 6 - 7 6 - 3 - 4 - 3 -
-
diff --git a/matita/matita/contribs/lambda_delta/Ground_2/xoa.ma b/matita/matita/contribs/lambda_delta/Ground_2/xoa.ma deleted file mode 100644 index 5a3a40024..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/xoa.ma +++ /dev/null @@ -1,175 +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 *) -(* *) -(**************************************************************************) - -(* This file was generated by xoa.native: do not edit *********************) - -include "basics/pts.ma". - -(* multiple existental quantifier (1, 2) *) - -inductive ex1_2 (A0,A1:Type[0]) (P0:A0→A1→Prop) : Prop ≝ - | ex1_2_intro: ∀x0,x1. P0 x0 x1 → ex1_2 ? ? ? -. - -interpretation "multiple existental quantifier (1, 2)" 'Ex P0 = (ex1_2 ? ? P0). - -(* multiple existental quantifier (2, 1) *) - -inductive ex2_1 (A0:Type[0]) (P0,P1:A0→Prop) : Prop ≝ - | ex2_1_intro: ∀x0. P0 x0 → P1 x0 → ex2_1 ? ? ? -. - -interpretation "multiple existental quantifier (2, 1)" 'Ex P0 P1 = (ex2_1 ? P0 P1). - -(* multiple existental quantifier (2, 2) *) - -inductive ex2_2 (A0,A1:Type[0]) (P0,P1:A0→A1→Prop) : Prop ≝ - | ex2_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → ex2_2 ? ? ? ? -. - -interpretation "multiple existental quantifier (2, 2)" 'Ex P0 P1 = (ex2_2 ? ? P0 P1). - -(* multiple existental quantifier (2, 3) *) - -inductive ex2_3 (A0,A1,A2:Type[0]) (P0,P1:A0→A1→A2→Prop) : Prop ≝ - | ex2_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → ex2_3 ? ? ? ? ? -. - -interpretation "multiple existental quantifier (2, 3)" 'Ex P0 P1 = (ex2_3 ? ? ? P0 P1). - -(* multiple existental quantifier (3, 1) *) - -inductive ex3_1 (A0:Type[0]) (P0,P1,P2:A0→Prop) : Prop ≝ - | ex3_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → ex3_1 ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 1)" 'Ex P0 P1 P2 = (ex3_1 ? P0 P1 P2). - -(* multiple existental quantifier (3, 2) *) - -inductive ex3_2 (A0,A1:Type[0]) (P0,P1,P2:A0→A1→Prop) : Prop ≝ - | ex3_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → ex3_2 ? ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 2)" 'Ex P0 P1 P2 = (ex3_2 ? ? P0 P1 P2). - -(* multiple existental quantifier (3, 3) *) - -inductive ex3_3 (A0,A1,A2:Type[0]) (P0,P1,P2:A0→A1→A2→Prop) : Prop ≝ - | ex3_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → ex3_3 ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (3, 3)" 'Ex P0 P1 P2 = (ex3_3 ? ? ? P0 P1 P2). - -(* multiple existental quantifier (4, 2) *) - -inductive ex4_2 (A0,A1:Type[0]) (P0,P1,P2,P3:A0→A1→Prop) : Prop ≝ - | ex4_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → P3 x0 x1 → ex4_2 ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (4, 2)" 'Ex P0 P1 P2 P3 = (ex4_2 ? ? P0 P1 P2 P3). - -(* multiple existental quantifier (4, 3) *) - -inductive ex4_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3:A0→A1→A2→Prop) : Prop ≝ - | ex4_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → ex4_3 ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (4, 3)" 'Ex P0 P1 P2 P3 = (ex4_3 ? ? ? P0 P1 P2 P3). - -(* multiple existental quantifier (4, 4) *) - -inductive ex4_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3:A0→A1→A2→A3→Prop) : Prop ≝ - | ex4_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → ex4_4 ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (4, 4)" 'Ex P0 P1 P2 P3 = (ex4_4 ? ? ? ? P0 P1 P2 P3). - -(* multiple existental quantifier (5, 2) *) - -inductive ex5_2 (A0,A1:Type[0]) (P0,P1,P2,P3,P4:A0→A1→Prop) : Prop ≝ - | ex5_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → P3 x0 x1 → P4 x0 x1 → ex5_2 ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (5, 2)" 'Ex P0 P1 P2 P3 P4 = (ex5_2 ? ? P0 P1 P2 P3 P4). - -(* multiple existental quantifier (5, 3) *) - -inductive ex5_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→Prop) : Prop ≝ - | ex5_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → P4 x0 x1 x2 → ex5_3 ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (5, 3)" 'Ex P0 P1 P2 P3 P4 = (ex5_3 ? ? ? P0 P1 P2 P3 P4). - -(* multiple existental quantifier (5, 4) *) - -inductive ex5_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→A3→Prop) : Prop ≝ - | ex5_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → ex5_4 ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (5, 4)" 'Ex P0 P1 P2 P3 P4 = (ex5_4 ? ? ? ? P0 P1 P2 P3 P4). - -(* multiple existental quantifier (6, 4) *) - -inductive ex6_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→Prop) : Prop ≝ - | ex6_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → P5 x0 x1 x2 x3 → ex6_4 ? ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (6, 4)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_4 ? ? ? ? P0 P1 P2 P3 P4 P5). - -(* multiple existental quantifier (6, 6) *) - -inductive ex6_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ - | ex6_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → ex6_6 ? ? ? ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (6, 6)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5). - -(* multiple existental quantifier (7, 6) *) - -inductive ex7_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ - | ex7_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → P6 x0 x1 x2 x3 x4 x5 → ex7_6 ? ? ? ? ? ? ? ? ? ? ? ? ? -. - -interpretation "multiple existental quantifier (7, 6)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6). - -(* multiple disjunction connective (3) *) - -inductive or3 (P0,P1,P2:Prop) : Prop ≝ - | or3_intro0: P0 → or3 ? ? ? - | or3_intro1: P1 → or3 ? ? ? - | or3_intro2: P2 → or3 ? ? ? -. - -interpretation "multiple disjunction connective (3)" 'Or P0 P1 P2 = (or3 P0 P1 P2). - -(* multiple disjunction connective (4) *) - -inductive or4 (P0,P1,P2,P3:Prop) : Prop ≝ - | or4_intro0: P0 → or4 ? ? ? ? - | or4_intro1: P1 → or4 ? ? ? ? - | or4_intro2: P2 → or4 ? ? ? ? - | or4_intro3: P3 → or4 ? ? ? ? -. - -interpretation "multiple disjunction connective (4)" 'Or P0 P1 P2 P3 = (or4 P0 P1 P2 P3). - -(* multiple conjunction connective (3) *) - -inductive and3 (P0,P1,P2:Prop) : Prop ≝ - | and3_intro: P0 → P1 → P2 → and3 ? ? ? -. - -interpretation "multiple conjunction connective (3)" 'And P0 P1 P2 = (and3 P0 P1 P2). - diff --git a/matita/matita/contribs/lambda_delta/Ground_2/xoa_notation.ma b/matita/matita/contribs/lambda_delta/Ground_2/xoa_notation.ma deleted file mode 100644 index edd014ef0..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/xoa_notation.ma +++ /dev/null @@ -1,194 +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 *) -(* *) -(**************************************************************************) - -(* This file was generated by xoa.native: do not edit *********************) - -(* multiple existental quantifier (1, 2) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) }. - -(* multiple existental quantifier (2, 1) *) - -notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) }. - -notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) }. - -(* multiple existental quantifier (2, 2) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) }. - -(* multiple existental quantifier (2, 3) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) }. - -(* multiple existental quantifier (3, 1) *) - -notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) (λ${ident x0}.$P2) }. - -notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) (λ${ident x0}:$T0.$P2) }. - -(* multiple existental quantifier (3, 2) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) }. - -(* multiple existental quantifier (3, 3) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) }. - -(* multiple existental quantifier (4, 2) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) (λ${ident x0}.λ${ident x1}.$P3) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P3) }. - -(* multiple existental quantifier (4, 3) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P3) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P3) }. - -(* multiple existental quantifier (4, 4) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) }. - -(* multiple existental quantifier (5, 2) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) (λ${ident x0}.λ${ident x1}.$P3) (λ${ident x0}.λ${ident x1}.$P4) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P4) }. - -(* multiple existental quantifier (5, 3) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P4) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P4) }. - -(* multiple existental quantifier (5, 4) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P4) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) }. - -(* multiple existental quantifier (6, 4) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P5) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P5) }. - -(* multiple existental quantifier (6, 6) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P5) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) }. - -(* multiple existental quantifier (7, 6) *) - -notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P6) }. - -notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)" - non associative with precedence 20 - for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P6) }. - -(* multiple disjunction connective (3) *) - -notation "hvbox(∨∨ term 29 P0 break | term 29 P1 break | term 29 P2)" - non associative with precedence 30 - for @{ 'Or $P0 $P1 $P2 }. - -(* multiple disjunction connective (4) *) - -notation "hvbox(∨∨ term 29 P0 break | term 29 P1 break | term 29 P2 break | term 29 P3)" - non associative with precedence 30 - for @{ 'Or $P0 $P1 $P2 $P3 }. - -(* multiple conjunction connective (3) *) - -notation "hvbox(∧∧ term 34 P0 break & term 34 P1 break & term 34 P2)" - non associative with precedence 35 - for @{ 'And $P0 $P1 $P2 }. - diff --git a/matita/matita/contribs/lambda_delta/Ground_2/xoa_props.ma b/matita/matita/contribs/lambda_delta/Ground_2/xoa_props.ma deleted file mode 100644 index 121eb7e74..000000000 --- a/matita/matita/contribs/lambda_delta/Ground_2/xoa_props.ma +++ /dev/null @@ -1,20 +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/xoa_notation.ma". -include "Ground_2/xoa.ma". - -lemma ex2_1_comm: ∀A0. ∀P0,P1:A0→Prop. (∃∃x0. P0 x0 & P1 x0) → ∃∃x0. P1 x0 & P0 x0. -#A0 #P0 #P1 * /2 width=3/ -qed. diff --git a/matita/matita/contribs/lambda_delta/Makefile b/matita/matita/contribs/lambda_delta/Makefile index 5d86a4dd2..49e7e1b23 100644 --- a/matita/matita/contribs/lambda_delta/Makefile +++ b/matita/matita/contribs/lambda_delta/Makefile @@ -59,43 +59,50 @@ stats: $(PACKAGES:%=%.stats) @printf ' %-10s' '' @printf '\x1B[0m\n' -# sum ######################################################################## - -tbls: $(PACKAGES:%=etc/ld_%_sum.tbl) - -etc/ld_%_sum.tbl: MAS = $(shell find $* -name "*.ma") - -%.tbl: V1 = $(shell ls $(MAS) | wc -l) -%.tbl: C1 = $(shell grep "inductive \|record " $(MAS) | wc -l) -%.tbl: C2 = $(shell grep "definition \|let rec " $(MAS) | wc -l) -%.tbl: C3 = $(shell grep "inductive \|record \|definition \|let rec " $(MAS) | wc -l) -%.tbl: P1 = $(shell grep "theorem " $(MAS) | wc -l) -%.tbl: P2 = $(shell grep "lemma " $(MAS) | wc -l) -%.tbl: P3 = $(shell grep "lemma \|theorem " $(MAS) | wc -l) - -etc/ld_%_sum.tbl: $(MAS) - @mkdir -p etc - @printf 'Summary for: $*\n' - @printf 'name "$(basename $(@F))"\n\n' > $@ - @printf 'table {\n' >> $@ - @printf ' class "grey" [ "category"\n' >> $@ - @printf ' [ "objects" * ]\n' >> $@ - @printf ' ]\n' >> $@ - @printf ' class "cyan" [ "volume"\n' >> $@ - @printf ' [ "files" "$(V1)" * ]\n' >> $@ - @printf ' [ 4 ]\n' >> $@ - @printf ' ]\n' >> $@ - @printf ' class "green" [ "propositions"\n' >> $@ - @printf ' [ "theorems" "$(P1)" * ]\n' >> $@ - @printf ' [ "lemmas" "$(P2)" * ]\n' >> $@ - @printf ' [ "total" "$(P3)" * ]\n' >> $@ - @printf ' ]\n' >> $@ - @printf ' class "yellow" [ "concepts"\n' >> $@ - @printf ' [ "declared" "$(C1)" * ]\n' >> $@ - @printf ' [ "defined" "$(C2)" * ]\n' >> $@ - @printf ' [ "total" "$(C3)" * ]\n' >> $@ - @printf ' ]\n' >> $@ - @printf '}\n\n' >> $@ - @printf 'class "component" { 0 }\n\n' >> $@ - @printf 'class "plane" { 1 } { 3 } { 5 }\n\n' >> $@ - @printf 'class "number" { 2 } { 4 } { 6 }\n\n' >> $@ +# summary #################################################################### + +define SUMMARY_TEMPLATE + TBL_$(1) := $(1)/ld_$(1)_sum.tbl + MAS_$(1) := $$(shell find $(1) -name "*.ma") + TBLS += $$(TBL_$(1)) + + $$(TBL_$(1)): V1 := $$(shell ls $$(MAS_$(1)) | wc -l) + $$(TBL_$(1)): V2 := $$(shell cat $$(MAS_$(1)) | wc -c) + $$(TBL_$(1)): C1 := $$(shell grep "inductive \|record " $$(MAS_$(1)) | wc -l) + $$(TBL_$(1)): C2 := $$(shell grep "definition \|let rec " $$(MAS_$(1)) | wc -l) + $$(TBL_$(1)): C3 := $$(shell grep "inductive \|record \|definition \|let rec " $$(MAS_$(1)) | wc -l) + $$(TBL_$(1)): P1 := $$(shell grep "theorem " $$(MAS_$(1)) | wc -l) + $$(TBL_$(1)): P2 := $$(shell grep "lemma " $$(MAS_$(1)) | wc -l) + $$(TBL_$(1)): P3 := $$(shell grep "lemma \|theorem " $$(MAS_$(1)) | wc -l) + + $$(TBL_$(1)): $$(MAS_$(1)) + @printf ' SUMMARY $(1)\n' + @printf 'name "$$(basename $$(@F))"\n\n' > $$@ + @printf 'table {\n' >> $$@ + @printf ' class "grey" [ "category"\n' >> $$@ + @printf ' [ "objects" * ]\n' >> $$@ + @printf ' ]\n' >> $$@ + @printf ' class "cyan" [ "sizes"\n' >> $$@ + @printf ' [ "files" "$$(V1)" ]\n' >> $$@ + @printf ' [ "bytes" "$$(V2)" ]\n' >> $$@ + @printf ' [ * ]\n' >> $$@ + @printf ' ]\n' >> $$@ + @printf ' class "green" [ "propositions"\n' >> $$@ + @printf ' [ "theorems" "$$(P1)" ]\n' >> $$@ + @printf ' [ "lemmas" "$$(P2)" ]\n' >> $$@ + @printf ' [ "total" "$$(P3)" ]\n' >> $$@ + @printf ' ]\n' >> $$@ + @printf ' class "yellow" [ "concepts"\n' >> $$@ + @printf ' [ "declared" "$$(C1)" ]\n' >> $$@ + @printf ' [ "defined" "$$(C2)" ]\n' >> $$@ + @printf ' [ "total" "$$(C3)" ]\n' >> $$@ + @printf ' ]\n' >> $$@ + @printf '}\n\n' >> $$@ + @printf 'class "component" { 0 }\n\n' >> $$@ + @printf 'class "plane" { 1 } { 3 } { 5 }\n\n' >> $$@ + @printf 'class "number" { 2 } { 4 } { 6 }\n\n' >> $$@ +endef + +$(foreach PKG, $(PACKAGES), $(eval $(call SUMMARY_TEMPLATE,$(PKG)))) + +tbls: $(TBLS) diff --git a/matita/matita/contribs/lambda_delta/apps_2/functional/lift.ma b/matita/matita/contribs/lambda_delta/apps_2/functional/lift.ma new file mode 100644 index 000000000..aefa1576d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/apps_2/functional/lift.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/tri.ma". +include "Basic_2/substitution/lift.ma". +include "Apps_2/functional/notation.ma". + +(* RELOCATION ***************************************************************) + +let rec flift d e U on U ≝ match U with +[ TAtom I ⇒ match I with + [ Sort _ ⇒ U + | LRef i ⇒ #(tri … i d i (i + e) (i + e)) + | GRef _ ⇒ U + ] +| TPair I V T ⇒ match I with + [ Bind2 I ⇒ ⓑ{I} (flift d e V). (flift (d+1) e T) + | Flat2 I ⇒ ⓕ{I} (flift d e V). (flift d e T) + ] +]. + +interpretation "functional relocation" 'Lift d e T = (flift d e T). + +(* Main properties **********************************************************) + +theorem flift_lift: ∀T,d,e. ⇧[d, e] T ≡ ↑[d, e] T. +#T elim T -T +[ * #i #d #e // + elim (lt_or_eq_or_gt i d) #Hid normalize + [ >(tri_lt ?????? Hid) /2 width=1/ + | /2 width=1/ + | >(tri_gt ?????? Hid) /3 width=2/ + ] +| * /2/ +] +qed. + +(* Main inversion properties ************************************************) + +theorem flift_inv_lift: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → ↑[d, e] T1 = T2. +#d #e #T1 #T2 #H elim H -d -e -T1 -T2 normalize // +[ #i #d #e #Hid >(tri_lt ?????? Hid) // +| #i #d #e #Hid + elim (le_to_or_lt_eq … Hid) -Hid #Hid + [ >(tri_gt ?????? Hid) // + | destruct // + ] +] +qed-. + +(* Derived properties *******************************************************) + +lemma flift_join: ∀e1,e2,T. ⇧[e1, e2] ↑[0, e1] T ≡ ↑[0, e1 + e2] T. +#e1 #e2 #T +lapply (flift_lift T 0 (e1+e2)) #H +elim (lift_split … H e1 e1 ? ? ?) -H // #U #H +>(flift_inv_lift … H) -H // +qed. diff --git a/matita/matita/contribs/lambda_delta/apps_2/functional/notation.ma b/matita/matita/contribs/lambda_delta/apps_2/functional/notation.ma new file mode 100644 index 000000000..48df845c1 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/apps_2/functional/notation.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "functional" COMPONENT ********************************) + +notation "hvbox( ↑ [ d , break e ] break T )" + non associative with precedence 55 + for @{ 'Lift $d $e $T }. + +notation "hvbox( [ d ← break V ] break T )" + non associative with precedence 55 + for @{ 'Subst $V $d $T }. + +notation "hvbox( T1 ⇨ break T2 )" + non associative with precedence 45 + for @{ 'SRed $T1 $T2 }. diff --git a/matita/matita/contribs/lambda_delta/apps_2/functional/rtm.ma b/matita/matita/contribs/lambda_delta/apps_2/functional/rtm.ma new file mode 100644 index 000000000..845e8a04d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/apps_2/functional/rtm.ma @@ -0,0 +1,85 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_vector.ma". +include "Basic_2/grammar/genv.ma". + +(* REDUCTION AND TYPE MACHINE ***********************************************) + +(* machine local environment *) +inductive xenv: Type[0] ≝ +| XAtom: xenv (* empty *) +| XQuad: xenv → bind2 → nat → xenv → term → xenv (* entry *) +. + +interpretation "atom (ext. local environment)" + 'Star = XAtom. + +interpretation "environment construction (quad)" + 'DxItem4 L I u K V = (XQuad L I u K V). + +(* machine stack *) +definition stack: Type[0] ≝ list2 xenv term. + +(* machine status *) +record rtm: Type[0] ≝ +{ rg: genv; (* global environment *) + ru: nat; (* current de Bruijn's level *) + re: xenv; (* extended local environment *) + rs: stack; (* application stack *) + rt: term (* code *) +}. + +(* initial state *) +definition rtm_i: genv → term → rtm ≝ + λG,T. mk_rtm G 0 (⋆) (⟠) T. + +(* update code *) +definition rtm_t: rtm → term → rtm ≝ + λM,T. match M with + [ mk_rtm G u E _ _ ⇒ mk_rtm G u E (⟠) T + ]. + +(* update closure *) +definition rtm_u: rtm → xenv → term → rtm ≝ + λM,E,T. match M with + [ mk_rtm G u _ _ _ ⇒ mk_rtm G u E (⟠) T + ]. + +(* get global environment *) +definition rtm_g: rtm → genv ≝ + λM. match M with + [ mk_rtm G _ _ _ _ ⇒ G + ]. + +(* get local reference level *) +definition rtm_l: rtm → nat ≝ + λM. match M with + [ mk_rtm _ u E _ _ ⇒ match E with + [ XAtom ⇒ u + | XQuad _ _ u _ _ ⇒ u + ] + ]. + +(* get stack *) +definition rtm_s: rtm → stack ≝ + λM. match M with + [ mk_rtm _ _ _ S _ ⇒ S + ]. + +(* get code *) +definition rtm_c: rtm → term ≝ + λM. match M with + [ mk_rtm _ _ _ _ T ⇒ T + ]. diff --git a/matita/matita/contribs/lambda_delta/apps_2/functional/rtm_step.ma b/matita/matita/contribs/lambda_delta/apps_2/functional/rtm_step.ma new file mode 100644 index 000000000..5b2a4eb2c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/apps_2/functional/rtm_step.ma @@ -0,0 +1,57 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "Apps_2/functional/rtm.ma". + +(* REDUCTION AND TYPE MACHINE ***********************************************) + +(* transitions *) +inductive rtm_step: relation rtm ≝ +| rtm_ldrop : ∀G,u,E,I,t,F,V,S,i. + rtm_step (mk_rtm G u (E. ④{I} {t, F, V}) S (#(i + 1))) + (mk_rtm G u E S (#i)) +| rtm_ldelta: ∀G,u,E,t,F,V,S. + rtm_step (mk_rtm G u (E. ④{Abbr} {t, F, V}) S (#0)) + (mk_rtm G u F S V) +| rtm_ltype : ∀G,u,E,t,F,V,S. + rtm_step (mk_rtm G u (E. ④{Abst} {t, F, V}) S (#0)) + (mk_rtm G u F S V) +| rtm_gdrop : ∀G,I,V,u,E,S,p. p < |G| → + rtm_step (mk_rtm (G. ⓑ{I} V) u E S (§p)) + (mk_rtm G u E S (§p)) +| rtm_gdelta: ∀G,V,u,E,S,p. p = |G| → + rtm_step (mk_rtm (G. ⓓV) u E S (§p)) + (mk_rtm G u E S V) +| rtm_gtype : ∀G,V,u,E,S,p. p = |G| → + rtm_step (mk_rtm (G. ⓛV) u E S (§p)) + (mk_rtm G u E S V) +| rtm_tau : ∀G,u,E,S,W,T. + rtm_step (mk_rtm G u E S (ⓣW. T)) + (mk_rtm G u E S T) +| rtm_appl : ∀G,u,E,S,V,T. + rtm_step (mk_rtm G u E S (ⓐV. T)) + (mk_rtm G u E ({E, V} :: S) T) +| rtm_beta : ∀G,u,E,F,V,S,W,T. + rtm_step (mk_rtm G u E ({F, V} :: S) (ⓛW. T)) + (mk_rtm G u (E. ④{Abbr} {u, F, V}) S T) +| rtm_push : ∀G,u,E,W,T. + rtm_step (mk_rtm G u E ⟠ (ⓛW. T)) + (mk_rtm G (u + 1) (E. ④{Abst} {u, E, W}) ⟠ T) +| rtm_theta : ∀G,u,E,S,V,T. + rtm_step (mk_rtm G u E S (ⓓV. T)) + (mk_rtm G u (E. ④{Abbr} {u, E, V}) S T) +. + +interpretation "sequential reduction (RTM)" + 'SRed O1 O2 = (rtm_step O1 O2). diff --git a/matita/matita/contribs/lambda_delta/apps_2/functional/subst.ma b/matita/matita/contribs/lambda_delta/apps_2/functional/subst.ma new file mode 100644 index 000000000..6aa75625d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/apps_2/functional/subst.ma @@ -0,0 +1,73 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/delift_lift.ma". +include "Apps_2/functional/lift.ma". + +(* CORE SUBSTITUTION ********************************************************) + +let rec fsubst W d U on U ≝ match U with +[ TAtom I ⇒ match I with + [ Sort _ ⇒ U + | LRef i ⇒ tri … i d (#i) (↑[0, i] W) (#(i-1)) + | GRef _ ⇒ U + ] +| TPair I V T ⇒ match I with + [ Bind2 I ⇒ ⓑ{I} (fsubst W d V). (fsubst W (d+1) T) + | Flat2 I ⇒ ⓕ{I} (fsubst W d V). (fsubst W d T) + ] +]. + +interpretation "functional core substitution" 'Subst V d T = (fsubst V d T). + +(* Main properties **********************************************************) + +theorem fsubst_delift: ∀K,V,T,L,d. + ⇩[0, d] L ≡ K. ⓓV → L ⊢ T [d, 1] ≡ [d ← V] T. +#K #V #T elim T -T +[ * #i #L #d #HLK normalize in ⊢ (? ? ? ? ? %); /2 width=3/ + elim (lt_or_eq_or_gt i d) #Hid + [ -HLK >(tri_lt ?????? Hid) /3 width=3/ + | destruct >tri_eq /4 width=4 by tpss_strap, tps_subst, le_n, ex2_1_intro/ (**) (* too slow without trace *) + | -HLK >(tri_gt ?????? Hid) /3 width=3/ + ] +| * /3 width=1/ /4 width=1/ +] +qed. + +(* Main inversion properties ************************************************) + +theorem fsubst_inv_delift: ∀K,V,T1,L,T2,d. ⇩[0, d] L ≡ K. ⓓV → + L ⊢ T1 [d, 1] ≡ T2 → [d ← V] T1 = T2. +#K #V #T1 elim T1 -T1 +[ * #i #L #T2 #d #HLK #H + [ -HLK >(delift_fwd_sort1 … H) -H // + | elim (lt_or_eq_or_gt i d) #Hid normalize + [ -HLK >(delift_fwd_lref1_lt … H) -H // /2 width=1/ + | destruct + elim (delift_fwd_lref1_be … H ? ?) -H // #K0 #V0 #V2 #HLK0 + lapply (ldrop_mono … HLK0 … HLK) -HLK0 -HLK #H >minus_plus (delift_inv_refl_O2 … HV2) -V >(flift_inv_lift … HVT2) -V2 // + | -HLK >(delift_fwd_lref1_ge … H) -H // /2 width=1/ + ] + | -HLK >(delift_fwd_gref1 … H) -H // + ] +| * #I #V1 #T1 #IHV1 #IHT1 #L #X #d #HLK #H + [ elim (delift_fwd_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + <(IHV1 … HV12) -IHV1 -HV12 // <(IHT1 … HT12) -IHT1 -HT12 // /2 width=1/ + | elim (delift_fwd_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + <(IHV1 … HV12) -IHV1 -HV12 // <(IHT1 … HT12) -IHT1 -HT12 // + ] +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/Basic_1.txt b/matita/matita/contribs/lambda_delta/basic_2/Basic_1.txt new file mode 100644 index 000000000..bc4582ede --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/Basic_1.txt @@ -0,0 +1,388 @@ +# waiting #################################################################### + +aplus/props aplus_reg_r +aplus/props aplus_assoc +aplus/props aplus_asucc +aplus/props aplus_sort_O_S_simpl +aplus/props aplus_sort_S_S_simpl +aplus/props aplus_asort_O_simpl +aplus/props aplus_asort_le_simpl +aplus/props aplus_asort_simpl +aplus/props aplus_ahead_simpl +aplus/props aplus_asucc_false +aplus/props aplus_inj +aprem/fwd aprem_gen_sort +aprem/fwd aprem_gen_head_O +aprem/fwd aprem_gen_head_S +aprem/props aprem_repl +aprem/props aprem_asucc +arity/aprem arity_aprem +arity/cimp arity_cimp_conf +arity/fwd arity_gen_sort +arity/fwd arity_gen_lref +arity/fwd arity_gen_bind +arity/fwd arity_gen_abst +arity/fwd arity_gen_appl +arity/fwd arity_gen_cast +arity/fwd arity_gen_appls +arity/fwd arity_gen_lift +arity/lift1 arity_lift1 +arity/pr3 arity_sred_wcpr0_pr0 +arity/pr3 arity_sred_wcpr0_pr1 +arity/pr3 arity_sred_pr2 +arity/pr3 arity_sred_pr3 +arity/props node_inh +arity/props arity_lift +arity/props arity_mono +arity/props arity_repellent +arity/props arity_appls_cast +arity/props arity_appls_abbr +arity/props arity_appls_bind +arity/subst0 arity_gen_cvoid_subst0 +arity/subst0 arity_gen_cvoid +arity/subst0 arity_fsubst0 +arity/subst0 arity_subst0 +asucc/fwd asucc_gen_sort +asucc/fwd asucc_gen_head +cnt/props cnt_lift +C/props clt_wf__q_ind +C/props clt_wf_ind +C/props chead_ctail +C/props clt_thead (ctail) +C/props c_tail_ind +csuba/arity csuba_arity +csuba/arity csuba_arity_rev +csuba/arity arity_appls_appl +csuba/clear csuba_clear_conf +csuba/clear csuba_clear_trans +csuba/drop csuba_drop_abbr +csuba/drop csuba_drop_abst +csuba/drop csuba_drop_abst_rev +csuba/drop csuba_drop_abbr_rev +csuba/fwd csuba_gen_abbr +csuba/fwd csuba_gen_void +csuba/fwd csuba_gen_abst +csuba/fwd csuba_gen_flat +csuba/fwd csuba_gen_bind +csuba/fwd csuba_gen_abst_rev +csuba/fwd csuba_gen_void_rev +csuba/fwd csuba_gen_abbr_rev +csuba/fwd csuba_gen_flat_rev +csuba/fwd csuba_gen_bind_rev +csuba/getl csuba_getl_abbr +csuba/getl csuba_getl_abst +csuba/getl csuba_getl_abst_rev +csuba/getl csuba_getl_abbr_rev +csuba/props csuba_refl +csubc/arity csubc_arity_conf +csubc/arity csubc_arity_trans +csubc/csuba csubc_csuba +csubc/drop1 drop1_csubc_trans +csubc/drop drop_csubc_trans +csubt/clear csubt_clear_conf +csubt/csuba csubt_csuba +csubt/drop csubt_drop_flat +csubt/drop csubt_drop_abbr +csubt/drop csubt_drop_abst +csubt/fwd csubt_gen_abbr +csubt/fwd csubt_gen_abst +csubt/fwd csubt_gen_flat +csubt/fwd csubt_gen_bind +csubt/getl csubt_getl_abbr +csubt/getl csubt_getl_abst +csubt/pc3 csubt_pr2 +csubt/pc3 csubt_pc3 +csubt/props csubt_refl +csubt/ty3 csubt_ty3 +csubt/ty3 csubt_ty3_ld +csubv/clear csubv_clear_conf +csubv/clear csubv_clear_conf_void +csubv/drop csubv_drop_conf +csubv/getl csubv_getl_conf +csubv/getl csubv_getl_conf_void +csubv/props csubv_bind_same +csubv/props csubv_refl +drop1/props drop1_cons_tail +drop/props drop_ctail +ex0/props aplus_gz_le +ex0/props aplus_gz_ge +ex0/props next_plus_gz +ex0/props leqz_leq +ex0/props leq_leqz +ex1/props ex1__leq_sort_SS +ex1/props ex1_arity +ex1/props ex1_ty3 +ex2/props ex2_nf2 +ex2/props ex2_arity +fsubst0/fwd fsubst0_gen_base +leq/asucc asucc_repl +leq/asucc asucc_inj +leq/asucc leq_asucc +leq/asucc leq_ahead_asucc_false +leq/asucc leq_asucc_false +leq/fwd leq_gen_sort1 +leq/fwd leq_gen_head1 +leq/fwd leq_gen_sort2 +leq/fwd leq_gen_head2 +leq/props ahead_inj_snd +leq/props leq_refl +leq/props leq_eq +leq/props leq_sym +leq/props leq_trans +leq/props leq_ahead_false_1 +leq/props leq_ahead_false_2 +lift1/fwd lift1_cons_tail + +# check ############################################################# + +lift1/fwd lifts1_flat +lift1/fwd lifts1_nil +lift1/fwd lifts1_cons +lift1/props lift1_free +lift/props thead_x_lift_y_y +lift/props lifts_tapp + +# waiting #################################################################### + +lift/props lifts_inj +llt/props lweight_repl +llt/props llt_repl +llt/props llt_trans +llt/props llt_head_sx +llt/props llt_head_dx +llt/props llt_wf__q_ind +llt/props llt_wf_ind +next_plus/props next_plus_assoc +next_plus/props next_plus_next +next_plus/props next_plus_lt +nf2/arity arity_nf2_inv_all +nf2/dec nf2_dec + +nf2/fwd nf2_gen_lref +nf2/fwd nf2_gen_abst +nf2/fwd nf2_gen_cast +nf2/fwd nf2_gen_beta +nf2/fwd nf2_gen_flat +nf2/fwd nf2_gen__nf2_gen_aux +nf2/fwd nf2_gen_abbr + +nf2/fwd nf2_gen_void +nf2/iso nf2_iso_appls_lref +nf2/pr3 nf2_pr3_unfold +nf2/pr3 nf2_pr3_confluence + +nf2/props nfs2_tapp +nf2/props nf2_appls_lref + +pc1/props pc1_pr0_r +pc1/props pc1_pr0_x +pc1/props pc1_refl +pc1/props pc1_pr0_u +pc1/props pc1_s +pc1/props pc1_head_1 +pc1/props pc1_head_2 +pc1/props pc1_t +pc1/props pc1_pr0_u2 +pc1/props pc1_head + +pc3/dec pc3_dec +pc3/dec pc3_abst_dec +pc3/fsubst0 pc3_pr2_fsubst0 +pc3/fsubst0 pc3_pr2_fsubst0_back +pc3/fsubst0 pc3_fsubst0 +pc3/fwd pc3_gen_sort +pc3/fwd pc3_gen_abst +pc3/fwd pc3_gen_abst_shift +pc3/fwd pc3_gen_lift +pc3/fwd pc3_gen_not_abst +pc3/fwd pc3_gen_lift_abst +pc3/fwd pc3_gen_sort_abst +pc3/left pc3_ind_left__pc3_left_pr3 +pc3/left pc3_ind_left__pc3_left_trans +pc3/left pc3_ind_left__pc3_left_sym +pc3/left pc3_ind_left__pc3_left_pc3 +pc3/left pc3_ind_left__pc3_pc3_left +pc3/left pc3_ind_left +pc3/nf2 pc3_nf2 +pc3/nf2 pc3_nf2_unfold +pc3/pc1 pc3_pc1 +pc3/props pc3_pr2_r +pc3/props pc3_pr2_x +pc3/props pc3_pr3_r +pc3/props pc3_pr3_x +pc3/props pc3_pr3_t +pc3/props pc3_s +pc3/props pc3_thin_dx +pc3/props pc3_head_1 +pc3/props pc3_head_2 +pc3/props pc3_pr2_u +pc3/props pc3_t +pc3/props pc3_pr2_u2 +pc3/props pc3_pr3_conf +pc3/props pc3_head_12 +pc3/props pc3_head_21 +pc3/props pc3_pr0_pr2_t +pc3/props pc3_pr2_pr2_t +pc3/props pc3_pr2_pr3_t +pc3/props pc3_pr3_pc3_t +pc3/props pc3_lift +pc3/props pc3_eta +pc3/subst1 pc3_gen_cabbr +pc3/wcpr0 pc3_wcpr0__pc3_wcpr0_t_aux +pc3/wcpr0 pc3_wcpr0_t +pc3/wcpr0 pc3_wcpr0 +pr0/fwd pr0_gen_void + +# check ############################################################# + +pr0/dec nf0_dec +pr0/subst1 pr0_subst1_back +pr0/subst1 pr0_subst1_fwd +pr1/pr1 pr1_strip +pr1/pr1 pr1_confluence +pr1/props pr1_pr0 +pr1/props pr1_t +pr1/props pr1_head_1 +pr1/props pr1_head_2 +pr1/props pr1_comp +pr1/props pr1_eta +pr2/clen pr2_gen_ctail +pr2/fwd pr2_gen_void +pr2/props pr2_ctail +pr2/subst1 pr2_gen_cabbr + +pr3/fwd pr3_gen_abst +pr3/fwd pr3_gen_lref +pr3/fwd pr3_gen_void +pr3/fwd pr3_gen_appl +pr3/fwd pr3_gen_bind +pr3/iso pr3_iso_appls_abbr +pr3/iso pr3_iso_appl_bind +pr3/iso pr3_iso_appls_appl_bind +pr3/iso pr3_iso_appls_bind +pr3/iso pr3_iso_beta +pr3/iso pr3_iso_appls_beta +pr3/pr1 pr3_pr1 +pr3/pr3 pr3_strip +pr3/props pr3_thin_dx +pr3/props pr3_head_1 +pr3/props pr3_head_2 +pr3/props pr3_head_21 +pr3/props pr3_head_12 +pr3/props pr3_flat +pr3/props pr3_eta +pr3/subst1 pr3_subst1 +pr3/subst1 pr3_gen_cabbr +sn3/props sn3_cpr3_trans + +sn3/props sn3_shift +sn3/props sn3_change +sn3/props sn3_gen_def +sn3/props sn3_cdelta +sn3/props sn3_appl_lref +sn3/props sn3_appl_abbr +sn3/props sn3_appl_cast +sn3/props sn3_appl_beta +sn3/props sn3_appl_appls +sn3/props sn3_appls_lref +sn3/props sn3_appls_cast +sn3/props sn3_appls_bind +sn3/props sn3_appls_beta +sn3/props sn3_appls_abbr +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 +sty1/props sty1_appl +sty1/props sty1_lift +sty1/props sty1_correct +sty1/props sty1_abbr +sty1/props sty1_cast2 +subst/fwd subst_sort +subst/fwd subst_lref_lt +subst/fwd subst_lref_eq +subst/fwd subst_lref_gt +subst/fwd subst_head +subst/props subst_lift_SO +subst/props subst_subst0 +T/dec binder_dec +T/dec abst_dec +tlist/props tslt_wf__q_ind +tlist/props tslt_wf_ind +tlist/props theads_tapp +tlist/props tcons_tapp_ex +tlist/props tlist_ind_rev +ty3/arity ty3_arity +ty3/arity_props ty3_predicative +ty3/arity_props ty3_repellent +ty3/arity_props ty3_acyclic +ty3/arity_props ty3_sn3 +ty3/dec ty3_inference +ty3/fsubst0 ty3_fsubst0 +ty3/fsubst0 ty3_csubst0 +ty3/fsubst0 ty3_subst0 +ty3/fwd ty3_gen_sort +ty3/fwd ty3_gen_lref +ty3/fwd ty3_gen_bind +ty3/fwd ty3_gen_appl +ty3/fwd ty3_gen_cast +ty3/fwd tys3_gen_nil +ty3/fwd tys3_gen_cons +ty3/fwd_nf2 ty3_gen_appl_nf2 +ty3/fwd_nf2 ty3_inv_lref_nf2_pc3 +ty3/fwd_nf2 ty3_inv_lref_nf2 +ty3/fwd_nf2 ty3_inv_appls_lref_nf2 +ty3/fwd_nf2 ty3_inv_lref_lref_nf2 +ty3/nf2 ty3_nf2_inv_abst_premise_csort +ty3/nf2 ty3_nf2_inv_all +ty3/nf2 ty3_nf2_inv_sort +ty3/nf2 ty3_nf2_gen__ty3_nf2_inv_abst_aux +ty3/nf2 ty3_nf2_inv_abst +ty3/pr3 ty3_sred_wcpr0_pr0 +ty3/pr3 ty3_sred_pr0 +ty3/pr3 ty3_sred_pr1 +ty3/pr3 ty3_sred_pr2 +ty3/pr3 ty3_sred_pr3 +ty3/pr3_props ty3_cred_pr2 +ty3/pr3_props ty3_cred_pr3 +ty3/pr3_props ty3_gen_lift +ty3/pr3_props ty3_tred +ty3/pr3_props ty3_sconv_pc3 +ty3/pr3_props ty3_sred_back +ty3/pr3_props ty3_sconv +ty3/props ty3_lift +ty3/props ty3_correct +ty3/props ty3_unique +ty3/props ty3_gen_abst_abst +ty3/props ty3_typecheck +ty3/props ty3_getl_subst0 +ty3/sty0 ty3_sty0 +ty3/subst1 ty3_gen_cabbr +ty3/subst1 ty3_gen_cvoid +wf3/clear wf3_clear_conf +wf3/clear clear_wf3_trans +wf3/fwd wf3_gen_sort1 +wf3/fwd wf3_gen_bind1 +wf3/fwd wf3_gen_flat1 +wf3/fwd wf3_gen_head2 +wf3/getl wf3_getl_conf +wf3/getl getl_wf3_trans +wf3/props wf3_mono +wf3/props wf3_total +wf3/props ty3_shift1 +wf3/props wf3_idem +wf3/props wf3_ty3 +wf3/ty3 wf3_pr2_conf +wf3/ty3 wf3_pr3_conf +wf3/ty3 wf3_pc3_conf +wf3/ty3 wf3_ty3_conf + +# check ###################################################################### diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/acp.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/acp.ma new file mode 100644 index 000000000..8be75a31f --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/acp.ma @@ -0,0 +1,55 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ldrops.ma". + +(* ABSTRACT COMPUTATION PROPERTIES ******************************************) + +definition CP1 ≝ λRR:lenv→relation term. λRS:relation term. + ∀L,k. NF … (RR L) RS (⋆k). + +definition CP2 ≝ λRR:lenv→relation term. λRS:relation term. + ∀L,K,W,i. ⇩[0,i] L ≡ K. ⓛW → NF … (RR L) RS (#i). + +definition CP3 ≝ λRR:lenv→relation term. λRP:lenv→predicate term. + ∀L,V,k. RP L (ⓐ⋆k.V) → RP L V. + +definition CP4 ≝ λRR:lenv→relation term. λRS:relation term. + ∀L0,L,T,T0,d,e. NF … (RR L) RS T → + ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → NF … (RR L0) RS T0. + +definition CP4s ≝ λRR:lenv→relation term. λRS:relation term. + ∀L0,L,des. ⇩*[des] L0 ≡ L → + ∀T,T0. ⇧*[des] T ≡ T0 → + NF … (RR L) RS T → NF … (RR L0) RS T0. + +(* requirements for abstract computation properties *) +record acp (RR:lenv->relation term) (RS:relation term) (RP:lenv→predicate term) : Prop ≝ +{ cp1: CP1 RR RS; + cp2: CP2 RR RS; + cp3: CP3 RR RP; + cp4: CP4 RR RS +}. + +(* Basic properties *********************************************************) + +(* Basic_1: was: nf2_lift1 *) +lemma acp_lifts: ∀RR,RS. CP4 RR RS → CP4s RR RS. +#RR #RS #HRR #L1 #L2 #des #H elim H -L1 -L2 -des +[ #L #T1 #T2 #H #HT1 + <(lifts_inv_nil … H) -H // +| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2 + elim (lifts_inv_cons … H) -H /3 width=9/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/acp_aaa.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/acp_aaa.ma new file mode 100644 index 000000000..1cab5d4b8 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/acp_aaa.ma @@ -0,0 +1,101 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/lifts_lifts.ma". +include "Basic_2/unfold/ldrops_ldrops.ma". +include "Basic_2/static/aaa_lifts.ma". +include "Basic_2/static/aaa_aaa.ma". +include "Basic_2/computation/lsubc_ldrops.ma". + +(* ABSTRACT COMPUTATION PROPERTIES ******************************************) + +(* Main propertis ***********************************************************) + +(* Basic_1: was: sc3_arity_csubc *) +theorem aacr_aaa_csubc_lifts: ∀RR,RS,RP. + acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L1,T,A. L1 ⊢ T ÷ A → ∀L0,des. ⇩*[des] L0 ≡ L1 → + ∀T0. ⇧*[des] T ≡ T0 → ∀L2. L2 [RP] ⊑ L0 → + ⦃L2, T0⦄ [RP] ϵ 〚A〛. +#RR #RS #RP #H1RP #H2RP #L1 #T #A #H elim H -L1 -T -A +[ #L #k #L0 #des #HL0 #X #H #L2 #HL20 + >(lifts_inv_sort1 … H) -H + lapply (aacr_acr … H1RP H2RP ⓪) #HAtom + @(s2 … HAtom … ◊) // /2 width=2/ +| #I #L1 #K1 #V1 #B #i #HLK1 #HKV1B #IHB #L0 #des #HL01 #X #H #L2 #HL20 + lapply (aacr_acr … H1RP H2RP B) #HB + elim (lifts_inv_lref1 … H) -H #i1 #Hi1 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK1) #HK1b + elim (ldrops_ldrop_trans … HL01 … HLK1) #X #des1 #i0 #HL0 #H #Hi0 #Hdes1 + >(at_mono … Hi1 … Hi0) -i1 + elim (ldrops_inv_skip2 … Hdes1 … H) -des1 #K0 #V0 #des0 #Hdes0 #HK01 #HV10 #H destruct + elim (lsubc_ldrop_O1_trans … HL20 … HL0) -HL0 #X #HLK2 #H + elim (lsubc_inv_pair2 … H) -H * + [ #K2 #HK20 #H destruct + generalize in match HLK2; generalize in match I; -HLK2 -I * #HLK2 + [ elim (lift_total V0 0 (i0 +1)) #V #HV0 + elim (lifts_lift_trans … Hi0 … Hdes0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2 + @(s4 … HB … ◊ … HV0 HLK2) /3 width=7/ (* uses IHB HL20 V2 HV0 *) + | @(s2 … HB … ◊) // /2 width=3/ + ] + | -HLK1 -IHB -HL01 -HL20 -HK1b -Hi0 -Hdes0 + #K2 #V2 #A2 #HKV2A #HKV0A #_ #H1 #H2 destruct + lapply (ldrop_fwd_ldrop2 … HLK2) #HLK2b + lapply (aaa_lifts … HK01 … HV10 HKV1B) -HKV1B -HK01 -HV10 #HKV0B + >(aaa_mono … HKV0A … HKV0B) in HKV2A; -HKV0A -HKV0B #HKV2B + elim (lift_total V2 0 (i0 +1)) #V #HV2 + @(s4 … HB … ◊ … HV2 HLK2) + @(s7 … HB … HKV2B) // + ] +| #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20 + elim (lifts_inv_bind1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct + lapply (aacr_acr … H1RP H2RP A) #HA + lapply (aacr_acr … H1RP H2RP B) #HB + lapply (s1 … HB) -HB #HB + @(s5 … HA … ◊ ◊) // /3 width=5/ +| #L #W #T #B #A #HLWB #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL02 + elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct + @(aacr_abst … H1RP H2RP) + [ lapply (aacr_acr … H1RP H2RP B) #HB + @(s1 … HB) /2 width=5/ + | -IHB + #L3 #V3 #T3 #des3 #HL32 #HT03 #HB + elim (lifts_total des3 W0) #W2 #HW02 + elim (ldrops_lsubc_trans … H1RP H2RP … HL32 … HL02) -L2 #L2 #HL32 #HL20 + lapply (aaa_lifts … L2 W2 … (des @ des3) … HLWB) -HLWB /2 width=3/ #HLW2B + @(IHA (L2. ⓛW2) … (des + 1 @ des3 + 1)) -IHA + /2 width=3/ /3 width=5/ + ] +| #L #V #T #B #A #_ #_ #IHB #IHA #L0 #des #HL0 #X #H #L2 #HL20 + elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct + /3 width=10/ +| #L #V #T #A #_ #_ #IH1A #IH2A #L0 #des #HL0 #X #H #L2 #HL20 + elim (lifts_inv_flat1 … H) -H #V0 #T0 #HV0 #HT0 #H destruct + lapply (aacr_acr … H1RP H2RP A) #HA + lapply (s1 … HA) #H + @(s6 … HA … ◊) /2 width=5/ /3 width=5/ +] +qed. + +(* Basic_1: was: sc3_arity *) +lemma aacr_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L,T,A. L ⊢ T ÷ A → ⦃L, T⦄ [RP] ϵ 〚A〛. +/2 width=8/ qed. + +lemma acp_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L,T,A. L ⊢ T ÷ A → RP L T. +#RR #RS #RP #H1RP #H2RP #L #T #A #HT +lapply (aacr_acr … H1RP H2RP A) #HA +@(s1 … HA) /2 width=4/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/acp_cr.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/acp_cr.ma new file mode 100644 index 000000000..6eb71054b --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/acp_cr.ma @@ -0,0 +1,174 @@ +(**************************************************************************) +(* ___ *) +(* ||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/aarity.ma". +include "Basic_2/unfold/gr2_gr2.ma". +include "Basic_2/unfold/lifts_lift_vector.ma". +include "Basic_2/unfold/ldrops_ldrop.ma". +include "Basic_2/computation/acp.ma". + +(* ABSTRACT COMPUTATION PROPERTIES ******************************************) + +(* Note: this is Girard's CR1 *) +definition S1 ≝ λRP,C:lenv→predicate term. + ∀L,T. C L T → RP L T. + +(* Note: this is Tait's iii, or Girard's CR4 *) +definition S2 ≝ λRR:lenv→relation term. λRS:relation term. λRP,C:lenv→predicate term. + ∀L,Vs. all … (RP L) Vs → + ∀T. 𝐒[T] → NF … (RR L) RS T → C L (ⒶVs.T). + +(* Note: this is Tait's ii *) +definition S3 ≝ λRP,C:lenv→predicate term. + ∀L,Vs,V,T,W. C L (ⒶVs. ⓓV. T) → RP L W → C L (ⒶVs. ⓐV. ⓛW. T). + +definition S4 ≝ λRP,C:lenv→predicate term. ∀L,K,Vs,V1,V2,i. + C L (ⒶVs. V2) → ⇧[0, i + 1] V1 ≡ V2 → + ⇩[0, i] L ≡ K. ⓓV1 → C L (Ⓐ Vs. #i). + +definition S5 ≝ λRP,C:lenv→predicate term. + ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → + ∀V,T. C (L. ⓓV) (ⒶV2s. T) → RP L V → C L (ⒶV1s. ⓓV. T). + +definition S6 ≝ λRP,C:lenv→predicate term. + ∀L,Vs,T,W. C L (ⒶVs. T) → RP L W → C L (ⒶVs. ⓣW. T). + +definition S7 ≝ λC:lenv→predicate term. ∀L2,L1,T1,d,e. + C L1 T1 → ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → C L2 T2. + +definition S7s ≝ λC:lenv→predicate term. + ∀L1,L2,des. ⇩*[des] L2 ≡ L1 → + ∀T1,T2. ⇧*[des] T1 ≡ T2 → C L1 T1 → C L2 T2. + +(* properties of the abstract candidate of reducibility *) +record acr (RR:lenv->relation term) (RS:relation term) (RP,C:lenv→predicate term) : Prop ≝ +{ s1: S1 RP C; + s2: S2 RR RS RP C; + s3: S3 RP C; + s4: S4 RP C; + s5: S5 RP C; + s6: S6 RP C; + s7: S7 C +}. + +(* the abstract candidate of reducibility associated to an atomic arity *) +let rec aacr (RP:lenv→predicate term) (A:aarity) (L:lenv) on A: predicate term ≝ +λT. match A with +[ AAtom ⇒ RP L T +| APair B A ⇒ ∀L0,V0,T0,des. aacr RP B L0 V0 → ⇩*[des] L0 ≡ L → ⇧*[des] T ≡ T0 → + aacr RP A L0 (ⓐV0. T0) +]. + +interpretation + "candidate of reducibility of an atomic arity (abstract)" + 'InEInt RP L T A = (aacr RP A L T). + +(* Basic properties *********************************************************) + +(* Basic_1: was: sc3_lift1 *) +lemma acr_lifts: ∀C. S7 C → S7s C. +#C #HC #L1 #L2 #des #H elim H -L1 -L2 -des +[ #L #T1 #T2 #H #HT1 + <(lifts_inv_nil … H) -H // +| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2 + elim (lifts_inv_cons … H) -H /3 width=9/ +] +qed. + +lemma rp_lifts: ∀RR,RS,RP. acr RR RS RP (λL,T. RP L T) → + ∀des,L0,L,V,V0. ⇩*[des] L0 ≡ L → ⇧*[des] V ≡ V0 → + RP L V → RP L0 V0. +#RR #RS #RP #HRP #des #L0 #L #V #V0 #HL0 #HV0 #HV +@acr_lifts /width=6/ +@(s7 … HRP) +qed. + +(* Basic_1: was only: sns3_lifts1 *) +lemma rp_liftsv_all: ∀RR,RS,RP. acr RR RS RP (λL,T. RP L T) → + ∀des,L0,L,Vs,V0s. ⇧*[des] Vs ≡ V0s → ⇩*[des] L0 ≡ L → + all … (RP L) Vs → all … (RP L0) V0s. +#RR #RS #RP #HRP #des #L0 #L #Vs #V0s #H elim H -Vs -V0s normalize // +#T1s #T2s #T1 #T2 #HT12 #_ #IHT2s #HL0 * #HT1 #HT1s +@conj /2 width=1/ /2 width=6 by rp_lifts/ +qed. + +(* Basic_1: was: + sc3_sn3 sc3_abst sc3_appl sc3_abbr sc3_bind sc3_cast sc3_lift +*) +lemma aacr_acr: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀A. acr RR RS RP (aacr RP A). +#RR #RS #RP #H1RP #H2RP #A elim A -A normalize // +#B #A #IHB #IHA @mk_acr normalize +[ #L #T #H + lapply (H ? (⋆0) ? ⟠ ? ? ?) -H + [1,3: // |2,4: skip + | @(s2 … IHB … ◊) // /2 width=2/ + | #H @(cp3 … H1RP … 0) @(s1 … IHA) // + ] +| #L #Vs #HVs #T #H1T #H2T #L0 #V0 #X #des #HB #HL0 #H + elim (lifts_inv_applv1 … H) -H #V0s #T0 #HV0s #HT0 #H destruct + lapply (s1 … IHB … HB) #HV0 + @(s2 … IHA … (V0 :: V0s)) /2 width=4 by lifts_simple_dx/ /3 width=6/ +| #L #Vs #U #T #W #HA #HW #L0 #V0 #X #des #HB #HL0 #H + elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct + elim (lifts_inv_flat1 … HY) -HY #U0 #X #HU0 #HX #H destruct + elim (lifts_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT0 #H destruct + @(s3 … IHA … (V0 :: V0s)) /2 width=6 by rp_lifts/ /4 width=5/ +| #L #K #Vs #V1 #V2 #i #HA #HV12 #HLK #L0 #V0 #X #des #HB #HL0 #H + elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct + elim (lifts_inv_lref1 … HY) -HY #i0 #Hi0 #H destruct + elim (ldrops_ldrop_trans … HL0 … HLK) #X #des0 #i1 #HL02 #H #Hi1 #Hdes0 + >(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02 + elim (ldrops_inv_skip2 … Hdes0 … H) -H -des0 #L2 #W1 #des0 #Hdes0 #HLK #HVW1 #H destruct + elim (lift_total W1 0 (i0 + 1)) #W2 #HW12 + elim (lifts_lift_trans … Hdes0 … HVW1 … HW12) // -Hdes0 -Hi0 #V3 #HV13 #HVW2 + >(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2 + @(s4 … IHA … (V0 :: V0s) … HW12 HL02) /3 width=4/ +| #L #V1s #V2s #HV12s #V #T #HA #HV #L0 #V10 #X #des #HB #HL0 #H + elim (lifts_inv_applv1 … H) -H #V10s #Y #HV10s #HY #H destruct + elim (lifts_inv_bind1 … HY) -HY #V0 #T0 #HV0 #HT0 #H destruct + elim (lift_total V10 0 1) #V20 #HV120 + elim (liftv_total 0 1 V10s) #V20s #HV120s + @(s5 … IHA … (V10 :: V10s) (V20 :: V20s)) /2 width=1/ /2 width=6 by rp_lifts/ + @(HA … (des + 1)) /2 width=1/ + [ @(s7 … IHB … HB … HV120) /2 width=1/ + | @lifts_applv // + elim (liftsv_liftv_trans_le … HV10s … HV120s) -V10s #V10s #HV10s #HV120s + >(liftv_mono … HV12s … HV10s) -V1s // + ] +| #L #Vs #T #W #HA #HW #L0 #V0 #X #des #HB #HL0 #H + elim (lifts_inv_applv1 … H) -H #V0s #Y #HV0s #HY #H destruct + elim (lifts_inv_flat1 … HY) -HY #W0 #T0 #HW0 #HT0 #H destruct + @(s6 … IHA … (V0 :: V0s)) /2 width=6 by rp_lifts/ /3 width=4/ +| /3 width=7/ +] +qed. + +lemma aacr_abst: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L,W,T,A,B. RP L W → ( + ∀L0,V0,T0,des. ⇩*[des] L0 ≡ L → ⇧*[des + 1] T ≡ T0 → + ⦃L0, V0⦄ [RP] ϵ 〚B〛 → ⦃L0. ⓓV0, T0⦄ [RP] ϵ 〚A〛 + ) → + ⦃L, ⓛW. T⦄ [RP] ϵ 〚②B. A〛. +#RR #RS #RP #H1RP #H2RP #L #W #T #A #B #HW #HA #L0 #V0 #X #des #HB #HL0 #H +lapply (aacr_acr … H1RP H2RP A) #HCA +lapply (aacr_acr … H1RP H2RP B) #HCB +elim (lifts_inv_bind1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct +lapply (s1 … HCB) -HCB #HCB +@(s3 … HCA … ◊) /2 width=6 by rp_lifts/ +@(s5 … HCA … ◊ ◊) // /2 width=1/ /2 width=3/ +qed. + +(* Basic_1: removed theorems 2: sc3_arity_gen sc3_repl *) +(* Basic_1: removed local theorems 1: sc3_sn3_abst *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs.ma new file mode 100644 index 000000000..12cd0791f --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs.ma @@ -0,0 +1,85 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Basic_1: includes: pr3_pr2 *) +definition cprs: lenv → relation term ≝ + λL. TC … (cpr L). + +interpretation "context-sensitive parallel computation (term)" + 'PRedStar L T1 T2 = (cprs L T1 T2). + +(* Basic eliminators ********************************************************) + +lemma cprs_ind: ∀L,T1. ∀R:predicate term. R T1 → + (∀T,T2. L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → R T → R T2) → + ∀T2. L ⊢ T1 ➡* T2 → R T2. +#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // +qed-. + +axiom cprs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 → + (∀T1,T. L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → R T → R T1) → + ∀T1. L ⊢ T1 ➡* T2 → R T1. + +(* Basic properties *********************************************************) + +(* Basic_1: was: pr3_refl *) +lemma cprs_refl: ∀L,T. L ⊢ T ➡* T. +/2 width=1/ qed. + +lemma cprs_strap1: ∀L,T1,T,T2. + L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → L ⊢ T1 ➡* T2. +/2 width=3/ qed. + +(* Basic_1: was: pr3_step *) +lemma cprs_strap2: ∀L,T1,T,T2. + L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → L ⊢ T1 ➡* T2. +/2 width=3/ qed. + +(* Note: it does not hold replacing |L1| with |L2| *) +lemma cprs_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡* T2 → + ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡* T2. +/3 width=3/ +qed. + +lemma cprs_flat_dx: ∀I,L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L ⊢ T1 ➡* T2 → + L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. +#I #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind … HT12) -T2 /3 width=1/ +#T #T2 #_ #HT2 #IHT2 +@(cprs_strap1 … IHT2) -IHT2 /2 width=1/ +qed. + +(* Basic inversion lemmas ***************************************************) + +(* Basic_1: was: pr3_gen_sort *) +lemma cprs_inv_sort1: ∀L,U2,k. L ⊢ ⋆k ➡* U2 → U2 = ⋆k. +#L #U2 #k #H @(cprs_ind … H) -U2 // +#U2 #U #_ #HU2 #IHU2 destruct +>(cpr_inv_sort1 … HU2) -HU2 // +qed-. + +(* Basic_1: was: pr3_gen_cast *) +lemma cprs_inv_cast1: ∀L,W1,T1,U2. L ⊢ ⓣW1.T1 ➡* U2 → L ⊢ T1 ➡* U2 ∨ + ∃∃W2,T2. L ⊢ W1 ➡* W2 & L ⊢ T1 ➡* T2 & U2 = ⓣW2.T2. +#L #W1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ +#U2 #U #_ #HU2 * /3 width=3/ * +#W #T #HW1 #HT1 #H destruct +elim (cpr_inv_cast1 … HU2) -HU2 /3 width=3/ * +#W2 #T2 #HW2 #HT2 #H destruct /4 width=5/ +qed-. + +(* Basic_1: removed theorems 2: clear_pr3_trans pr3_cflat *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_cprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_cprs.ma new file mode 100644 index 000000000..16c1a6256 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_cprs.ma @@ -0,0 +1,91 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr_lift.ma". +include "Basic_2/reducibility/cpr_cpr.ma". +include "Basic_2/reducibility/lcpr_cpr.ma". +include "Basic_2/computation/cprs_lcpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Advanced properties ******************************************************) + +lemma cprs_abbr_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 → + L ⊢ ⓓV1. T1 ➡* ⓓV2. T2. +#L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1 +[ /3 width=5/ +| #T1 #T #HT1 #_ #IHT1 + @(cprs_strap2 … IHT1) -IHT1 /2 width=1/ +] +qed. + +lemma cpr_abbr: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡ T2 → + L ⊢ ⓓV1. T1 ➡* ⓓV2. T2. +/3 width=1/ qed. + +lemma cpr_abbr2: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡ T2 → + L ⊢ ⓓV1. T1 ➡* ⓓV2. T2. +#L #V1 #V2 #HV12 #T1 #T2 #HT12 +lapply (lcpr_cpr_trans (L. ⓓV1) … HT12) /2 width=1/ +qed. + +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was pr3_gen_appl *) +lemma cprs_inv_appl1: ∀L,V1,T1,U2. L ⊢ ⓐV1. T1 ➡* U2 → + ∨∨ ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 & + U2 = ⓐV2. T2 + | ∃∃V2,W,T. L ⊢ V1 ➡* V2 & + L ⊢ T1 ➡* ⓛW. T & L ⊢ ⓓV2. T ➡* U2 + | ∃∃V0,V2,V,T. L ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 & + L ⊢ T1 ➡* ⓓV. T & L ⊢ ⓓV. ⓐV2. T ➡* U2. +#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ +#U #U2 #_ #HU2 * * +[ #V0 #T0 #HV10 #HT10 #H destruct + elim (cpr_inv_appl1 … HU2) -HU2 * + [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/ + | #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct /4 width=7/ + | #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct + @or3_intro2 @(ex4_4_intro … HV2 HT10) /2 width=3/ /3 width=1/ (**) (* explicit constructor. /5 width=8/ is too slow because TC_transitive gets in the way *) + ] +| /4 width=8/ +| /4 width=10/ +] +qed-. + +(* Main propertis ***********************************************************) + +(* Basic_1: was: pr3_confluence *) +theorem cprs_conf: ∀L,T1,T. L ⊢ T ➡* T1 → ∀T2. L ⊢ T ➡* T2 → + ∃∃T0. L ⊢ T1 ➡* T0 & L ⊢ T2 ➡* T0. +/3 width=3/ qed. + +(* Basic_1: was: pr3_t *) +theorem cprs_trans: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T ➡* T2 → L ⊢ T1 ➡* T2. +/2 width=3/ qed. + +lemma cprs_flat: ∀I,L,T1,T2. L ⊢ T1 ➡* T2 → ∀V1,V2. L ⊢ V1 ➡* V2 → + L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. +#I #L #T1 #T2 #HT12 #V1 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/ +#V #V2 #_ #HV2 #IHV1 +@(cprs_trans … IHV1) -IHV1 /2 width=1/ +qed. + +(* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *) +lemma lcpr_cprs_trans: ∀L1,L2. L1 ⊢ ➡ L2 → + ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2. +#L1 #L2 #HL12 #T1 #T2 #H @(cprs_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT2 +@(cprs_trans … IHT2) /2 width=3/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcpr.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcpr.ma new file mode 100644 index 000000000..7816989a7 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcpr.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/reducibility/ltpr_tps.ma". +include "Basic_2/reducibility/cpr_ltpss.ma". +include "Basic_2/reducibility/lcpr.ma". +include "Basic_2/computation/cprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Properties concerning context-sensitive parallel reduction on lenv's *****) + +lemma ltpr_tpss_trans: ∀L1,L2. L1 ➡ L2 → ∀T1,T2,d,e. L2 ⊢ T1 [d, e] ▶* T2 → + ∃∃T. L1 ⊢ T1 [d, e] ▶* T & L1 ⊢ T ➡* T2. +#L1 #L2 #HL12 #T1 #T2 #d #e #H @(tpss_ind … H) -T2 +[ /2 width=3/ +| #T #T2 #_ #HT2 * #T0 #HT10 #HT0 + elim (ltpr_tps_trans … HT2 … HL12) -L2 #T3 #HT3 #HT32 + @(ex2_1_intro … HT10) -T1 (**) (* explicit constructors *) + @(cprs_strap1 … T3 …) /2 width=1/ -HT32 + @(cprs_strap1 … HT0) -HT0 /3 width=3/ +] +qed. + +(* Basic_1: was just: pr3_pr0_pr2_t *) +lemma ltpr_cpr_trans: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡* T2. +#L1 #L2 #HL12 #T1 #T2 * #T #HT1 +<(ltpr_fwd_length … HL12) #HT2 +elim (ltpr_tpss_trans … HL12 … HT2) -L2 /3 width=3/ +qed. + +(* Basic_1: was just: pr3_pr2_pr2_t *) +lemma lcpr_cpr_trans: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡* T2. +#L1 #L2 * /3 width=7/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcprs.ma new file mode 100644 index 000000000..6d4f5ec8c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcprs.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/computation/cprs_cprs.ma". +include "Basic_2/computation/lcprs_lcprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Properties exploiting context-senstive computation on local environments *) + +(* Basic_1: was just: pr3_pr3_pr3_t *) +lemma lcprs_cprs_trans: ∀L1,L2. L1 ⊢ ➡* L2 → + ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2. +#L1 #L2 #HL12 @(lcprs_ind … HL12) -L2 // /3 width=3/ +qed. + +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was pr3_gen_abbr *) +lemma cprs_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡* U2 → + (∃∃V2,T2. L ⊢ V1 ➡* V2 & L. ⓓV1 ⊢ T1 ➡* T2 & + U2 = ⓓV2. T2 + ) ∨ + ∃∃U. ⇧[0, 1] U2 ≡ U & L. ⓓV1 ⊢ T1 ➡* U. +#L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ +#U0 #U2 #_ #HU02 * * +[ #V0 #T0 #HV10 #HT10 #H destruct + elim (cpr_inv_abbr1 … HU02) -HU02 * + [ #V #V2 #T2 #HV0 #HV2 #HT02 #H destruct + lapply (cpr_intro … HV0 … HV2) -HV2 #HV02 + lapply (ltpr_cpr_trans (L.ⓓV0) … HT02) /2 width=1/ -V #HT02 + lapply (lcprs_cprs_trans (L. ⓓV1) … HT02) -HT02 /2 width=1/ /4 width=5/ + | -V0 #T2 #HT20 #HTU2 + elim (lift_total U2 0 1) #U0 #HU20 + lapply (cpr_lift (L.ⓓV1) … HT20 … HU20 HTU2) -T2 /2 width=1/ /4 width=5/ + ] +| #U1 #HU01 #HTU1 + elim (lift_total U2 0 1) #U #HU2 + lapply (cpr_lift (L.ⓓV1) … HU01 … HU2 HU02) -U0 /2 width=1/ /4 width=5/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.ma new file mode 100644 index 000000000..ccff2f1fe --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.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/reducibility/cpr_lift.ma". +include "Basic_2/computation/cprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Relocation properties ****************************************************) + +(* Basic_1: was: pr3_lift *) +lemma cprs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 → + ∀T2. K ⊢ T1 ➡* T2 → ∀U2. ⇧[d, e] T2 ≡ U2 → + L ⊢ U1 ➡* U2. +#L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #HT12 @(cprs_ind … HT12) -T2 +[ -HLK #T2 #HT12 + <(lift_mono … HTU1 … HT12) -T1 // +| -HTU1 #T #T2 #_ #HT2 #IHT2 #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (cpr_lift … HLK … HTU … HTU2 … HT2) -T2 -HLK /3 width=3/ +] +qed. + +(* Basic_1: was: pr3_gen_lift *) +lemma cprs_inv_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → + ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀U2. L ⊢ U1 ➡* U2 → + ∃∃T2. ⇧[d, e] T2 ≡ U2 & K ⊢ T1 ➡* T2. +#L #K #d #e #HLK #T1 #U1 #HTU1 #U2 #HU12 @(cprs_ind … HU12) -U2 /2 width=3/ +-HTU1 #U #U2 #_ #HU2 * #T #HTU #HT1 +elim (cpr_inv_lift … HLK … HTU … HU2) -U -HLK /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc.ma new file mode 100644 index 000000000..e623c7b90 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc.ma @@ -0,0 +1,56 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tstc.ma". +(* +include "Basic_2/reducibility/cpr_lift.ma". +include "Basic_2/reducibility/lcpr_cpr.ma". +*) +include "Basic_2/computation/cprs_cprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Forward lemmas involving same top term constructor ***********************) +(* +lemma cpr_fwd_beta: ∀L,V,W,T,U. L ⊢ ⓐV. ⓛW. T ➡ U → + ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⓓV. T ➡* U. +#L #V #W #T #U #H +elim (cpr_inv_appl1 … H) -H * +[ #V0 #X #_ #_ #H destruct /2 width=1/ +| #V0 #W0 #T1 #T2 #HV0 #HT12 #H1 #H2 destruct + lapply (lcpr_cpr_trans (L. ⓓV) … HT12) -HT12 /2 width=1/ /3 width=1/ +| #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #H destruct +] +qed-. + +lemma cpr_fwd_theta: ∀L,V1,V,T,U. L ⊢ ⓐV1. ⓓV. T ➡ U → + ∀V2. ⇧[0, 1] V1 ≡ V2 → ⓐV1. ⓓV. T ≃ U ∨ + L ⊢ ⓓV. ⓐV2. T ➡* U. +#L #V1 #V #T #U #H #V2 #HV12 +elim (cpr_inv_appl1 … H) -H * +[ -HV12 #V0 #X #_ #_ #H destruct /2 width=1/ +| -HV12 #V0 #W #T1 #T2 #_ #_ #H destruct +| #V0 #V3 #W1 #W2 #T1 #T2 #HV10 #HW12 #HT12 #HV03 #H1 #H2 destruct + lapply (cpr_lift (L.ⓓW1) … HV12 … HV03 … HV10) -V0 -HV12 /2 width=1/ #HV23 + lapply (lcpr_cpr_trans (L. ⓓW1) … HT12) -HT12 /2 width=1/ #HT12 + /4 width=1/ +] +qed-. +*) +lemma cprs_fwd_tau: ∀L,W,T,U. L ⊢ ⓣW. T ➡* U → + ⓣW. T ≃ U ∨ L ⊢ T ➡* U. +#L #W #T #U #H +elim (cprs_inv_cast1 … H) -H /2 width=1/ * +#W0 #T0 #_ #_ #H destruct /2 width=1/ +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma new file mode 100644 index 000000000..3e6f4a20d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.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 "Basic_2/grammar/tstc_vector.ma". +include "Basic_2/substitution/lift_vector.ma". +include "Basic_2/computation/cprs_tstc.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Vector form of forward lemmas involving same top term constructor ********) +(* +lemma cpr_fwd_beta_vector: ∀L,V,W,T,U,Vs. L ⊢ ⒶVs. ⓐV. ⓛW. T ➡ U → + ⒶVs. ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⒶVs. ⓓV. T ➡* U. +#L #V #W #T #U * /2 width=1 by cpr_fwd_beta/ +#V0 #Vs #H +elim (cpr_inv_appl1_simple … H ?) -H +[ #V1 #T1 #_ #_ #H destruct /2 width=1/ +| elim Vs -Vs // +] +qed-. + +lemma cpr_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → + ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡ U → + ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U. +#L #V1s #V2s * -V1s -V2s /3 width=1/ +#V1s #V2s #V1a #V2a #HV12a * -V1s -V2s /2 width=1 by cpr_fwd_theta/ -HV12a +#V1s #V2s #V1b #V2b #_ #_ #V #U #T #H +elim (cpr_inv_appl1_simple … H ?) -H // +#V0 #T0 #_ #_ #H destruct /2 width=1/ +qed-. +*) + +(* Basic_1: was: pr3_iso_appls_cast *) +lemma cprs_fwd_tau_vector: ∀L,Vs,W,T,U. L ⊢ ⒶVs. ⓣW. T ➡* U → + ⒶVs. ⓣW. T ≃ U ∨ L ⊢ ⒶVs. T ➡* U. +#L #Vs elim Vs -Vs /2 width=1 by cprs_fwd_tau/ +#V #Vs #IHVs #W #T #U #H +elim (cprs_inv_appl1 … H) -H * +[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/ +| #V0 #W0 #T0 #HV0 #HT0 #HU + elim (IHVs … HT0) -IHVs -HT0 #HT0 + [ elim (tstc_inv_bind_appls_simple … HT0 ?) // + | @or_intror + @(cprs_trans … HU) -HU + @(cprs_strap1 … (ⓐV0.ⓛW0.T0)) /2 width=1/ -HV0 -HT0 /3 width=1/ + ] +| #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU + elim (IHVs … HT0) -IHVs -HT0 #HT0 + [ elim (tstc_inv_bind_appls_simple … HT0 ?) // + | @or_intror + @(cprs_trans … HU) -HU + @(cprs_strap1 … (ⓐV0.ⓓV2.T0)) /2 width=1/ -HV0 -HT0 /3 width=3/ +] +qed-. + +axiom cprs_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → + ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡* U → + ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma new file mode 100644 index 000000000..043fa9a27 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma @@ -0,0 +1,85 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr.ma". +include "Basic_2/reducibility/cnf.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) + +definition csn: lenv → predicate term ≝ λL. SN … (cpr L) (eq …). + +interpretation + "context-sensitive strong normalization (term)" + 'SN L T = (csn L T). + +(* Basic eliminators ********************************************************) + +lemma csn_ind: ∀L. ∀R:predicate term. + (∀T1. L ⊢ ⬇* T1 → + (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → R T2) → + R T1 + ) → + ∀T. L ⊢ ⬇* T → R T. +#L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1 +@H0 -H0 /3 width=1/ -IHT1 /4 width=1/ +qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: sn3_pr2_intro *) +lemma csn_intro: ∀L,T1. + (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇* T2) → L ⊢ ⬇* T1. +#L #T1 #H +@(SN_intro … H) +qed. + +(* Basic_1: was: sn3_nf2 *) +lemma csn_cnf: ∀L,T. L ⊢ 𝐍[T] → L ⊢ ⬇* T. +/2 width=1/ qed. + +lemma csn_cpr_trans: ∀L,T1. L ⊢ ⬇* T1 → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ ⬇* T2. +#L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12 +@csn_intro #T #HLT2 #HT2 +elim (term_eq_dec T1 T2) #HT12 +[ -IHT1 -HLT12 destruct /3 width=1/ +| -HT1 -HT2 /3 width=4/ +qed. + +(* Basic_1: was: sn3_cast *) +lemma csn_cast: ∀L,W. L ⊢ ⬇* W → ∀T. L ⊢ ⬇* T → L ⊢ ⬇* ⓣW. T. +#L #W #HW elim HW -W #W #_ #IHW #T #HT @(csn_ind … HT) -T #T #HT #IHT +@csn_intro #X #H1 #H2 +elim (cpr_inv_cast1 … H1) -H1 +[ * #W0 #T0 #HLW0 #HLT0 #H destruct + elim (eq_false_inv_tpair_sn … H2) -H2 + [ /3 width=3/ + | -HLW0 * #H destruct /3 width=1/ + ] +| /3 width=3/ +] +qed. + +(* Basic forward lemmas *****************************************************) + +fact csn_fwd_flat_dx_aux: ∀L,U. L ⊢ ⬇* U → ∀I,V,T. U = ⓕ{I} V. T → L ⊢ ⬇* T. +#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct +@csn_intro #T2 #HLT2 #HT2 +@(IH (ⓕ{I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/ +qed. + +(* Basic_1: was: sn3_gen_flat *) +lemma csn_fwd_flat_dx: ∀I,L,V,T. L ⊢ ⬇* ⓕ{I} V. T → L ⊢ ⬇* T. +/2 width=5/ qed-. + +(* Basic_1: removed theorems 3: sn3_gen_cflat sn3_cflat sn3_bind *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_aaa.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_aaa.ma new file mode 100644 index 000000000..421a3a7f7 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_aaa.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/computation/acp_aaa.ma". +include "Basic_2/computation/csn_lcpr_vector.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) + +(* Properties concerning atomic arity assignment ****************************) + +lemma csn_aaa: ∀L,T,A. L ⊢ T ÷ A → L ⊢ ⬇* T. +#L #T #A #H +@(acp_aaa … csn_acp csn_acr … H) +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_cpr.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_cpr.ma new file mode 100644 index 000000000..11647e9dc --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_cpr.ma @@ -0,0 +1,41 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr_cpr.ma". +include "Basic_2/computation/csn.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) + +(* Advanced forvard lemmas **************************************************) + +fact csn_fwd_pair_sn_aux: ∀L,U. L ⊢ ⬇* U → ∀I,V,T. U = ②{I} V. T → L ⊢ ⬇* V. +#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct +@csn_intro #V2 #HLV2 #HV2 +@(IH (②{I} V2. T)) -IH // /2 width=1/ -HLV2 #H destruct /2 width=1/ +qed. + +(* Basic_1: was: sn3_gen_head *) +lemma csn_fwd_pair_sn: ∀I,L,V,T. L ⊢ ⬇* ②{I} V. T → L ⊢ ⬇* V. +/2 width=5/ qed. + +fact csn_fwd_bind_dx_aux: ∀L,U. L ⊢ ⬇* U → + ∀I,V,T. U = ⓑ{I} V. T → L. ⓑ{I} V ⊢ ⬇* T. +#L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct +@csn_intro #T2 #HLT2 #HT2 +@(IH (ⓑ{I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/ +qed. + +(* Basic_1: was: sn3_gen_bind *) +lemma csn_fwd_bind_dx: ∀I,L,V,T. L ⊢ ⬇* ⓑ{I} V. T → L. ⓑ{I} V ⊢ ⬇* T. +/2 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_cprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_cprs.ma new file mode 100644 index 000000000..bf6683d3c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_cprs.ma @@ -0,0 +1,111 @@ +(**************************************************************************) +(* ___ *) +(* ||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/computation/cprs.ma". +include "Basic_2/computation/csn.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) + +(* Properties concerning context-sensitive computation on terms *************) + +definition csns: lenv → predicate term ≝ λL. SN … (cprs L) (eq …). + +interpretation + "context-sensitive strong normalization (term)" + 'SNStar L T = (csns L T). + +(* Basic eliminators ********************************************************) + +lemma csns_ind: ∀L. ∀R:predicate term. + (∀T1. L ⊢ ⬇** T1 → + (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → R T2) → R T1 + ) → + ∀T. L ⊢ ⬇** T → R T. +#L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1 +@H0 -H0 /3 width=1/ -IHT1 /4 width=1/ +qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: sn3_intro *) +lemma csns_intro: ∀L,T1. + (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → L ⊢ ⬇** T2) → L ⊢ ⬇** T1. +#L #T1 #H +@(SN_intro … H) +qed. + +fact csns_intro_aux: ∀L,T1. + (∀T,T2. L ⊢ T ➡* T2 → T1 = T → (T1 = T2 → False) → L ⊢ ⬇** T2) → L ⊢ ⬇** T1. +/4 width=3/ qed-. + +(* Basic_1: was: sn3_pr3_trans (old version) *) +lemma csns_cprs_trans: ∀L,T1. L ⊢ ⬇** T1 → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ ⬇** T2. +#L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12 +@csns_intro #T #HLT2 #HT2 +elim (term_eq_dec T1 T2) #HT12 +[ -IHT1 -HLT12 destruct /3 width=1/ +| -HT1 -HT2 /3 width=4/ +qed. + +(* Basic_1: was: sn3_pr2_intro (old version) *) +lemma csns_intro_cpr: ∀L,T1. + (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇** T2) → + L ⊢ ⬇** T1. +#L #T1 #H +@csns_intro_aux #T #T2 #H @(cprs_ind_dx … H) -T +[ -H #H destruct #H + elim (H ?) // +| #T0 #T #HLT1 #HLT2 #IHT #HT10 #HT12 destruct + elim (term_eq_dec T0 T) #HT0 + [ -HLT1 -HLT2 -H /3 width=1/ + | -IHT -HT12 /4 width=3/ + ] +] +qed. + +(* Main properties **********************************************************) + +theorem csn_csns: ∀L,T. L ⊢ ⬇* T → L ⊢ ⬇** T. +#L #T #H @(csn_ind … H) -T /4 width=1/ +qed. + +theorem csns_csn: ∀L,T. L ⊢ ⬇** T → L ⊢ ⬇* T. +#L #T #H @(csns_ind … H) -T /4 width=1/ +qed. + +(* Basic_1: was: sn3_pr3_trans *) +lemma csn_cprs_trans: ∀L,T1. L ⊢ ⬇* T1 → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ ⬇* T2. +/4 width=3/ qed. + +(* Basic_1: was: nf2_sn3 *) +lemma csn_cwn: ∀L,T1. L ⊢ ⬇* T1 → + ∃∃T2. L ⊢ T1 ➡* T2 & L ⊢ 𝐍[T2]. +#L #T1 #H elim H -T1 #T1 #_ #IHT1 +elim (cnf_dec L T1) +[ -IHT1 /2 width=3/ +| * #T2 #HLT12 #HT12 + elim (IHT1 T2 ? ?) -IHT1 // /2 width=1/ -HT12 /3 width=3/ +] +qed. + +(* Main eliminators *********************************************************) + +lemma csn_ind_cprs: ∀L. ∀R:predicate term. + (∀T1. L ⊢ ⬇* T1 → + (∀T2. L ⊢ T1 ➡* T2 → (T1 = T2 → False) → R T2) → R T1 + ) → + ∀T. L ⊢ ⬇* T → R T. +#L #R #H0 #T1 #H @(csns_ind … (csn_csns … H)) -T1 #T1 #HT1 #IHT1 +@H0 -H0 /2 width=1/ -HT1 /3 width=1/ +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma new file mode 100644 index 000000000..303be5997 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma @@ -0,0 +1,146 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tstc_tstc.ma". +include "Basic_2/computation/cprs_cprs.ma". +include "Basic_2/computation/csn_lift.ma". +include "Basic_2/computation/csn_cpr.ma". +include "Basic_2/computation/csn_cprs.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) + +(* Advanced properties ******************************************************) + +lemma csn_lcpr_conf: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T. L1 ⊢ ⬇* T → L2 ⊢ ⬇* T. +#L1 #L2 #HL12 #T #H @(csn_ind_cprs … H) -T #T #_ #IHT +@csn_intro #T0 #HLT0 #HT0 +@IHT /2 width=2/ -IHT -HT0 /2 width=3/ +qed. + +lemma csn_abbr: ∀L,V. L ⊢ ⬇* V → ∀T. L. ⓓV ⊢ ⬇* T → L ⊢ ⬇* ⓓV. T. +#L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_cprs … HT) -T #T #HT #IHT +@csn_intro #X #H1 #H2 +elim (cpr_inv_abbr1 … H1) -H1 * +[ #V0 #V1 #T1 #HLV0 #HLV01 #HLT1 #H destruct + lapply (cpr_intro … HLV0 HLV01) -HLV01 #HLV1 + lapply (ltpr_cpr_trans (L. ⓓV) … HLT1) /2 width=1/ -V0 #HLT1 + elim (eq_false_inv_tpair_sn … H2) -H2 + [ #HV1 @IHV // /2 width=1/ -HV1 + @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -HLV1 /2 width=3/ + | -IHV -HLV1 * #H destruct /3 width=1/ + ] +| -IHV -IHT -H2 #T0 #HT0 #HLT0 + lapply (csn_inv_lift … HT … HT0) -HT /2 width=3/ +] +qed. + +fact csn_appl_beta_aux: ∀L,W. L ⊢ ⬇* W → ∀U. L ⊢ ⬇* U → + ∀V,T. U = ⓓV. T → L ⊢ ⬇* ⓐV. ⓛW. T. +#L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V #T #H destruct +lapply (csn_fwd_pair_sn … HVT) #HV +lapply (csn_fwd_bind_dx … HVT) #HT -HVT +@csn_intro #X #H #H2 +elim (cpr_inv_appl1 … H) -H * +[ #V0 #Y #HLV0 #H #H0 destruct + elim (cpr_inv_abst1 … H Abbr V) -H #W0 #T0 #HLW0 #HLT0 #H destruct + elim (eq_false_inv_beta … H2) -H2 + [ -IHVT #HW0 @IHW -IHW [1,5: // |3: skip ] -HLW0 /2 width=1/ -HW0 + @csn_abbr /2 width=3/ -HV + @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/ + | -IHW -HLW0 -HV -HT * #H #HVT0 destruct + @(IHVT … HVT0) -IHVT -HVT0 // /2 width=1/ + ] +| -IHW -IHVT -H2 #V0 #W0 #T0 #T1 #HLV0 #HLT01 #H1 #H2 destruct + lapply (lcpr_cpr_trans (L. ⓓV) … HLT01) -HLT01 /2 width=1/ #HLT01 + @csn_abbr /2 width=3/ -HV + @(csn_lcpr_conf (L. ⓓV)) /2 width=1/ -V0 /2 width=3/ +| -IHW -IHVT -HV -HT -H2 #V0 #V1 #W0 #W1 #T0 #T1 #_ #_ #_ #_ #H destruct +] +qed. + +(* Basic_1: was: sn3_beta *) +lemma csn_appl_beta: ∀L,W. L ⊢ ⬇* W → ∀V,T. L ⊢ ⬇* (ⓓV. T) → (**) + L ⊢ ⬇* ⓐV. ⓛW. T. +/2 width=3/ qed. + +fact csn_appl_theta_aux: ∀L,U. L ⊢ ⬇* U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → + ∀V,T. U = ⓓV. ⓐV2. T → L ⊢ ⬇* ⓐV1. ⓓV. T. +#L #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct +lapply (csn_fwd_pair_sn … HVT) #HV +lapply (csn_fwd_bind_dx … HVT) -HVT #HVT +@csn_intro #X #HL #H +elim (cpr_inv_appl1 … HL) -HL * +[ -HV #V0 #Y #HLV10 #HL #H0 destruct + elim (cpr_inv_abbr1 … HL) -HL * + [ #V3 #V4 #T3 #HV3 #HLV34 #HLT3 #H0 destruct + lapply (cpr_intro … HV3 HLV34) -HLV34 #HLV34 + elim (lift_total V0 0 1) #V5 #HV05 + elim (term_eq_dec (ⓓV.ⓐV2.T) (ⓓV4.ⓐV5.T3)) + [ -IHVT #H0 destruct + elim (eq_false_inv_tpair_sn … H) -H + [ -HLV10 -HLV34 -HV3 -HLT3 -HVT + >(lift_inj … HV12 … HV05) -V5 + #H elim (H ?) // + | * #_ #H elim (H ?) // + ] + | -H -HVT #H + lapply (cpr_lift (L. ⓓV) … HV12 … HV05 HLV10) -HLV10 -HV12 /2 width=1/ #HV25 + lapply (ltpr_cpr_trans (L. ⓓV) … HLT3) /2 width=1/ -HLT3 #HLT3 + @(IHVT … H … HV05) -IHVT // -H -HV05 /3 width=1/ + ] + | -H -IHVT #T0 #HT0 #HLT0 + @(csn_cpr_trans … (ⓐV1.T0)) /2 width=1/ -V0 -Y + @(csn_inv_lift … 0 1 HVT) /2 width=1/ + ] +| -HV -HV12 -HVT -IHVT -H #V0 #W0 #T0 #T1 #_ #_ #H destruct +| -IHVT -H #V0 #V3 #W0 #W1 #T0 #T1 #HLV10 #HLW01 #HLT01 #HV03 #H1 #H2 destruct + lapply (cpr_lift (L. ⓓW0) … HV12 … HV03 HLV10) -HLV10 -HV12 -HV03 /2 width=1/ #HLV23 + lapply (lcpr_cpr_trans (L. ⓓW0) … HLT01) -HLT01 /2 width=1/ #HLT01 + @csn_abbr /2 width=3/ -HV + @(csn_lcpr_conf (L. ⓓW0)) /2 width=1/ -W1 + @(csn_cprs_trans … HVT) -HVT /2 width=1/ +] +qed. + +(* Basic_1: was: sn3_appl_bind *) +lemma csn_appl_theta: ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → + ∀L,V,T. L ⊢ ⬇* ⓓV. ⓐV2. T → L ⊢ ⬇* ⓐV1. ⓓV. T. +/2 width=5/ qed. + +(* Basic_1: was only: sn3_appl_appl *) +lemma csn_appl_simple_tstc: ∀L,V. L ⊢ ⬇* V → ∀T1. + L ⊢ ⬇* T1 → + (∀T2. L ⊢ T1 ➡* T2 → (T1 ≃ T2 → False) → L ⊢ ⬇* ⓐV. T2) → + 𝐒[T1] → L ⊢ ⬇* ⓐV. T1. +#L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #H @(csn_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1 +@csn_intro #X #HL #H +elim (cpr_inv_appl1_simple … HL ?) -HL // +#V0 #T0 #HLV0 #HLT10 #H0 destruct +elim (eq_false_inv_tpair_sn … H) -H +[ -IHT1 #HV0 + @(csn_cpr_trans … (ⓐV0.T1)) /2 width=1/ -HLT10 + @IHV -IHV // -H1T1 -H3T1 /2 width=1/ -HV0 + #T2 #HLT12 #HT12 + @(csn_cpr_trans … (ⓐV.T2)) /2 width=1/ -HLV0 + @H2T1 -H2T1 // -HLT12 /2 width=1/ +| -IHV -H1T1 -HLV0 * #H #H1T10 destruct + elim (tstc_dec T1 T0) #H2T10 + [ @IHT1 -IHT1 // /2 width=1/ -H1T10 /2 width=3/ -H3T1 + #T2 #HLT02 #HT02 + @H2T1 -H2T1 /2 width=3/ -HLT10 -HLT02 /3 width=3/ + | -IHT1 -H3T1 -H1T10 + @H2T1 -H2T1 /2 width=1/ + ] +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr_vector.ma new file mode 100644 index 000000000..4ff8c039b --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr_vector.ma @@ -0,0 +1,85 @@ +(**************************************************************************) +(* ___ *) +(* ||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/computation/acp_cr.ma". +include "Basic_2/computation/cprs_tstc_vector.ma". +include "Basic_2/computation/csn_lcpr.ma". +include "Basic_2/computation/csn_vector.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERM VECTORS **********************) + +(* Advanced properties ******************************************************) +(* +(* Basic_1: was only: sn3_appl_appls *) +lemma csn_appl_appls_simple_tstc: ∀L,Vs,V,T1. L ⊢ ⬇* V → L ⊢ ⬇* T1 → + (∀T2. L ⊢ ⒶVs.T1 ➡* T2 → (ⒶVs.T1 ≃ T2 → False) → L ⊢ ⬇* ⓐV. T2) → + 𝐒[T1] → L ⊢ ⬇* ⓐV. ⒶVs. T1. +#L * +[ #V #T1 #HV + @csn_appl_simple_tstc // +| #V0 #Vs #V #T1 #HV #H1T1 #H2T1 #H3T1 + @csn_appl_simple_tstc // -HV + [ @H2T1 +] +qed. +*) +lemma csn_applv_theta: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → + ∀V,T. L ⊢ ⬇* ⓓV. ⒶV2s. T → L ⊢ ⬇* V → + L ⊢ ⬇* ⒶV1s. ⓓV. T. +#L #V1s #V2s * -V1s -V2s /2 width=1/ +#V1s #V2s #V1 #V2 #HV12 #H +generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1 +elim H -V1s -V2s /2 width=3/ +#V1s #V2s #V1 #V2 #HV12 #HV12s #IHV12s #W1 #W2 #HW12 #V #T #H #HV +lapply (csn_appl_theta … HW12 … H) -H -HW12 #H +lapply (csn_fwd_pair_sn … H) #HW1 +lapply (csn_fwd_flat_dx … H) #H1 +@csn_appl_simple_tstc // -HW1 /2 width=3/ -IHV12s -HV -H1 #X #H1 #H2 +elim (cprs_fwd_theta_vector … (V2::V2s) … H1) -H1 /2 width=1/ -HV12s -HV12 +[ -H #H elim (H2 ?) -H2 // +| -H2 #H1 @(csn_cprs_trans … H) -H /2 width=1/ +] +qed. + +lemma csn_applv_tau: ∀L,W. L ⊢ ⬇* W → + ∀Vs,T. L ⊢ ⬇* ⒶVs. T → + L ⊢ ⬇* ⒶVs. ⓣW. T. +#L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW +#V #Vs #IHV #T #H1T +lapply (csn_fwd_pair_sn … H1T) #HV +lapply (csn_fwd_flat_dx … H1T) #H2T +@csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T +[ #X #H #H0 + elim (cprs_fwd_tau_vector … H) -H #H + [ -H1T elim (H0 ?) -H0 // + | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ + ] +| -H1T elim Vs -Vs // +] +qed. +(* +theorem csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬇* T). +@mk_acr // +[ +| +| +| #L #V1 #V2 #HV12 #V #T #H #HVT + @(csn_applv_theta … HV12) -HV12 // + @(csn_abbr) // +| /2 width=1/ +| @csn_lift +] +qed. +*) +axiom csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬇* T). diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lift.ma new file mode 100644 index 000000000..01d983c0c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lift.ma @@ -0,0 +1,100 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cnf_lift.ma". +include "Basic_2/computation/acp.ma". +include "Basic_2/computation/csn.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) + +(* Relocation properties ****************************************************) + +(* Basic_1: was: sn3_lift *) +lemma csn_lift: ∀L2,L1,T1,d,e. L1 ⊢ ⬇* T1 → + ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → L2 ⊢ ⬇* T2. +#L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12 +@csn_intro #T #HLT2 #HT2 +elim (cpr_inv_lift … HL21 … HT12 … HLT2) -HLT2 #T0 #HT0 #HLT10 +@(IHT1 … HLT10) // -L1 -L2 #H destruct +>(lift_mono … HT0 … HT12) in HT2; -T0 /2 width=1/ +qed. + +(* Basic_1: was: sn3_gen_lift *) +lemma csn_inv_lift: ∀L2,L1,T1,d,e. L1 ⊢ ⬇* T1 → + ∀T2. ⇩[d, e] L1 ≡ L2 → ⇧[d, e] T2 ≡ T1 → L2 ⊢ ⬇* T2. +#L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21 +@csn_intro #T #HLT2 #HT2 +elim (lift_total T d e) #T0 #HT0 +lapply (cpr_lift … HL12 … HT21 … HT0 HLT2) -HLT2 #HLT10 +@(IHT1 … HLT10) // -L1 -L2 #H destruct +>(lift_inj … HT0 … HT21) in HT2; -T0 /2 width=1/ +qed. + +(* Advanced properties ******************************************************) + +(* Basic_1: was: sn3_abbr *) +lemma csn_lref_abbr: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓓV → K ⊢ ⬇* V → L ⊢ ⬇* #i. +#L #K #V #i #HLK #HV +@csn_intro #X #H #Hi +elim (cpr_inv_lref1 … H) -H +[ #H destruct elim (Hi ?) // +| -Hi * #K0 #V0 #V1 #HLK0 #HV01 #HV1 #_ + lapply (ldrop_mono … HLK0 … HLK) -HLK #H destruct + lapply (ldrop_fwd_ldrop2 … HLK0) -HLK0 #HLK + @(csn_lift … HLK HV1) -HLK -HV1 + @(csn_cpr_trans … HV) -HV + @(cpr_intro … HV01) -HV01 // +] +qed. + +lemma csn_abst: ∀L,W. L ⊢ ⬇* W → ∀I,V,T. L. ⓑ{I} V ⊢ ⬇* T → L ⊢ ⬇* ⓛW. T. +#L #W #HW elim HW -W #W #_ #IHW #I #V #T #HT @(csn_ind … HT) -T #T #HT #IHT +@csn_intro #X #H1 #H2 +elim (cpr_inv_abst1 … H1 I V) -H1 +#W0 #T0 #HLW0 #HLT0 #H destruct +elim (eq_false_inv_tpair_sn … H2) -H2 +[ /3 width=5/ +| -HLW0 * #H destruct /3 width=1/ +] +qed. + +lemma csn_appl_simple: ∀L,V. L ⊢ ⬇* V → ∀T1. + (∀T2. L ⊢ T1 ➡ T2 → (T1 = T2 → False) → L ⊢ ⬇* ⓐV. T2) → + 𝐒[T1] → L ⊢ ⬇* ⓐV. T1. +#L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1 +@csn_intro #X #H1 #H2 +elim (cpr_inv_appl1_simple … H1 ?) // -H1 +#V0 #T0 #HLV0 #HLT10 #H destruct +elim (eq_false_inv_tpair_dx … H2) -H2 +[ -IHV -HT1 #HT10 + @(csn_cpr_trans … (ⓐV.T0)) /2 width=1/ -HLV0 + @IHT1 -IHT1 // /2 width=1/ +| -HLT10 * #H #HV0 destruct + @IHV -IHV // -HT1 /2 width=1/ -HV0 + #T2 #HLT02 #HT02 + @(csn_cpr_trans … (ⓐV.T2)) /2 width=1/ -HLV0 + @IHT1 -IHT1 // -HLT02 /2 width=1/ +] +qed. + +(* Main properties **********************************************************) + +theorem csn_acp: acp cpr (eq …) (csn …). +@mk_acp +[ /2 width=1/ +| /2 width=3/ +| /2 width=5/ +| @cnf_lift +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_vector.ma new file mode 100644 index 000000000..bddaaa89d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_vector.ma @@ -0,0 +1,50 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_vector.ma". +include "Basic_2/computation/csn.ma". + +(* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERM VECTORS **********************) + +inductive csnv (L:lenv): predicate (list term) ≝ +| csnv_nil: csnv L ◊ +| csn_cons: ∀Ts,T. L ⊢ ⬇* T → csnv L Ts → csnv L (T :: Ts) +. + +interpretation + "context-sensitive strong normalization (term vector)" + 'SN L Ts = (csnv L Ts). + +(* Basic inversion lemmas ***************************************************) + +fact csnv_inv_cons_aux: ∀L,Ts. L ⊢ ⬇* Ts → ∀U,Us. Ts = U :: Us → + L ⊢ ⬇* U ∧ L ⊢ ⬇* Us. +#L #Ts * -Ts +[ #U #Us #H destruct +| #Ts #T #HT #HTs #U #Us #H destruct /2 width=1/ +] +qed. + +lemma csnv_inv_cons: ∀L,T,Ts. L ⊢ ⬇* T :: Ts → L ⊢ ⬇* T ∧ L ⊢ ⬇* Ts. +/2 width=3/ qed-. + +include "Basic_2/computation/csn_cpr.ma". + +lemma csn_fwd_applv: ∀L,T,Vs. L ⊢ ⬇* Ⓐ Vs. T → L ⊢ ⬇* Vs ∧ L ⊢ ⬇* T. +#L #T #Vs elim Vs -Vs /2 width=1/ +#V #Vs #IHVs #HVs +lapply (csn_fwd_pair_sn … HVs) #HV +lapply (csn_fwd_flat_dx … HVs) -HVs #HVs +elim (IHVs HVs) -IHVs -HVs /3 width=1/ +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs.ma new file mode 100644 index 000000000..627a0f58d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs.ma @@ -0,0 +1,44 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/lcpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************) + +definition lcprs: relation lenv ≝ TC … lcpr. + +interpretation + "context-sensitive parallel computation (environment)" + 'CPRedStar L1 L2 = (lcprs L1 L2). + +(* Basic eliminators ********************************************************) + +lemma lcprs_ind: ∀L1. ∀R:predicate lenv. R L1 → + (∀L,L2. L1 ⊢ ➡* L → L ⊢ ➡ L2 → R L → R L2) → + ∀L2. L1 ⊢ ➡* L2 → R L2. +#L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // +qed-. + +(* Basic properties *********************************************************) + +lemma lcprs_refl: ∀L. L ⊢ ➡* L. +/2 width=1/ qed. + +lemma lcprs_strap1: ∀L1,L,L2. + L1 ⊢ ➡* L → L ⊢ ➡ L2 → L1 ⊢ ➡* L2. +/2 width=3/ qed. + +lemma lcprs_strap2: ∀L1,L,L2. + L1 ⊢ ➡ L → L ⊢ ➡* L2 → L1 ⊢ ➡* L2. +/2 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_cprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_cprs.ma new file mode 100644 index 000000000..a45e8ee9f --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_cprs.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/lcpr_cpr.ma". +include "Basic_2/computation/cprs.ma". +include "Basic_2/computation/lcprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************) + +(* Advanced properties ******************************************************) + +lemma lcprs_pair_dx: ∀I,L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. L2 ⊢ V1 ➡* V2 → + L1. ⓑ{I} V1 ⊢ ➡* L2. ⓑ{I} V2. +#I #L1 #L2 #HL12 #V1 #V2 #H @(cprs_ind … H) -V2 +/3 width=1/ /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_lcprs.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_lcprs.ma new file mode 100644 index 000000000..935b83d24 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lcprs_lcprs.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||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/computation/lcprs_cprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS *************) + +(* Main properties **********************************************************) + +theorem lcprs_trans: ∀L1,L. L1 ⊢ ➡* L → ∀L2. L ⊢ ➡* L2 → L1 ⊢ ➡* L2. +/2 width=3/ qed. + +lemma lcprs_pair: ∀L1,L2. L1 ⊢ ➡* L2 → ∀V1,V2. L2 ⊢ V1 ➡* V2 → + ∀I. L1. ⓑ{I} V1 ⊢ ➡* L2. ⓑ{I} V2. +#L1 #L2 #H @(lcprs_ind … H) -L2 /2 width=1/ +#L #L2 #_ #HL2 #IHL1 #V1 #V2 #HV12 #I +@(lcprs_trans … (L.ⓑ{I}V1)) /2 width=1/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc.ma new file mode 100644 index 000000000..e61c2081d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc.ma @@ -0,0 +1,104 @@ +(**************************************************************************) +(* ___ *) +(* ||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.ma". +include "Basic_2/computation/acp_cr.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) + +inductive lsubc (RP:lenv→predicate term): relation lenv ≝ +| lsubc_atom: lsubc RP (⋆) (⋆) +| lsubc_pair: ∀I,L1,L2,V. lsubc RP L1 L2 → lsubc RP (L1. ⓑ{I} V) (L2. ⓑ{I} V) +| lsubc_abbr: ∀L1,L2,V,W,A. ⦃L1, V⦄ [RP] ϵ 〚A〛 → L2 ⊢ W ÷ A → + lsubc RP L1 L2 → lsubc RP (L1. ⓓV) (L2. ⓛW) +. + +interpretation + "local environment refinement (abstract candidates of reducibility)" + 'CrSubEq L1 RP L2 = (lsubc RP L1 L2). + +(* Basic inversion lemmas ***************************************************) + +fact lsubc_inv_atom1_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → L1 = ⋆ → L2 = ⋆. +#RP #L1 #L2 * -L1 -L2 +[ // +| #I #L1 #L2 #V #_ #H destruct +| #L1 #L2 #V #W #A #_ #_ #_ #H destruct +] +qed. + +(* Basic_1: was: csubc_gen_sort_r *) +lemma lsubc_inv_atom1: ∀RP,L2. ⋆ [RP] ⊑ L2 → L2 = ⋆. +/2 width=4/ qed-. + +fact lsubc_inv_pair1_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀I,K1,V. L1 = K1. ⓑ{I} V → + (∃∃K2. K1 [RP] ⊑ K2 & L2 = K2. ⓑ{I} V) ∨ + ∃∃K2,W,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & + K1 [RP] ⊑ K2 & + L2 = K2. ⓛW & I = Abbr. +#RP #L1 #L2 * -L1 -L2 +[ #I #K1 #V #H destruct +| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3/ +| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K1 #V #H destruct /3 width=7/ +] +qed. + +(* Basic_1: was: csubc_gen_head_r *) +lemma lsubc_inv_pair1: ∀RP,I,K1,L2,V. K1. ⓑ{I} V [RP] ⊑ L2 → + (∃∃K2. K1 [RP] ⊑ K2 & L2 = K2. ⓑ{I} V) ∨ + ∃∃K2,W,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & + K1 [RP] ⊑ K2 & + L2 = K2. ⓛW & I = Abbr. +/2 width=3/ qed-. + +fact lsubc_inv_atom2_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → L2 = ⋆ → L1 = ⋆. +#RP #L1 #L2 * -L1 -L2 +[ // +| #I #L1 #L2 #V #_ #H destruct +| #L1 #L2 #V #W #A #_ #_ #_ #H destruct +] +qed. + +(* Basic_1: was: csubc_gen_sort_l *) +lemma lsubc_inv_atom2: ∀RP,L1. L1 [RP] ⊑ ⋆ → L1 = ⋆. +/2 width=4/ qed-. + +fact lsubc_inv_pair2_aux: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀I,K2,W. L2 = K2. ⓑ{I} W → + (∃∃K1. K1 [RP] ⊑ K2 & L1 = K1. ⓑ{I} W) ∨ + ∃∃K1,V,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & + K1 [RP] ⊑ K2 & + L1 = K1. ⓓV & I = Abst. +#RP #L1 #L2 * -L1 -L2 +[ #I #K2 #W #H destruct +| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3/ +| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K2 #W #H destruct /3 width=7/ +] +qed. + +(* Basic_1: was: csubc_gen_head_l *) +lemma lsubc_inv_pair2: ∀RP,I,L1,K2,W. L1 [RP] ⊑ K2. ⓑ{I} W → + (∃∃K1. K1 [RP] ⊑ K2 & L1 = K1. ⓑ{I} W) ∨ + ∃∃K1,V,A. ⦃K1, V⦄ [RP] ϵ 〚A〛 & K2 ⊢ W ÷ A & + K1 [RP] ⊑ K2 & + L1 = K1. ⓓV & I = Abst. +/2 width=3/ qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: csubc_refl *) +lemma lsubc_refl: ∀RP,L. L [RP] ⊑ L. +#RP #L elim L -L // /2 width=1/ +qed. + +(* Basic_1: removed theorems 2: csubc_clear_conf csubc_getl_conf *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrop.ma new file mode 100644 index 000000000..03ea7b628 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrop.ma @@ -0,0 +1,68 @@ +(**************************************************************************) +(* ___ *) +(* ||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/computation/acp_cr.ma". +include "Basic_2/computation/lsubc.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) + +(* Properties concerning basic local environment slicing ********************) + +(* Basic_1: was: csubc_drop_conf_O *) +(* Note: the constant (0) can not be generalized *) +lemma lsubc_ldrop_O1_trans: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → + ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 [RP] ⊑ K2. +#RP #L1 #L2 #H elim H -L1 -L2 +[ #X #e #H + >(ldrop_inv_atom1 … H) -H /2 width=3/ +| #I #L1 #L2 #V #_ #IHL12 #X #e #H + elim (ldrop_inv_O1 … H) -H * #He #H destruct + [ elim (IHL12 L2 0 ?) -IHL12 // #X #H <(ldrop_inv_refl … H) -H /3 width=3/ + | elim (IHL12 … H) -L2 /3 width=3/ + ] +| #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #X #e #H + elim (ldrop_inv_O1 … H) -H * #He #H destruct + [ elim (IHL12 L2 0 ?) -IHL12 // #X #H <(ldrop_inv_refl … H) -H /3 width=7/ + | elim (IHL12 … H) -L2 /3 width=3/ + ] +qed-. + +(* Basic_1: was: csubc_drop_conf_rev *) +lemma ldrop_lsubc_trans: ∀RR,RS,RP. + acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 [RP] ⊑ K2 → + ∃∃L2. L1 [RP] ⊑ L2 & ⇩[d, e] L2 ≡ K2. +#RR #RS #RP #Hacp #Hacr #L1 #K1 #d #e #H elim H -L1 -K1 -d -e +[ #d #e #X #H + >(lsubc_inv_atom1 … H) -H /2 width=3/ +| #L1 #I #V1 #X #H + elim (lsubc_inv_pair1 … H) -H * + [ #K1 #HLK1 #H destruct /3 width=3/ + | #K1 #W1 #A #HV1 #HW1 #HLK1 #H1 #H2 destruct /3 width=3/ + ] +| #L1 #K1 #I #V1 #e #_ #IHLK1 #K2 #HK12 + elim (IHLK1 … HK12) -K1 /3 width=5/ +| #L1 #K1 #I #V1 #V2 #d #e #HLK1 #HV21 #IHLK1 #X #H + elim (lsubc_inv_pair1 … H) -H * + [ #K2 #HK12 #H destruct + elim (IHLK1 … HK12) -K1 /3 width=5/ + | #K2 #W2 #A #HV2 #HW2 #HK12 #H1 #H2 destruct + elim (IHLK1 … HK12) #K #HL1K #HK2 + lapply (aacr_acr … Hacp Hacr A) -Hacp -Hacr #HA + lapply (s7 … HA … HV2 … HLK1 HV21) -HV2 + elim (lift_total W2 d e) /4 width=9/ + ] +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrops.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrops.ma new file mode 100644 index 000000000..95879a234 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrops.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||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/computation/lsubc_ldrop.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) + +(* Properties concerning generic local environment slicing ******************) + +(* Basic_1: was: csubc_drop1_conf_rev *) +lemma ldrops_lsubc_trans: ∀RR,RS,RP. + acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L1,K1,des. ⇩*[des] L1 ≡ K1 → ∀K2. K1 [RP] ⊑ K2 → + ∃∃L2. L1 [RP] ⊑ L2 & ⇩*[des] L2 ≡ K2. +#RR #RS #RP #Hacp #Hacr #L1 #K1 #des #H elim H -L1 -K1 -des +[ /2 width=3/ +| #L1 #L #K1 #des #d #e #_ #HLK1 #IHL #K2 #HK12 + elim (ldrop_lsubc_trans … Hacp Hacr … HLK1 … HK12) -Hacp -Hacr -K1 #K #HLK #HK2 + elim (IHL … HLK) -IHL -HLK /3 width=5/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_lsuba.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_lsuba.ma new file mode 100644 index 000000000..41c906285 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_lsuba.ma @@ -0,0 +1,26 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lsuba.ma". +include "Basic_2/computation/acp_aaa.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) + +(* properties concerning lenv refinement for atomic arity assignment ********) + +lemma lsubc_lsuba: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L1,L2. L1 ÷⊑ L2 → L1 [RP] ⊑ L2. +#RR #RS #RP #H1RP #H2RP #L1 #L2 #H elim H -L1 -L2 +// /2 width=1/ /3 width=4/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/conversion/cpc.ma b/matita/matita/contribs/lambda_delta/basic_2/conversion/cpc.ma new file mode 100644 index 000000000..e3058349b --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/conversion/cpc.ma @@ -0,0 +1,37 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************) + +definition cpc: lenv → relation term ≝ + λL,T1,T2. L ⊢ T1 ➡ T2 ∨ L ⊢ T2 ➡ T1. + +interpretation + "context-sensitive parallel conversion (term)" + 'PConv L T1 T2 = (cpc L T1 T2). + +(* Basic properties *********************************************************) + +lemma cpc_refl: ∀L,T. L ⊢ T ⬌ T. +/2 width=1/ qed. + +lemma cpc_sym: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → L ⊢ T2 ⬌ T1. +#L #T1 #T2 * /2 width=1/ +qed. + +lemma cpc_cpr: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → ∃∃T. L ⊢ T1 ➡ T & L ⊢ T2 ➡ T. +#L #T1 #T2 * /2 width=3/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/conversion/cpc_cpc.ma b/matita/matita/contribs/lambda_delta/basic_2/conversion/cpc_cpc.ma new file mode 100644 index 000000000..b4920c35c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/conversion/cpc_cpc.ma @@ -0,0 +1,23 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/conversion/cpc.ma". + +(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************) + +(* Main properties **********************************************************) + +theorem cpc_conf: ∀L,T0,T1,T2. L ⊢ T0 ⬌ T1 → L ⊢ T0 ⬌ T2 → + ∃∃T. L ⊢ T1 ⬌ T & L ⊢ T2 ⬌ T. +/3 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/equivalence/cpcs.ma b/matita/matita/contribs/lambda_delta/basic_2/equivalence/cpcs.ma new file mode 100644 index 000000000..6004ee539 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/equivalence/cpcs.ma @@ -0,0 +1,47 @@ +(**************************************************************************) +(* ___ *) +(* ||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/conversion/cpc.ma". + +(* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************) + +definition cpcs: lenv → relation term ≝ + λL. TC … (cpc L). + +interpretation "context-sensitive parallel equivalence (term)" + 'PConvStar L T1 T2 = (cpcs L T1 T2). + +(* Basic eliminators ********************************************************) + +lemma cpcs_ind: ∀L,T1. ∀R:predicate term. R T1 → + (∀T,T2. L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → R T → R T2) → + ∀T2. L ⊢ T1 ⬌* T2 → R T2. +#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // +qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: pc3_refl *) +lemma cpcs_refl: ∀L,T. L ⊢ T ⬌* T. +/2 width=1/ qed. + +lemma cpcs_strap1: ∀L,T1,T,T2. + L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → L ⊢ T1 ⬌* T2. +/2 width=3/ qed. + +lemma cpcs_strap2: ∀L,T1,T,T2. + L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +/2 width=3/ qed. + +(* Basic_1: removed theorems 1: clear_pc3_trans *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/aarity.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/aarity.ma new file mode 100644 index 000000000..15418582b --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/aarity.ma @@ -0,0 +1,77 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +(* THE FORMAL SYSTEM λδ - MATITA SOURCE FILES + * Support for abstract candidates of reducibility closed: 2012 January 27 + * Confluence of context-sensitive parallel reduction closed: 2011 September 21 + * Confluence of context-free parallel reduction closed: 2011 September 6 + * Specification started: 2011 April 17 + * - Patience on me to gain peace and perfection! - + * [ suggested invocation to start formal specifications with ] + *) + +include "Ground_2/star.ma". +include "Basic_2/notation.ma". + +(* ATOMIC ARITY *************************************************************) + +inductive aarity: Type[0] ≝ + | AAtom: aarity (* atomic aarity construction *) + | APair: aarity → aarity → aarity (* binary aarity construction *) +. + +interpretation "aarity construction (atomic)" + 'Item0 = AAtom. + +interpretation "aarity construction (binary)" + 'SnItem2 A1 A2 = (APair A1 A2). + +(* Basic inversion lemmas ***************************************************) + +lemma discr_apair_xy_x: ∀A,B. ②B. A = B → False. +#A #B elim B -B +[ #H destruct +| #Y #X #IHY #_ #H destruct + -H >e0 in e1; normalize (**) (* destruct: one quality is not simplified, the destucted equality is not erased *) + /2 width=1/ +] +qed-. + +lemma discr_tpair_xy_y: ∀B,A. ②B. A = A → False. +#B #A elim A -A +[ #H destruct +| #Y #X #_ #IHX #H destruct + -H (**) (* destruct: the destucted equality is not erased *) + /2 width=1/ +] +qed-. + +(* Basic properties *********************************************************) + +lemma aarity_eq_dec: ∀A1,A2:aarity. Decidable (A1 = A2). +#A1 elim A1 -A1 +[ #A2 elim A2 -A2 /2 width=1/ + #B2 #A2 #_ #_ @or_intror #H destruct +| #B1 #A1 #IHB1 #IHA1 #A2 elim A2 -A2 + [ -IHB1 -IHA1 @or_intror #H destruct + | #B2 #A2 #_ #_ elim (IHB1 B2) -IHB1 + [ #H destruct elim (IHA1 A2) -IHA1 + [ #H destruct /2 width=1/ + | #HA12 @or_intror #H destruct /2 width=1/ + ] + | -IHA1 #HB12 @or_intror #H destruct /2 width=1/ + ] + ] +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/cl_shift.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/cl_shift.ma new file mode 100644 index 000000000..3a259cecf --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/cl_shift.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.ma". + +(* SHIFT OF A CLOSURE *******************************************************) + +let rec shift L T on L ≝ match L with +[ LAtom ⇒ T +| LPair L I V ⇒ shift L (ⓑ{I} V. T) +]. + +interpretation "shift (closure)" 'Append L T = (shift L T). diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/cl_weight.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/cl_weight.ma new file mode 100644 index 000000000..925b84a34 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/cl_weight.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_weight.ma". +include "Basic_2/grammar/cl_shift.ma". + +(* WEIGHT OF A CLOSURE ******************************************************) + +definition cw: lenv → term → ? ≝ λL,T. #[L] + #[T]. + +interpretation "weight (closure)" 'Weight L T = (cw L T). + +(* Basic properties *********************************************************) + +(* Basic_1: was: flt_wf__q_ind *) + +(* Basic_1: was: flt_wf_ind *) +axiom cw_wf_ind: ∀R:lenv→predicate term. + (∀L2,T2. (∀L1,T1. #[L1,T1] < #[L2,T2] → R L1 T1) → R L2 T2) → + ∀L,T. R L T. + +(* Basic_1: was: flt_shift *) +lemma cw_shift: ∀K,I,V,T. #[K. ⓑ{I} V, T] < #[K, ②{I} V. T]. +normalize // +qed. + +lemma tw_shift: ∀L,T. #[L, T] ≤ #[L @ T]. +#L elim L // +#K #I #V #IHL #T +@transitive_le [3: @IHL |2: /2 width=2/ | skip ] +qed. + +(* Basic_1: removed theorems 6: + flt_thead_sx flt_thead_dx flt_arith0 flt_arith1 flt_arith2 flt_trans +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/genv.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/genv.ma new file mode 100644 index 000000000..349f8708a --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/genv.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 "Ground_2/list.ma". +include "Basic_2/grammar/term.ma". + +(* GLOBAL ENVIRONMENTS ******************************************************) + +(* global environments *) +definition genv ≝ list2 bind2 term. + +interpretation "sort (global environment)" + 'Star = (nil2 bind2 term). + +interpretation "environment construction (binary)" + 'DxItem2 L I T = (cons2 bind2 term I T L). + +interpretation "environment binding construction (binary)" + 'DxBind2 L I T = (cons2 bind2 term I T L). + +interpretation "abbreviation (global environment)" + 'DxAbbr L T = (cons2 bind2 term Abbr T L). + +interpretation "abstraction (global environment)" + 'DxAbst L T = (cons2 bind2 term Abst T L). + +(* Basic properties *********************************************************) + +axiom genv_eq_dec: ∀T1,T2:genv. Decidable (T1 = T2). diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/item.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/item.ma new file mode 100644 index 000000000..7e4ee3839 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/item.ma @@ -0,0 +1,67 @@ +(**************************************************************************) +(* ___ *) +(* ||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/arith.ma". +include "Basic_2/notation.ma". + +(* ITEMS ********************************************************************) + +(* atomic items *) +inductive item0: Type[0] ≝ + | Sort: nat → item0 (* sort: starting at 0 *) + | LRef: nat → item0 (* reference by index: starting at 0 *) + | GRef: nat → item0 (* reference by position: starting at 0 *) +. + +(* binary binding items *) +inductive bind2: Type[0] ≝ + | Abbr: bind2 (* abbreviation *) + | Abst: bind2 (* abstraction *) +. + +(* binary non-binding items *) +inductive flat2: Type[0] ≝ + | Appl: flat2 (* application *) + | Cast: flat2 (* explicit type annotation *) +. + +(* binary items *) +inductive item2: Type[0] ≝ + | Bind2: bind2 → item2 (* binding item *) + | Flat2: flat2 → item2 (* non-binding item *) +. + +coercion Bind2. + +coercion Flat2. + +(* Basic properties *********************************************************) + +axiom item0_eq_dec: ∀I1,I2:item0. Decidable (I1 = I2). + +(* Basic_1: was: bind_dec *) +axiom bind2_eq_dec: ∀I1,I2:bind2. Decidable (I1 = I2). + +(* Basic_1: was: flat_dec *) +axiom flat2_eq_dec: ∀I1,I2:flat2. Decidable (I1 = I2). + +(* Basic_1: was: kind_dec *) +axiom item2_eq_dec: ∀I1,I2:item2. Decidable (I1 = I2). + +(* Basic_1: removed theorems 21: + s_S s_plus s_plus_sym s_minus minus_s_s s_le s_lt s_inj s_inc + s_arith0 s_arith1 + r_S r_plus r_plus_sym r_minus r_dis s_r r_arith0 r_arith1 + not_abbr_abst bind_dec_not +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv.ma new file mode 100644 index 000000000..5aacaf57c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv.ma @@ -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/grammar/term.ma". + +(* LOCAL ENVIRONMENTS *******************************************************) + +(* local environments *) +inductive lenv: Type[0] ≝ +| LAtom: lenv (* empty *) +| LPair: lenv → bind2 → term → lenv (* binary binding construction *) +. + +interpretation "sort (local environment)" + 'Star = LAtom. + +interpretation "environment construction (binary)" + 'DxItem2 L I T = (LPair L I T). + +interpretation "environment binding construction (binary)" + 'DxBind2 L I T = (LPair L I T). + +interpretation "abbreviation (local environment)" + 'DxAbbr L T = (LPair L Abbr T). + +interpretation "abstraction (local environment)" + 'DxAbst L T = (LPair L Abst T). diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_length.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_length.ma new file mode 100644 index 000000000..357301333 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_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.ma". + +(* LENGTH OF A LOCAL ENVIRONMENT ********************************************) + +let rec length L ≝ match L with +[ LAtom ⇒ 0 +| LPair L _ _ ⇒ length L + 1 +]. + +interpretation "length (local environment)" 'card L = (length L). diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_weight.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_weight.ma new file mode 100644 index 000000000..414d3f1ab --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_weight.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_weight.ma". +include "Basic_2/grammar/lenv.ma". + +(* WEIGHT OF A LOCAL ENVIRONMENT ********************************************) + +let rec lw L ≝ match L with +[ LAtom ⇒ 0 +| LPair L _ V ⇒ lw L + #[V] +]. + +interpretation "weight (local environment)" 'Weight L = (lw L). + +(* Basic_1: removed theorems 2: clt_cong clt_head *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/lsubs.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/lsubs.ma new file mode 100644 index 000000000..2c27f0f9e --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/lsubs.ma @@ -0,0 +1,96 @@ +(**************************************************************************) +(* ___ *) +(* ||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". + +(* LOCAL ENVIRONMENT REFINEMENT FOR SUBSTITUTION ****************************) + +inductive lsubs: nat → nat → relation lenv ≝ +| lsubs_sort: ∀d,e. lsubs d e (⋆) (⋆) +| lsubs_OO: ∀L1,L2. lsubs 0 0 L1 L2 +| lsubs_abbr: ∀L1,L2,V,e. lsubs 0 e L1 L2 → + lsubs 0 (e + 1) (L1. ⓓV) (L2.ⓓV) +| lsubs_abst: ∀L1,L2,I,V1,V2,e. lsubs 0 e L1 L2 → + lsubs 0 (e + 1) (L1. ⓛV1) (L2.ⓑ{I} V2) +| lsubs_skip: ∀L1,L2,I1,I2,V1,V2,d,e. + lsubs d e L1 L2 → lsubs (d + 1) e (L1. ⓑ{I1} V1) (L2. ⓑ{I2} V2) +. + +interpretation + "local environment refinement (substitution)" + 'SubEq L1 d e L2 = (lsubs d e L1 L2). + +definition lsubs_conf: ∀S. (lenv → relation S) → Prop ≝ λS,R. + ∀L1,s1,s2. R L1 s1 s2 → + ∀L2,d,e. L1 [d, e] ≼ L2 → R L2 s1 s2. + +(* Basic properties *********************************************************) + +lemma TC_lsubs_conf: ∀S,R. lsubs_conf S R → lsubs_conf S (λL. (TC … (R L))). +#S #R #HR #L1 #s1 #s2 #H elim H -s2 +[ /3 width=5/ +| #s #s2 #_ #Hs2 #IHs1 #L2 #d #e #HL12 + lapply (HR … Hs2 … HL12) -HR -Hs2 -HL12 /3 width=3/ +] +qed. + +lemma lsubs_bind_eq: ∀L1,L2,e. L1 [0, e] ≼ L2 → ∀I,V. + L1. ⓑ{I} V [0, e + 1] ≼ L2.ⓑ{I} V. +#L1 #L2 #e #HL12 #I #V elim I -I /2 width=1/ +qed. + +lemma lsubs_refl: ∀d,e,L. L [d, e] ≼ L. +#d elim d -d +[ #e elim e -e // #e #IHe #L elim L -L // /2 width=1/ +| #d #IHd #e #L elim L -L // /2 width=1/ +] +qed. + +lemma lsubs_skip_lt: ∀L1,L2,d,e. L1 [d - 1, e] ≼ L2 → 0 < d → + ∀I1,I2,V1,V2. L1. ⓑ{I1} V1 [d, e] ≼ L2. ⓑ{I2} V2. + +#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) // /2 width=1/ +qed. + +(* Basic inversion lemmas ***************************************************) + +(* Basic forward lemmas ***************************************************) + +fact lsubs_fwd_length_full1_aux: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → + d = 0 → e = |L1| → |L1| ≤ |L2|. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize +[ // +| /2 width=1/ +| /3 width=1/ +| /3 width=1/ +| #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma lsubs_fwd_length_full1: ∀L1,L2. L1 [0, |L1|] ≼ L2 → |L1| ≤ |L2|. +/2 width=5/ qed-. + +fact lsubs_fwd_length_full2_aux: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → + d = 0 → e = |L2| → |L2| ≤ |L1|. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize +[ // +| /2 width=1/ +| /3 width=1/ +| /3 width=1/ +| #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma lsubs_fwd_length_full2: ∀L1,L2. L1 [0, |L2|] ≼ L2 → |L2| ≤ |L1|. +/2 width=5/ qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/term.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/term.ma new file mode 100644 index 000000000..6e8370d49 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/term.ma @@ -0,0 +1,117 @@ +(**************************************************************************) +(* ___ *) +(* ||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/item.ma". + +(* TERMS ********************************************************************) + +(* terms *) +inductive term: Type[0] ≝ + | TAtom: item0 → term (* atomic item construction *) + | TPair: item2 → term → term → term (* binary item construction *) +. + +interpretation "term construction (atomic)" + 'Item0 I = (TAtom I). + +interpretation "term construction (binary)" + 'SnItem2 I T1 T2 = (TPair I T1 T2). + +interpretation "term binding construction (binary)" + 'SnBind2 I T1 T2 = (TPair (Bind2 I) T1 T2). + +interpretation "term flat construction (binary)" + 'SnFlat2 I T1 T2 = (TPair (Flat2 I) T1 T2). + +interpretation "sort (term)" + 'Star k = (TAtom (Sort k)). + +interpretation "local reference (term)" + 'LRef i = (TAtom (LRef i)). + +interpretation "global reference (term)" + 'GRef p = (TAtom (GRef p)). + +interpretation "abbreviation (term)" + 'SnAbbr T1 T2 = (TPair (Bind2 Abbr) T1 T2). + +interpretation "abstraction (term)" + 'SnAbst T1 T2 = (TPair (Bind2 Abst) T1 T2). + +interpretation "application (term)" + 'SnAppl T1 T2 = (TPair (Flat2 Appl) T1 T2). + +interpretation "native type annotation (term)" + 'SnCast T1 T2 = (TPair (Flat2 Cast) T1 T2). + +(* Basic properties *********************************************************) + +(* Basic_1: was: term_dec *) +axiom term_eq_dec: ∀T1,T2:term. Decidable (T1 = T2). + +(* Basic inversion lemmas ***************************************************) + +lemma discr_tpair_xy_x: ∀I,T,V. ②{I} V. T = V → False. +#I #T #V elim V -V +[ #J #H destruct +| #J #W #U #IHW #_ #H destruct + -H >e0 in e1; normalize (**) (* destruct: one quality is not simplified, the destucted equality is not erased *) + /2 width=1/ +] +qed-. + +(* Basic_1: was: thead_x_y_y *) +lemma discr_tpair_xy_y: ∀I,V,T. ②{I} V. T = T → False. +#I #V #T elim T -T +[ #J #H destruct +| #J #W #U #_ #IHU #H destruct + -H (**) (* destruct: the destucted equality is not erased *) + /2 width=1/ +] +qed-. + +lemma eq_false_inv_tpair_sn: ∀I,V1,T1,V2,T2. + (②{I} V1. T1 = ②{I} V2. T2 → False) → + (V1 = V2 → False) ∨ (V1 = V2 ∧ (T1 = T2 → False)). +#I #V1 #T1 #V2 #T2 #H +elim (term_eq_dec V1 V2) /3 width=1/ #HV12 destruct +@or_intror @conj // #HT12 destruct /2 width=1/ +qed-. + +lemma eq_false_inv_tpair_dx: ∀I,V1,T1,V2,T2. + (②{I} V1. T1 = ②{I} V2. T2 → False) → + (T1 = T2 → False) ∨ (T1 = T2 ∧ (V1 = V2 → False)). +#I #V1 #T1 #V2 #T2 #H +elim (term_eq_dec T1 T2) /3 width=1/ #HT12 destruct +@or_intror @conj // #HT12 destruct /2 width=1/ +qed-. + +lemma eq_false_inv_beta: ∀V1,V2,W1,W2,T1,T2. + (ⓐV1. ⓛW1. T1 = ⓐV2. ⓛW2 .T2 →False) → + (W1 = W2 → False) ∨ + (W1 = W2 ∧ (ⓓV1. T1 = ⓓV2. T2 → False)). +#V1 #V2 #W1 #W2 #T1 #T2 #H +elim (eq_false_inv_tpair_sn … H) -H +[ #HV12 elim (term_eq_dec W1 W2) /3 width=1/ + #H destruct @or_intror @conj // #H destruct /2 width=1/ +| * #H1 #H2 destruct + elim (eq_false_inv_tpair_sn … H2) -H2 /3 width=1/ + * #H #HT12 destruct + @or_intror @conj // #H destruct /2 width=1/ +] +qed. + +(* Basic_1: removed theorems 3: + not_void_abst not_abbr_void not_abst_void +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma new file mode 100644 index 000000000..9b6c08270 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma @@ -0,0 +1,41 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term.ma". + +(* SIMPLE (NEUTRAL) TERMS ***************************************************) + +inductive simple: predicate term ≝ + | simple_atom: ∀I. simple (⓪{I}) + | simple_flat: ∀I,V,T. simple (ⓕ{I} V. T) +. + +interpretation "simple (term)" 'Simple T = (simple T). + +(* Basic inversion lemmas ***************************************************) + +fact simple_inv_bind_aux: ∀T. 𝐒[T] → ∀J,W,U. T = ⓑ{J} W. U → False. +#T * -T +[ #I #J #W #U #H destruct +| #I #V #T #J #W #U #H destruct +] +qed. + +lemma simple_inv_bind: ∀I,V,T. 𝐒[ⓑ{I} V. T] → False. +/2 width=6/ qed-. + +lemma simple_inv_pair: ∀I,V,T. 𝐒[②{I}V.T] → ∃J. I = Flat2 J. +* /2 width=2/ #I #V #T #H +elim (simple_inv_bind … H) +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/term_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_vector.ma new file mode 100644 index 000000000..2b3f8880e --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_vector.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||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/list.ma". +include "Basic_2/grammar/term.ma". + +(* TERMS ********************************************************************) + +let rec applv Vs T on Vs ≝ + match Vs with + [ nil ⇒ T + | cons hd tl ⇒ ⓐhd. (applv tl T) + ]. + +interpretation "application o vevtor (term)" + 'SnApplV Vs T = (applv Vs T). diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/term_weight.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_weight.ma new file mode 100644 index 000000000..1cf414b38 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/term_weight.ma @@ -0,0 +1,43 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/grammar/term.ma". + +(* WEIGHT OF A TERM *********************************************************) + +let rec tw T ≝ match T with +[ TAtom _ ⇒ 1 +| TPair _ V T ⇒ tw V + tw T + 1 +]. + +interpretation "weight (term)" 'Weight T = (tw T). + +(* Basic properties *********************************************************) + +(* Basic_1: was: tweight_lt *) +lemma tw_pos: ∀T. 1 ≤ #[T]. +#T elim T -T // +qed. + +(* Basic eliminators ********************************************************) + +axiom tw_wf_ind: ∀R:predicate term. + (∀T2. (∀T1. #[T1] < #[T2] → R T1) → R T2) → + ∀T. R T. + +(* Basic_1: removed theorems 11: + wadd_le wadd_lt wadd_O weight_le weight_eq weight_add_O + weight_add_S tlt_trans tlt_head_sx tlt_head_dx tlt_wf_ind + removed local theorems 1: q_ind +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/tshf.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/tshf.ma new file mode 100644 index 000000000..364a530e1 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/tshf.ma @@ -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/grammar/term_simple.ma". + +(* SAME HEAD TERM FORMS *****************************************************) + +inductive tshf: relation term ≝ + | tshf_atom: ∀I. tshf (⓪{I}) (⓪{I}) + | tshf_abst: ∀V1,V2,T1,T2. tshf (ⓛV1. T1) (ⓛV2. T2) + | tshf_appl: ∀V1,V2,T1,T2. tshf T1 T2 → 𝐒[T1] → 𝐒[T2] → + tshf (ⓐV1. T1) (ⓐV2. T2) +. + +interpretation "same head form (term)" 'napart T1 T2 = (tshf T1 T2). + +(* Basic properties *********************************************************) + +lemma tshf_sym: ∀T1,T2. T1 ≈ T2 → T2 ≈ T1. +#T1 #T2 #H elim H -T1 -T2 /2 width=1/ +qed. + +lemma tshf_refl2: ∀T1,T2. T1 ≈ T2 → T2 ≈ T2. +#T1 #T2 #H elim H -T1 -T2 // /2 width=1/ +qed. + +lemma tshf_refl1: ∀T1,T2. T1 ≈ T2 → T1 ≈ T1. +/3 width=2/ qed. + +lemma simple_tshf_repl_dx: ∀T1,T2. T1 ≈ T2 → 𝐒[T1] → 𝐒[T2]. +#T1 #T2 #H elim H -T1 -T2 // +#V1 #V2 #T1 #T2 #H +elim (simple_inv_bind … H) +qed. (**) (* remove from index *) + +lemma simple_tshf_repl_sn: ∀T1,T2. T1 ≈ T2 → 𝐒[T2] → 𝐒[T1]. +/3 width=3/ qed-. + +(* Basic inversion lemmas ***************************************************) + +fact tshf_inv_bind1_aux: ∀T1,T2. T1 ≈ T2 → ∀I,W1,U1. T1 = ⓑ{I}W1.U1 → + ∃∃W2,U2. I = Abst & T2 = ⓛW2. U2. +#T1 #T2 * -T1 -T2 +[ #J #I #W1 #U1 #H destruct +| #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/ +| #V1 #V2 #T1 #T2 #H_ #_ #_ #I #W1 #U1 #H destruct +] +qed. + +lemma tshf_inv_bind1: ∀I,W1,U1,T2. ⓑ{I}W1.U1 ≈ T2 → + ∃∃W2,U2. I = Abst & T2 = ⓛW2. U2. +/2 width=5/ qed-. + +fact tshf_inv_flat1_aux: ∀T1,T2. T1 ≈ T2 → ∀I,W1,U1. T1 = ⓕ{I}W1.U1 → + ∃∃W2,U2. U1 ≈ U2 & 𝐒[U1] & 𝐒[U2] & + I = Appl & T2 = ⓐW2. U2. +#T1 #T2 * -T1 -T2 +[ #J #I #W1 #U1 #H destruct +| #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct +| #V1 #V2 #T1 #T2 #HT12 #HT1 #HT2 #I #W1 #U1 #H destruct /2 width=5/ +] +qed. + +lemma tshf_inv_flat1: ∀I,W1,U1,T2. ⓕ{I}W1.U1 ≈ T2 → + ∃∃W2,U2. U1 ≈ U2 & 𝐒[U1] & 𝐒[U2] & + I = Appl & T2 = ⓐW2. U2. +/2 width=4/ qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc.ma new file mode 100644 index 000000000..c4c44c9f3 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc.ma @@ -0,0 +1,107 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_simple.ma". + +(* SAME TOP TERM CONSTRUCTOR ************************************************) + +inductive tstc: relation term ≝ + | tstc_atom: ∀I. tstc (⓪{I}) (⓪{I}) + | tstc_pair: ∀I,V1,V2,T1,T2. tstc (②{I} V1. T1) (②{I} V2. T2) +. + +interpretation "same top constructor (term)" 'Iso T1 T2 = (tstc T1 T2). + +(* Basic inversion lemmas ***************************************************) + +fact tstc_inv_atom1_aux: ∀T1,T2. T1 ≃ T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I}. +#T1 #T2 * -T1 -T2 // +#J #V1 #V2 #T1 #T2 #I #H destruct +qed. + +(* Basic_1: was: iso_gen_sort iso_gen_lref *) +lemma tstc_inv_atom1: ∀I,T2. ⓪{I} ≃ T2 → T2 = ⓪{I}. +/2 width=3/ qed-. + +fact tstc_inv_pair1_aux: ∀T1,T2. T1 ≃ T2 → ∀I,W1,U1. T1 = ②{I}W1.U1 → + ∃∃W2,U2. T2 = ②{I}W2. U2. +#T1 #T2 * -T1 -T2 +[ #J #I #W1 #U1 #H destruct +| #J #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/ +] +qed. + +(* Basic_1: was: iso_gen_head *) +lemma tstc_inv_pair1: ∀I,W1,U1,T2. ②{I}W1.U1 ≃ T2 → + ∃∃W2,U2. T2 = ②{I}W2. U2. +/2 width=5/ qed-. + +fact tstc_inv_atom2_aux: ∀T1,T2. T1 ≃ T2 → ∀I. T2 = ⓪{I} → T1 = ⓪{I}. +#T1 #T2 * -T1 -T2 // +#J #V1 #V2 #T1 #T2 #I #H destruct +qed. + +lemma tstc_inv_atom2: ∀I,T1. T1 ≃ ⓪{I} → T1 = ⓪{I}. +/2 width=3/ qed-. + +fact tstc_inv_pair2_aux: ∀T1,T2. T1 ≃ T2 → ∀I,W2,U2. T2 = ②{I}W2.U2 → + ∃∃W1,U1. T1 = ②{I}W1. U1. +#T1 #T2 * -T1 -T2 +[ #J #I #W2 #U2 #H destruct +| #J #V1 #V2 #T1 #T2 #I #W2 #U2 #H destruct /2 width=3/ +] +qed. + +lemma tstc_inv_pair2: ∀I,T1,W2,U2. T1 ≃ ②{I}W2.U2 → + ∃∃W1,U1. T1 = ②{I}W1. U1. +/2 width=5/ qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: iso_refl *) +lemma tstc_refl: ∀T. T ≃ T. +#T elim T -T // +qed. + +lemma tstc_sym: ∀T1,T2. T1 ≃ T2 → T2 ≃ T1. +#T1 #T2 #H elim H -T1 -T2 // +qed. + +lemma tstc_dec: ∀T1,T2. Decidable (T1 ≃ T2). +* #I1 [2: #V1 #T1 ] * #I2 [2,4: #V2 #T2 ] +[ elim (item2_eq_dec I1 I2) #HI12 + [ destruct /2 width=1/ + | @or_intror #H + elim (tstc_inv_pair1 … H) -H #V #T #H destruct /2 width=1/ + ] +| @or_intror #H + lapply (tstc_inv_atom1 … H) -H #H destruct +| @or_intror #H + lapply (tstc_inv_atom2 … H) -H #H destruct +| elim (item0_eq_dec I1 I2) #HI12 + [ destruct /2 width=1/ + | @or_intror #H + lapply (tstc_inv_atom2 … H) -H #H destruct /2 width=1/ + ] +] +qed. + +lemma simple_tstc_repl_dx: ∀T1,T2. T1 ≃ T2 → 𝐒[T1] → 𝐒[T2]. +#T1 #T2 * -T1 -T2 // +#I #V1 #V2 #T1 #T2 #H +elim (simple_inv_pair … H) -H #J #H destruct // +qed. (**) (* remove from index *) + +lemma simple_tstc_repl_sn: ∀T1,T2. T1 ≃ T2 → 𝐒[T2] → 𝐒[T1]. +/3 width=3/ qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc_tstc.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc_tstc.ma new file mode 100644 index 000000000..a32d045b5 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc_tstc.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tstc.ma". + +(* SAME TOP TERM CONSTRUCTOR ************************************************) + +(* Main properties **********************************************************) + +(* Basic_1: was: iso_trans *) +theorem tstc_trans: ∀T1,T. T1 ≃ T → ∀T2. T ≃ T2 → T1 ≃ T2. +#T1 #T * -T1 -T // +#I #V1 #V #T1 #T #X #H +elim (tstc_inv_pair1 … H) -H #V2 #T2 #H destruct // +qed. + +theorem tstc_canc_sn: ∀T,T1. T ≃ T1 → ∀T2. T ≃ T2 → T1 ≃ T2. +/3 width=3/ qed. + +theorem tstc_canc_dx: ∀T1,T. T1 ≃ T → ∀T2. T2 ≃ T → T1 ≃ T2. +/3 width=3/ qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc_vector.ma new file mode 100644 index 000000000..a3b38fcbf --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/tstc_vector.ma @@ -0,0 +1,33 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_vector.ma". +include "Basic_2/grammar/tstc.ma". + +(* SAME TOP TERM CONSTRUCTOR ************************************************) + +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was only: iso_flats_lref_bind_false iso_flats_flat_bind_false *) +lemma tstc_inv_bind_appls_simple: ∀I,Vs,V2,T1,T2. ⒶVs.T1 ≃ ⓑ{I} V2. T2 → + 𝐒[T1] → False. +#I #Vs #V2 #T1 #T2 #H +elim (tstc_inv_pair2 … H) -H #V0 #T0 +elim Vs -Vs normalize +[ #H destruct #H + @(simple_inv_bind … H) +| #V #Vs #_ #H destruct +] +qed. + diff --git a/matita/matita/contribs/lambda_delta/basic_2/names.txt b/matita/matita/contribs/lambda_delta/basic_2/names.txt new file mode 100644 index 000000000..dfa83ebdd --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/names.txt @@ -0,0 +1,40 @@ +NAMING CONVENTIONS FOR METAVARIABLES + +A,B : arity +C,D : candidate of reducibility +E,F : RTM environment +G : global environment +H : reserved: transient premise +IH : reserved: inductive premise +I,J : item +K,L : local environment +M,N : reserved: future use +O : +P,Q : reserved: future use +R : generic predicate (relation) +S : RTM stack +T,U,V,W: term +X,Y,Z : reserved: transient objet denoted by a capital letter + +d : relocation depth +e : relocation height +h : sort hierarchy parameter +i,j : local reference position index (de Bruijn's) +k : sort index +p,q : global reference position +t,u : local reference position level (de Bruijn's) +x,y,z : reserved: transient objet denoted by a small letter + +NAMING CONVENTIONS FOR CONSTRUCTORS + +0: atomic +2: binary + +A: application to vector + +a: application +b: binder +d: abbreviation +f: flat +l: abstraction +t: native type annotation diff --git a/matita/matita/contribs/lambda_delta/basic_2/notation.ma b/matita/matita/contribs/lambda_delta/basic_2/notation.ma new file mode 100644 index 000000000..a9fca999c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/notation.ma @@ -0,0 +1,307 @@ +(**************************************************************************) +(* ___ *) +(* ||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 λδ ****************************************) + +(* Grammar ******************************************************************) + +notation "hvbox( ⓪ )" + non associative with precedence 90 + for @{ 'Item0 }. + +notation "hvbox( ⓪ { I } )" + non associative with precedence 90 + for @{ 'Item0 $I }. + +notation "hvbox( ⋆ )" + non associative with precedence 90 + for @{ 'Star }. + +notation "hvbox( ⋆ term 90 k )" + non associative with precedence 90 + for @{ 'Star $k }. + +notation "hvbox( # term 90 i )" + non associative with precedence 90 + for @{ 'LRef $i }. + +notation "hvbox( § term 90 p )" + non associative with precedence 90 + for @{ 'GRef $p }. + +notation "hvbox( ② term 90 T1 . break term 90 T )" + non associative with precedence 90 + for @{ 'SnItem2 $T1 $T }. + +notation "hvbox( ② { I } break term 90 T1 . break term 90 T )" + non associative with precedence 90 + for @{ 'SnItem2 $I $T1 $T }. + +notation "hvbox( ⓑ { I } break term 90 T1 . break term 90 T )" + non associative with precedence 90 + for @{ 'SnBind2 $I $T1 $T }. + +notation "hvbox( ⓕ { I } break term 90 T1 . break term 90 T )" + non associative with precedence 90 + for @{ 'SnFlat2 $I $T1 $T }. + +notation "hvbox( ⓓ term 90 T1 . break term 90 T2 )" + non associative with precedence 90 + for @{ 'SnAbbr $T1 $T2 }. + +notation "hvbox( ⓛ term 90 T1 . break term 90 T2 )" + non associative with precedence 90 + for @{ 'SnAbst $T1 $T2 }. + +notation "hvbox( ⓐ term 90 T1 . break term 90 T2 )" + non associative with precedence 90 + for @{ 'SnAppl $T1 $T2 }. + +notation "hvbox( ⓣ term 90 T1 . break term 90 T2 )" + non associative with precedence 90 + for @{ 'SnCast $T1 $T2 }. + +notation "hvbox( Ⓐ term 90 T1 . break term 90 T )" + non associative with precedence 90 + for @{ 'SnApplV $T1 $T }. + +notation > "hvbox( T . break ②{ I } break term 47 T1 )" + non associative with precedence 46 + for @{ 'DxBind2 $T $I $T1 }. + +notation "hvbox( T . break ⓑ { I } break term 90 T1 )" + non associative with precedence 89 + for @{ 'DxBind2 $T $I $T1 }. + +notation "hvbox( T1 . break ⓓ T2 )" + left associative with precedence 48 + for @{ 'DxAbbr $T1 $T2 }. + +notation "hvbox( T1 . break ⓛ T2 )" + left associative with precedence 49 + for @{ 'DxAbst $T1 $T2 }. + +notation "hvbox( T . break ④ { I } break { T1 , break T2 , break T3 } )" + non associative with precedence 47 + for @{ 'DxItem4 $T $I $T1 $T2 $T3 }. + +notation "hvbox( # [ x ] )" + non associative with precedence 90 + for @{ 'Weight $x }. + +notation "hvbox( # [ x , break y ] )" + non associative with precedence 90 + for @{ 'Weight $x $y }. + +notation "hvbox( 𝐒 [ T ] )" + non associative with precedence 45 + for @{ 'Simple $T }. + +notation "hvbox( L ⊢ break term 90 T1 ≈ break T2 )" + non associative with precedence 45 + for @{ 'Hom $L $T1 $T2 }. + +notation "hvbox( T1 ≃ break T2 )" + non associative with precedence 45 + for @{ 'Iso $T1 $T2 }. + +notation "hvbox( T1 break [ d , break e ] ≼ break T2 )" + non associative with precedence 45 + for @{ 'SubEq $T1 $d $e $T2 }. + +(* Substitution *************************************************************) + +notation "hvbox( L ⊢ break [ d , break e ] break 𝐑 [ T ] )" + non associative with precedence 45 + for @{ 'Reducible $L $d $e $T }. + +notation "hvbox( L ⊢ break [ d , break e ] break 𝐈 [ T ] )" + non associative with precedence 45 + for @{ 'NotReducible $L $d $e $T }. + +notation "hvbox( L ⊢ break [ d , break e ] break 𝐍 [ T ] )" + non associative with precedence 45 + for @{ 'Normal $L $d $e $T }. + +notation "hvbox( ⇧ [ d , break e ] break T1 ≡ break T2 )" + non associative with precedence 45 + for @{ 'RLift $d $e $T1 $T2 }. + +notation "hvbox( ⇩ [ e ] break L1 ≡ break L2 )" + non associative with precedence 45 + for @{ 'RDrop $e $L1 $L2 }. + +notation "hvbox( ⇩ [ d , break e ] break L1 ≡ break L2 )" + non associative with precedence 45 + for @{ 'RDrop $d $e $L1 $L2 }. + +notation "hvbox( T1 break [ d , break e ] ▶ break T2 )" + non associative with precedence 45 + for @{ 'PSubst $T1 $d $e $T2 }. + +notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ▶ break T2 )" + non associative with precedence 45 + for @{ 'PSubst $L $T1 $d $e $T2 }. + +(* Unfold *******************************************************************) + +notation "hvbox( @ [ T1 ] break f ≡ break T2 )" + non associative with precedence 45 + for @{ 'RAt $T1 $f $T2 }. + +notation "hvbox( T1 ▭ break T2 ≡ break T )" + non associative with precedence 45 + for @{ 'RMinus $T1 $T2 $T }. + +notation "hvbox( ⇧ * [ e ] break T1 ≡ break T2 )" + non associative with precedence 45 + for @{ 'RLiftStar $e $T1 $T2 }. + +notation "hvbox( ⇩ * [ e ] break L1 ≡ break L2 )" + non associative with precedence 45 + for @{ 'RDropStar $e $L1 $L2 }. + +notation "hvbox( T1 break [ d , break e ] ▶* break T2 )" + non associative with precedence 45 + for @{ 'PSubstStar $T1 $d $e $T2 }. + +notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ▶* break T2 )" + non associative with precedence 45 + for @{ 'PSubstStar $L $T1 $d $e $T2 }. + +notation "hvbox( T1 break [ d , break e ] ≡ break T2 )" + non associative with precedence 45 + for @{ 'TSubst $T1 $d $e $T2 }. + +notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ≡ break T2 )" + non associative with precedence 45 + for @{ 'TSubst $L $T1 $d $e $T2 }. + +(* Static Typing ************************************************************) + +notation "hvbox( L ⊢ break term 90 T ÷ break A )" + non associative with precedence 45 + for @{ 'AtomicArity $L $T $A }. + +notation "hvbox( T1 ÷ ⊑ break T2 )" + non associative with precedence 45 + for @{ 'CrSubEqA $T1 $T2 }. + +(* Reducibility *************************************************************) + +notation "hvbox( 𝐑 [ T ] )" + non associative with precedence 45 + for @{ 'Reducible $T }. + +notation "hvbox( L ⊢ break 𝐑 [ T ] )" + non associative with precedence 45 + for @{ 'Reducible $L $T }. + +notation "hvbox( 𝐈 [ T ] )" + non associative with precedence 45 + for @{ 'NotReducible $T }. + +notation "hvbox( L ⊢ break 𝐈 [ T ] )" + non associative with precedence 45 + for @{ 'NotReducible $L $T }. + +notation "hvbox( 𝐍 [ T ] )" + non associative with precedence 45 + for @{ 'Normal $T }. + +notation "hvbox( L ⊢ break 𝐍 [ T ] )" + non associative with precedence 45 + for @{ 'Normal $L $T }. + +notation "hvbox( 𝐖𝐇𝐑 [ T ] )" + non associative with precedence 45 + for @{ 'WHdReducible $T }. + +notation "hvbox( L ⊢ break 𝐖𝐇𝐑 [ T ] )" + non associative with precedence 45 + for @{ 'WHdReducible $L $T }. + +notation "hvbox( 𝐖𝐇𝐈 [ T ] )" + non associative with precedence 45 + for @{ 'NotWHdReducible $T }. + +notation "hvbox( L ⊢ break 𝐖𝐇𝐈 [ T ] )" + non associative with precedence 45 + for @{ 'NotWHdReducible $L $T }. + +notation "hvbox( 𝐖𝐇𝐍 [ T ] )" + non associative with precedence 45 + for @{ 'WHdNormal $T }. + +notation "hvbox( L ⊢ break 𝐖𝐇𝐍 [ T ] )" + non associative with precedence 45 + for @{ 'WHdNormal $L $T }. + +notation "hvbox( T1 ➡ break T2 )" + non associative with precedence 45 + for @{ 'PRed $T1 $T2 }. + +notation "hvbox( L ⊢ break term 90 T1 ➡ break T2 )" + non associative with precedence 45 + for @{ 'PRed $L $T1 $T2 }. + +notation "hvbox( L1 ⊢ ➡ break L2 )" + non associative with precedence 45 + for @{ 'CPRed $L1 $L2 }. + +(* Computation **************************************************************) + +notation "hvbox( T1 ➡* break T2 )" + non associative with precedence 45 + for @{ 'PRedStar $T1 $T2 }. + +notation "hvbox( L ⊢ break term 90 T1 ➡* break T2 )" + non associative with precedence 45 + for @{ 'PRedStar $L $T1 $T2 }. + +notation "hvbox( L1 ⊢ ➡* break L2 )" + non associative with precedence 45 + for @{ 'CPRedStar $L1 $L2 }. + +notation "hvbox( ⬇ * T )" + non associative with precedence 45 + for @{ 'SN $T }. + +notation "hvbox( L ⊢ ⬇ * T )" + non associative with precedence 45 + for @{ 'SN $L $T }. + +notation "hvbox( L ⊢ ⬇ * * T )" + non associative with precedence 45 + for @{ 'SNStar $L $T }. + +notation "hvbox( ⦃ L, break T ⦄ break [ R ] ϵ break 〚 A 〛 )" + non associative with precedence 45 + for @{ 'InEInt $R $L $T $A }. + +notation "hvbox( T1 break [ R ] ⊑ break T2 )" + non associative with precedence 45 + for @{ 'CrSubEq $T1 $R $T2 }. + +(* Conversion ***************************************************************) + +notation "hvbox( L ⊢ break term 90 T1 ⬌ break T2 )" + non associative with precedence 45 + for @{ 'PConv $L $T1 $T2 }. + +(* Equivalence **************************************************************) + +notation "hvbox( L ⊢ break term 90 T1 ⬌* break T2 )" + non associative with precedence 45 + for @{ 'PConvStar $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf.ma new file mode 100644 index 000000000..39ac692cd --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf.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/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) + +definition cnf: lenv → predicate term ≝ λL. NF … (cpr L) (eq …). + +interpretation + "context-sensitive normality (term)" + 'Normal L T = (cnf L T). + +(* Basic properties *********************************************************) + +(* Basic_1: was: nf2_sort *) +lemma cnf_sort: ∀L,k. L ⊢ 𝐍[⋆k]. +#L #k #X #H +>(cpr_inv_sort1 … H) // +qed. + +axiom cnf_dec: ∀L,T1. L ⊢ 𝐍[T1] ∨ + ∃∃T2. L ⊢ T1 ➡ T2 & (T1 = T2 → False). + +(* Basic_1: removed theorems 1: nf2_abst_shift *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_lift.ma new file mode 100644 index 000000000..b492eb9af --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_lift.ma @@ -0,0 +1,61 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr_lift.ma". +include "Basic_2/reducibility/cnf.ma". + +(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) + +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was only: nf2_csort_lref *) +lemma cnf_lref_atom: ∀L,i. ⇩[0, i] L ≡ ⋆ → L ⊢ 𝐍[#i]. +#L #i #HLK #X #H +elim (cpr_inv_lref1 … H) // * +#K0 #V0 #V1 #HLK0 #_ #_ #_ +lapply (ldrop_mono … HLK … HLK0) -L #H destruct +qed. + +(* Basic_1: was: nf2_lref_abst *) +lemma cnf_lref_abst: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓛV → L ⊢ 𝐍[#i]. +#L #K #V #i #HLK #X #H +elim (cpr_inv_lref1 … H) // * +#K0 #V0 #V1 #HLK0 #_ #_ #_ +lapply (ldrop_mono … HLK … HLK0) -L #H destruct +qed. + +(* Basic_1: was: nf2_abst *) +lemma cnf_abst: ∀I,L,V,W,T. L ⊢ 𝐍[W] → L. ⓑ{I} V ⊢ 𝐍[T] → L ⊢ 𝐍[ⓛW.T]. +#I #L #V #W #T #HW #HT #X #H +elim (cpr_inv_abst1 … H I V) -H #W0 #T0 #HW0 #HT0 #H destruct +>(HW … HW0) -W >(HT … HT0) -T // +qed. + +(* Basic_1: was only: nf2_appl_lref *) +lemma cnf_appl_simple: ∀L,V,T. L ⊢ 𝐍[V] → L ⊢ 𝐍[T] → 𝐒[T] → L ⊢ 𝐍[ⓐV.T]. +#L #V #T #HV #HT #HS #X #H +elim (cpr_inv_appl1_simple … H ?) -H // #V0 #T0 #HV0 #HT0 #H destruct +>(HV … HV0) -V >(HT … HT0) -T // +qed. + +(* Relocation properties ****************************************************) + +(* Basic_1: was: nf2_lift *) +lemma cnf_lift: ∀L0,L,T,T0,d,e. + L ⊢ 𝐍[T] → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → L0 ⊢ 𝐍[T0]. +#L0 #L #T #T0 #d #e #HLT #HL0 #HT0 #X #H +elim (cpr_inv_lift … HL0 … HT0 … H) -L0 #T1 #HT10 #HT1 +>(HLT … HT1) in HT0; -L #HT0 +>(lift_mono … HT10 … HT0) -T1 -X // +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma new file mode 100644 index 000000000..0ad09e4b4 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma @@ -0,0 +1,114 @@ +(**************************************************************************) +(* ___ *) +(* ||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_shift.ma". +include "Basic_2/unfold/tpss.ma". +include "Basic_2/reducibility/tpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Basic_1: includes: pr2_delta1 *) +definition cpr: lenv → relation term ≝ + λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T [0, |L|] ▶* T2. + +interpretation + "context-sensitive parallel reduction (term)" + 'PRed L T1 T2 = (cpr L T1 T2). + +(* Basic properties *********************************************************) + +lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 ➡ T2. +/4 width=3/ qed-. + +(* Basic_1: was by definition: pr2_free *) +lemma cpr_pr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. +/2 width=3/ qed. + +lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 ➡ T2. +/3 width=5/ qed. + +lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. +/2 width=1/ qed. + +(* Note: new property *) +(* Basic_1: was only: pr2_thin_dx *) +lemma cpr_flat: ∀I,L,V1,V2,T1,T2. + L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ +qed. + +lemma cpr_cast: ∀L,V,T1,T2. + L ⊢ T1 ➡ T2 → L ⊢ ⓣV. T1 ➡ T2. +#L #V #T1 #T2 * /3 width=3/ +qed. + +(* Note: it does not hold replacing |L1| with |L2| *) +(* Basic_1: was only: pr2_change *) +lemma cpr_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → + ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡ T2. +#L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 +lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 -HL12 /3 width=4/ +qed. + +(* Basic inversion lemmas ***************************************************) + +(* Basic_1: was: pr2_gen_csort *) +lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2. +#T1 #T2 * #T #HT normalize #HT2 +<(tpss_inv_refl_O2 … HT2) -HT2 // +qed-. + +(* Basic_1: was: pr2_gen_sort *) +lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. +#L #T2 #k * #X #H +>(tpr_inv_atom1 … H) -H #H +>(tpss_inv_sort1 … H) -H // +qed-. + +(* Basic_1: was pr2_gen_abbr *) +lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 → + (∃∃V,V2,T2. V1 ➡ V & L ⊢ V [O, |L|] ▶* V2 & + L. ⓓV ⊢ T1 ➡ T2 & + U2 = ⓓV2. T2 + ) ∨ + ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2. +#L #V1 #T1 #Y * #X #H1 #H2 +elim (tpr_inv_abbr1 … H1) -H1 * +[ #V #T0 #T #HV1 #HT10 #HT0 #H destruct + elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct + lapply (tps_lsubs_conf … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0 + lapply (tps_weak_all … HT0) -HT0 #HT0 + lapply (tpss_lsubs_conf … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2 + lapply (tpss_weak_all … HT2) -HT2 #HT2 + lapply (tpss_strap … HT0 HT2) -T /4 width=7/ +| /4 width=5/ +] +qed-. + +(* Basic_1: was: pr2_gen_cast *) +lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓣV1. T1 ➡ U2 → ( + ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & + U2 = ⓣV2. T2 + ) ∨ L ⊢ T1 ➡ U2. +#L #V1 #T1 #U2 * #X #H #HU2 +elim (tpr_inv_cast1 … H) -H /3 width=3/ +* #V #T #HV1 #HT1 #H destruct +elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ +qed-. + +(* Basic_1: removed theorems 4: + pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans + Basic_1: removed local theorems 3: + pr2_free_free pr2_free_delta pr2_delta_delta +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_cpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_cpr.ma new file mode 100644 index 000000000..92db70f0f --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_cpr.ma @@ -0,0 +1,63 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/tpr_tpr.ma". +include "Basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Advanced properties ******************************************************) + +lemma cpr_bind_sn: ∀I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 → + L ⊢ ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12 +@ex2_1_intro [2: @(tpr_delta … HV1 HT12) | skip ] /2 width=3/ (* /3 width=5/ is too slow *) +qed. + +(* Basic_1: was only: pr2_gen_cbind *) +lemma cpr_bind_dx: ∀I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 → + L ⊢ ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2 +elim (tpss_split_up … HT2 1 ? ?) -HT2 // #T0 (tpr_inv_atom1 … H) -H #H +elim (tpss_inv_lref1 … H) -H /2 width=1/ +* /3 width=6/ +qed-. + +(* Basic_1: was: pr2_gen_abst *) +lemma cpr_inv_abst1: ∀L,V1,T1,U2. L ⊢ ⓛV1. T1 ➡ U2 → ∀I,W. + ∃∃V2,T2. L ⊢ V1 ➡ V2 & L. ⓑ{I} W ⊢ T1 ➡ T2 & U2 = ⓛV2. T2. +#L #V1 #T1 #Y * #X #H1 #H2 #I #W +elim (tpr_inv_abst1 … H1) -H1 #V #T #HV1 #HT1 #H destruct +elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct +lapply (tpss_lsubs_conf … HT2 (L. ⓑ{I} W) ?) -HT2 /2 width=1/ /4 width=5/ +qed-. + +(* Basic_1: was pr2_gen_appl *) +lemma cpr_inv_appl1: ∀L,V1,U0,U2. L ⊢ ⓐV1. U0 ➡ U2 → + ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 & + U2 = ⓐV2. T2 + | ∃∃V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 & + U0 = ⓛW. T1 & + U2 = ⓓV2. T2 + | ∃∃V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 & + ⇧[0,1] V2 ≡ V & + U0 = ⓓW1. T1 & + U2 = ⓓW2. ⓐV. T2. +#L #V1 #U0 #Y * #X #H1 #H2 +elim (tpr_inv_appl1 … H1) -H1 * +[ #V #U #HV1 #HU0 #H destruct + elim (tpss_inv_flat1 … H2) -H2 #V2 #U2 #HV2 #HU2 #H destruct /4 width=5/ +| #V #W #T0 #T #HV1 #HT0 #H #H1 destruct + elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct + lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=8/ +| #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct + elim (tpss_inv_bind1 … H2) -H2 #W2 #X #HW02 #HX #HY destruct + elim (tpss_inv_flat1 … HX) -HX #V2 #T2 #HV2 #HT2 #H destruct + elim (tpss_inv_lift1_ge … HV2 … HV0 ?) -V // [3: /2 width=1/ |2: skip ] #V (ltpr_fwd_length … HL12) in HT2; #HT2 +elim (tpr_tpss_ltpr … HL12 … HT2) -L1 /3 width=3/ +qed. + +lemma cpr_ltpr_conf_tpss: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. L1 ⊢ T1 ➡ T2 → + ∀d,e,U1. L1 ⊢ T1 [d, e] ▶* U1 → + ∃∃U2. L2 ⊢ U1 ➡ U2 & L2 ⊢ T2 ➡ U2. +#L1 #L2 #HL12 #T1 #T2 #HT12 #d #e #U1 #HTU1 +elim (cpr_ltpr_conf_eq … HT12 … HL12) -HT12 #T #HT1 #HT2 +elim (cpr_tpss_ltpr … HL12 … HT1 … HTU1) -L1 -HT1 #U2 #HU12 #HTU2 +lapply (tpss_weak_all … HTU2) -HTU2 #HTU2 /3 width=5/ (**) (* /4 width=5/ is too slow *) +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_ltpss.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_ltpss.ma new file mode 100644 index 000000000..148a29645 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_ltpss.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ltpss_tpss.ma". +include "Basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Properties concerning partial unfold on local environments ***************) + +lemma ltpss_cpr_trans: ∀L1,L2,d,e. L1 [d, e] ▶* L2 → + ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡ T2. +#L1 #L2 #d #e #HL12 #T1 #T2 * +lapply (ltpss_weak_all … HL12) +<(ltpss_fwd_length … HL12) -HL12 /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr.ma new file mode 100644 index 000000000..80d14a2e5 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr.ma @@ -0,0 +1,37 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ltpss.ma". +include "Basic_2/reducibility/ltpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS *************) + +definition lcpr: relation lenv ≝ + λL1,L2. ∃∃L. L1 ➡ L & L [0, |L|] ▶* L2 +. + +interpretation + "context-sensitive parallel reduction (environment)" + 'CPRed L1 L2 = (lcpr L1 L2). + +(* Basic properties *********************************************************) + +lemma lcpr_refl: ∀L. L ⊢ ➡ L. +/2 width=3/ qed. + +(* Basic inversion lemmas ***************************************************) + +lemma lcpr_inv_atom1: ∀L2. ⋆ ⊢ ➡ L2 → L2 = ⋆. +#L2 * #L #HL >(ltpr_inv_atom1 … HL) -HL #HL2 >(ltpss_inv_atom1 … HL2) -HL2 // +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr_cpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr_cpr.ma new file mode 100644 index 000000000..7abe34742 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr_cpr.ma @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ltpss_ltpss.ma". +include "Basic_2/reducibility/cpr.ma". +include "Basic_2/reducibility/lcpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS *************) + +(* Advanced properties ****************************************************) + +lemma lcpr_pair: ∀L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. L2 ⊢ V1 ➡ V2 → + ∀I. L1. ⓑ{I} V1 ⊢ ➡ L2. ⓑ{I} V2. +#L1 #L2 * #L #HL1 #HL2 #V1 #V2 * +<(ltpss_fwd_length … HL2) /4 width=5/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma new file mode 100644 index 000000000..8e834b2da --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma @@ -0,0 +1,89 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/reducibility/tpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +inductive ltpr: relation lenv ≝ +| ltpr_stom: ltpr (⋆) (⋆) +| ltpr_pair: ∀K1,K2,I,V1,V2. + ltpr K1 K2 → V1 ➡ V2 → ltpr (K1. ⓑ{I} V1) (K2. ⓑ{I} V2) (*ⓑ*) +. + +interpretation + "context-free parallel reduction (environment)" + 'PRed L1 L2 = (ltpr L1 L2). + +(* Basic properties *********************************************************) + +lemma ltpr_refl: ∀L:lenv. L ➡ L. +#L elim L -L // /2 width=1/ +qed. + +(* Basic inversion lemmas ***************************************************) + +fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ➡ L2 → L1 = ⋆ → L2 = ⋆. +#L1 #L2 * -L1 -L2 +[ // +| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct +] +qed. + +(* Basic_1: was: wcpr0_gen_sort *) +lemma ltpr_inv_atom1: ∀L2. ⋆ ➡ L2 → L2 = ⋆. +/2 width=3/ qed-. + +fact ltpr_inv_pair1_aux: ∀L1,L2. L1 ➡ L2 → ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. +#L1 #L2 * -L1 -L2 +[ #K1 #I #V1 #H destruct +| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct /2 width=5/ +] +qed. + +(* Basic_1: was: wcpr0_gen_head *) +lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. ⓑ{I} V1 ➡ L2 → + ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. +/2 width=3/ qed-. + +fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ➡ L2 → L2 = ⋆ → L1 = ⋆. +#L1 #L2 * -L1 -L2 +[ // +| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct +] +qed. + +lemma ltpr_inv_atom2: ∀L1. L1 ➡ ⋆ → L1 = ⋆. +/2 width=3/ qed-. + +fact ltpr_inv_pair2_aux: ∀L1,L2. L1 ➡ L2 → ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. +#L1 #L2 * -L1 -L2 +[ #K2 #I #V2 #H destruct +| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct /2 width=5/ +] +qed. + +lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ➡ K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. +/2 width=3/ qed-. + +(* Basic forward lemmas *****************************************************) + +lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|. +#L1 #L2 #H elim H -L1 -L2 normalize // +qed-. + +(* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *) diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr_ldrop.ma new file mode 100644 index 000000000..090e8cfbf --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr_ldrop.ma @@ -0,0 +1,70 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/tpr_lift.ma". +include "Basic_2/reducibility/ltpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Basic_1: was: wcpr0_drop *) +lemma ltpr_ldrop_conf: ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀L2. L1 ➡ L2 → + ∃∃K2. ⇩[d, e] L2 ≡ K2 & K1 ➡ K2. +#L1 #K1 #d #e #H elim H -L1 -K1 -d -e +[ #d #e #X #H >(ltpr_inv_atom1 … H) -H /2 width=3/ +| #K1 #I #V1 #X #H + elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct /3 width=5/ +| #L1 #K1 #I #V1 #e #_ #IHLK1 #X #H + elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct + elim (IHLK1 … HL12) -L1 /3 width=3/ +| #L1 #K1 #I #V1 #W1 #d #e #_ #HWV1 #IHLK1 #X #H + elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct + elim (tpr_inv_lift … HV12 … HWV1) -V1 + elim (IHLK1 … HL12) -L1 /3 width=5/ +] +qed. + +(* Basic_1: was: wcpr0_drop_back *) +lemma ldrop_ltpr_trans: ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 ➡ K2 → + ∃∃L2. ⇩[d, e] L2 ≡ K2 & L1 ➡ L2. +#L1 #K1 #d #e #H elim H -L1 -K1 -d -e +[ #d #e #X #H >(ltpr_inv_atom1 … H) -H /2 width=3/ +| #K1 #I #V1 #X #H + elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct /3 width=5/ +| #L1 #K1 #I #V1 #e #_ #IHLK1 #K2 #HK12 + elim (IHLK1 … HK12) -K1 /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d #e #_ #HWV1 #IHLK1 #X #H + elim (ltpr_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (lift_total W2 d e) #V2 #HWV2 + lapply (tpr_lift … HW12 … HWV1 … HWV2) -W1 + elim (IHLK1 … HK12) -K1 /3 width=5/ +] +qed. + +fact ltpr_ldrop_trans_O1_aux: ∀L2,K2,d,e. ⇩[d, e] L2 ≡ K2 → ∀L1. L1 ➡ L2 → + d = 0 → ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 ➡ K2. +#L2 #K2 #d #e #H elim H -L2 -K2 -d -e +[ #d #e #X #H >(ltpr_inv_atom2 … H) -H /2 width=3/ +| #K2 #I #V2 #X #H + elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct /3 width=5/ +| #L2 #K2 #I #V2 #e #_ #IHLK2 #X #H #_ + elim (ltpr_inv_pair2 … H) -H #L1 #V1 #HL12 #HV12 #H destruct + elim (IHLK2 … HL12 ?) -L2 // /3 width=3/ +| #L2 #K2 #I #V2 #W2 #d #e #_ #_ #_ #L1 #_ + >commutative_plus normalize #H destruct +] +qed. + +lemma ltpr_ldrop_trans_O1: ∀L1,L2. L1 ➡ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → + ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 ➡ K2. +/2 width=5/ qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr_tps.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr_tps.ma new file mode 100644 index 000000000..91b30b014 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr_tps.ma @@ -0,0 +1,55 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/ltpr_ldrop.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Properties concerning parallel substitution on terms *********************) + +lemma ltpr_tps_trans: ∀L2,T1,T2,d,e. L2 ⊢ T1 [d, e] ▶ T2 → ∀L1. L1 ➡ L2 → + ∃∃T. L1 ⊢ T1 [d, e] ▶ T & T ➡ T2. +#L2 #T1 #T2 #d #e #H elim H -L2 -T1 -T2 -d -e +[ /2 width=3/ +| #L2 #K2 #V2 #W2 #i #d #e #Hdi #Hide #HLK2 #HVW2 #L1 #HL12 + elim (ltpr_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H + elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct -K2 + elim (lift_total V1 0 (i+1)) #W1 #HVW1 + lapply (tpr_lift … HV12 … HVW1 … HVW2) -V2 /3 width=4/ +| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 + elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 + elim (IHT12 (L1.ⓑ{I}V) ?) /2 width=1/ -L2 /3 width=5/ +| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 + elim (IHV12 … HL12) -IHV12 + elim (IHT12 … HL12) -L2 /3 width=5/ +] +qed. + +lemma ltpr_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 → ∀L2. L1 ➡ L2 → + ∃∃T. L2 ⊢ T1 [d, e] ▶ T & T2 ➡ T. +#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e +[ /2 width=3/ +| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L2 #HL12 + elim (ltpr_ldrop_conf … HLK1 … HL12) -L1 #X #HLK2 #H + elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct -K1 + elim (lift_total V2 0 (i+1)) #W2 #HVW2 + lapply (tpr_lift … HV12 … HVW1 … HVW2) -V1 /3 width=4/ +| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 + elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 + elim (IHT12 (L2.ⓑ{I}V) ?) /2 width=1/ -L1 /3 width=5/ +| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 + elim (IHV12 … HL12) -IHV12 + elim (IHT12 … HL12) -L1 /3 width=5/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tif.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tif.ma new file mode 100644 index 000000000..5b89755a7 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tif.ma @@ -0,0 +1,54 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/trf.ma". + +(* CONTEXT-FREE IRREDUCIBLE TERMS *******************************************) + +definition tif: predicate term ≝ λT. 𝐑[T] → False. + +interpretation "context-free irreducibility (term)" + 'NotReducible T = (tif T). + +(* Basic inversion lemmas ***************************************************) + +lemma tif_inv_abbr: ∀V,T. 𝐈[ⓓV.T] → False. +/2 width=1/ qed-. + +lemma tif_inv_abst: ∀V,T. 𝐈[ⓛV.T] → 𝐈[V] ∧ 𝐈[T]. +/4 width=1/ qed-. + +lemma tif_inv_appl: ∀V,T. 𝐈[ⓐV.T] → ∧∧ 𝐈[V] & 𝐈[T] & 𝐒[T]. +#V #T #HVT @and3_intro /3 width=1/ +generalize in match HVT; -HVT elim T -T // +* // * #U #T #_ #_ #H elim (H ?) -H /2 width=1/ +qed-. + +lemma tif_inv_cast: ∀V,T. 𝐈[ⓣV.T] → False. +/2 width=1/ qed-. + +(* Basic properties *********************************************************) + +lemma tif_atom: ∀I. 𝐈[⓪{I}]. +/2 width=4/ qed. + +lemma tif_abst: ∀V,T. 𝐈[V] → 𝐈[T] → 𝐈[ⓛV.T]. +#V #T #HV #HT #H +elim (trf_inv_abst … H) -H /2 width=1/ +qed. + +lemma tif_appl: ∀V,T. 𝐈[V] → 𝐈[T] → 𝐒[T] → 𝐈[ⓐV.T]. +#V #T #HV #HT #S #H +elim (trf_inv_appl … H) -H /2 width=1/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma new file mode 100644 index 000000000..5a7205cc0 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma @@ -0,0 +1,59 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/tpr.ma". + +(* CONTEXT-FREE NORMAL TERMS ************************************************) + +definition tnf: predicate term ≝ NF … tpr (eq …). + +interpretation + "context-free normality (term)" + 'Normal T = (tnf T). + +(* Basic inversion lemmas ***************************************************) + +lemma tnf_inv_abst: ∀V,T. 𝐍[ⓛV.T] → 𝐍[V] ∧ 𝐍[T]. +#V1 #T1 #HVT1 @conj +[ #V2 #HV2 lapply (HVT1 (ⓛV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct // +| #T2 #HT2 lapply (HVT1 (ⓛV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct // +] +qed-. + +lemma tnf_inv_appl: ∀V,T. 𝐍[ⓐV.T] → ∧∧ 𝐍[V] & 𝐍[T] & 𝐒[T]. +#V1 #T1 #HVT1 @and3_intro +[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct // +| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct // +| generalize in match HVT1; -HVT1 elim T1 -T1 * // * #W1 #U1 #_ #_ #H + [ elim (lift_total V1 0 1) #V2 #HV12 + lapply (H (ⓓW1.ⓐV2.U1) ?) -H /2 width=3/ -HV12 #H destruct + | lapply (H (ⓓV1.U1) ?) -H /2 width=1/ #H destruct +] +qed-. + +lemma tnf_inv_abbr: ∀V,T. 𝐍[ⓓV.T] → False. +#V #T #H elim (is_lift_dec T 0 1) +[ * #U #HTU + lapply (H U ?) -H /2 width=3/ #H destruct + elim (lift_inv_pair_xy_y … HTU) +| #HT + elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12 + lapply (H (ⓓV.T2) ?) -H /2 width=3/ -HT2 #H destruct /3 width=2/ +] +qed. + +lemma tnf_inv_cast: ∀V,T. 𝐍[ⓣV.T] → False. +#V #T #H lapply (H T ?) -H /2 width=1/ #H +@(discr_tpair_xy_y … H) +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf_tif.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf_tif.ma new file mode 100644 index 000000000..2888ff9d6 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf_tif.ma @@ -0,0 +1,74 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps_lift.ma". +include "Basic_2/reducibility/tif.ma". +include "Basic_2/reducibility/tnf.ma". + +(* CONTEXT-FREE NORMAL TERMS ************************************************) + +(* Main properties properties ***********************************************) + +lemma tpr_tif_eq: ∀T1,T2. T1 ➡ T2 → 𝐈[T1] → T1 = T2. +#T1 #T2 #H elim H -T1 -T2 +[ // +| * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H + [ elim (tif_inv_appl … H) -H #HV1 #HT1 #_ + >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // + | elim (tif_inv_cast … H) + ] +| #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H + elim (tif_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #H + [ -HT2 -IHV1 -IHT1 elim (tif_inv_abbr … H) + | <(tps_inv_refl_SO2 … HT2 ?) -HT2 // + elim (tif_inv_abst … H) -H #HV1 #HT1 + >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // + ] +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H + elim (tif_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #V1 #T1 #T2 #T #_ #_ #_ #H + elim (tif_inv_abbr … H) +| #V1 #T1 #T #_ #_ #H + elim (tif_inv_cast … H) +] +qed. + +theorem tif_tnf: ∀T1. 𝐈[T1] → 𝐍[T1]. +/2 width=1/ qed. + +(* Note: this property is unusual *) +lemma tnf_trf_false: ∀T1. 𝐑[T1] → 𝐍[T1] → False. +#T1 #H elim H -T1 +[ #V #T #_ #IHV #H elim (tnf_inv_abst … H) -H /2 width=1/ +| #V #T #_ #IHT #H elim (tnf_inv_abst … H) -H /2 width=1/ +| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/ +| #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/ +| #V #T #H elim (tnf_inv_abbr … H) +| #V #T #H elim (tnf_inv_cast … H) +| #V #W #T #H elim (tnf_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +] +qed. + +theorem tnf_tif: ∀T1. 𝐍[T1] → 𝐈[T1]. +/2 width=3/ qed. + +lemma tnf_abst: ∀V,T. 𝐍[V] → 𝐍[T] → 𝐍[ⓛV.T]. +/4 width=1/ qed. + +lemma tnf_appl: ∀V,T. 𝐍[V] → 𝐍[T] → 𝐒[T] → 𝐍[ⓐV.T]. +/4 width=1/ qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr.ma new file mode 100644 index 000000000..0edd15211 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr.ma @@ -0,0 +1,209 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +(* Basic_1: includes: pr0_delta1 *) +inductive tpr: relation term ≝ +| tpr_atom : ∀I. tpr (⓪{I}) (⓪{I}) +| tpr_flat : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 → + tpr (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +| tpr_beta : ∀V1,V2,W,T1,T2. + tpr V1 V2 → tpr T1 T2 → tpr (ⓐV1. ⓛW. T1) (ⓓV2. T2) +| tpr_delta: ∀I,V1,V2,T1,T2,T. + tpr V1 V2 → tpr T1 T2 → ⋆. ⓑ{I} V2 ⊢ T2 [0, 1] ▶ T → + tpr (ⓑ{I} V1. T1) (ⓑ{I} V2. T) +| tpr_theta: ∀V,V1,V2,W1,W2,T1,T2. + tpr V1 V2 → ⇧[0,1] V2 ≡ V → tpr W1 W2 → tpr T1 T2 → + tpr (ⓐV1. ⓓW1. T1) (ⓓW2. ⓐV. T2) +| tpr_zeta : ∀V,T,T1,T2. ⇧[0,1] T1 ≡ T → tpr T1 T2 → tpr (ⓓV. T) T2 +| tpr_tau : ∀V,T1,T2. tpr T1 T2 → tpr (ⓣV. T1) T2 +. + +interpretation + "context-free parallel reduction (term)" + 'PRed T1 T2 = (tpr T1 T2). + +(* Basic properties *********************************************************) + +lemma tpr_bind: ∀I,V1,V2,T1,T2. V1 ➡ V2 → T1 ➡ T2 → + ⓑ{I} V1. T1 ➡ ⓑ{I} V2. T2. +/2 width=3/ qed. + +(* Basic_1: was by definition: pr0_refl *) +lemma tpr_refl: ∀T. T ➡ T. +#T elim T -T // +#I elim I -I /2 width=1/ +qed. + +(* Basic inversion lemmas ***************************************************) + +fact tpr_inv_atom1_aux: ∀U1,U2. U1 ➡ U2 → ∀I. U1 = ⓪{I} → U2 = ⓪{I}. +#U1 #U2 * -U1 -U2 +[ // +| #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #T #_ #_ #_ #k #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct +| #V #T #T1 #T2 #_ #_ #k #H destruct +| #V #T1 #T2 #_ #k #H destruct +] +qed. + +(* Basic_1: was: pr0_gen_sort pr0_gen_lref *) +lemma tpr_inv_atom1: ∀I,U2. ⓪{I} ➡ U2 → U2 = ⓪{I}. +/2 width=3/ qed-. + +fact tpr_inv_bind1_aux: ∀U1,U2. U1 ➡ U2 → ∀I,V1,T1. U1 = ⓑ{I} V1. T1 → + (∃∃V2,T2,T. V1 ➡ V2 & T1 ➡ T2 & + ⋆. ⓑ{I} V2 ⊢ T2 [0, 1] ▶ T & + U2 = ⓑ{I} V2. T + ) ∨ + ∃∃T. ⇧[0,1] T ≡ T1 & T ➡ U2 & I = Abbr. +#U1 #U2 * -U1 -U2 +[ #J #I #V #T #H destruct +| #I1 #V1 #V2 #T1 #T2 #_ #_ #I #V #T #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #I #V #T #H destruct +| #I1 #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #I0 #V0 #T0 #H destruct /3 width=7/ +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #I0 #V0 #T0 #H destruct +| #V #T #T1 #T2 #HT1 #HT12 #I0 #V0 #T0 #H destruct /3 width=3/ +| #V #T1 #T2 #_ #I0 #V0 #T0 #H destruct +] +qed. + +lemma tpr_inv_bind1: ∀V1,T1,U2,I. ⓑ{I} V1. T1 ➡ U2 → + (∃∃V2,T2,T. V1 ➡ V2 & T1 ➡ T2 & + ⋆. ⓑ{I} V2 ⊢ T2 [0, 1] ▶ T & + U2 = ⓑ{I} V2. T + ) ∨ + ∃∃T. ⇧[0,1] T ≡ T1 & T ➡ U2 & I = Abbr. +/2 width=3/ qed-. + +(* Basic_1: was pr0_gen_abbr *) +lemma tpr_inv_abbr1: ∀V1,T1,U2. ⓓV1. T1 ➡ U2 → + (∃∃V2,T2,T. V1 ➡ V2 & T1 ➡ T2 & + ⋆. ⓓV2 ⊢ T2 [0, 1] ▶ T & + U2 = ⓓV2. T + ) ∨ + ∃∃T. ⇧[0,1] T ≡ T1 & T ➡ U2. +#V1 #T1 #U2 #H +elim (tpr_inv_bind1 … H) -H * /3 width=7/ +qed-. + +fact tpr_inv_flat1_aux: ∀U1,U2. U1 ➡ U2 → ∀I,V1,U0. U1 = ⓕ{I} V1. U0 → + ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 & + U2 = ⓕ{I} V2. T2 + | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 & + U0 = ⓛW. T1 & + U2 = ⓓV2. T2 & I = Appl + | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 & + ⇧[0,1] V2 ≡ V & + U0 = ⓓW1. T1 & + U2 = ⓓW2. ⓐV. T2 & + I = Appl + | (U0 ➡ U2 ∧ I = Cast). +#U1 #U2 * -U1 -U2 +[ #I #J #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=5/ +| #V1 #V2 #W #T1 #T2 #HV12 #HT12 #J #V #T #H destruct /3 width=8/ +| #I #V1 #V2 #T1 #T2 #T #_ #_ #_ #J #V0 #T0 #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #J #V0 #T0 #H destruct /3 width=12/ +| #V #T #T1 #T2 #_ #_ #J #V0 #T0 #H destruct +| #V #T1 #T2 #HT12 #J #V0 #T0 #H destruct /3 width=1/ +] +qed. + +lemma tpr_inv_flat1: ∀V1,U0,U2,I. ⓕ{I} V1. U0 ➡ U2 → + ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 & + U2 = ⓕ{I} V2. T2 + | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 & + U0 = ⓛW. T1 & + U2 = ⓓV2. T2 & I = Appl + | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 & + ⇧[0,1] V2 ≡ V & + U0 = ⓓW1. T1 & + U2 = ⓓW2. ⓐV. T2 & + I = Appl + | (U0 ➡ U2 ∧ I = Cast). +/2 width=3/ qed-. + +(* Basic_1: was pr0_gen_appl *) +lemma tpr_inv_appl1: ∀V1,U0,U2. ⓐV1. U0 ➡ U2 → + ∨∨ ∃∃V2,T2. V1 ➡ V2 & U0 ➡ T2 & + U2 = ⓐV2. T2 + | ∃∃V2,W,T1,T2. V1 ➡ V2 & T1 ➡ T2 & + U0 = ⓛW. T1 & + U2 = ⓓV2. T2 + | ∃∃V2,V,W1,W2,T1,T2. V1 ➡ V2 & W1 ➡ W2 & T1 ➡ T2 & + ⇧[0,1] V2 ≡ V & + U0 = ⓓW1. T1 & + U2 = ⓓW2. ⓐV. T2. +#V1 #U0 #U2 #H +elim (tpr_inv_flat1 … H) -H * /3 width=12/ #_ #H destruct +qed-. + +(* Note: the main property of simple terms *) +lemma tpr_inv_appl1_simple: ∀V1,T1,U. ⓐV1. T1 ➡ U → 𝐒[T1] → + ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & + U = ⓐV2. T2. +#V1 #T1 #U #H #HT1 +elim (tpr_inv_appl1 … H) -H * +[ /2 width=5/ +| #V2 #W #W1 #W2 #_ #_ #H #_ destruct + elim (simple_inv_bind … HT1) +| #V2 #V #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct + elim (simple_inv_bind … HT1) +] +qed-. + +(* Basic_1: was: pr0_gen_cast *) +lemma tpr_inv_cast1: ∀V1,T1,U2. ⓣV1. T1 ➡ U2 → + (∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓣV2. T2) + ∨ T1 ➡ U2. +#V1 #T1 #U2 #H +elim (tpr_inv_flat1 … H) -H * /3 width=5/ +[ #V2 #W #W1 #W2 #_ #_ #_ #_ #H destruct +| #V2 #W #W1 #W2 #T2 #U1 #_ #_ #_ #_ #_ #_ #H destruct +] +qed-. + +fact tpr_inv_lref2_aux: ∀T1,T2. T1 ➡ T2 → ∀i. T2 = #i → + ∨∨ T1 = #i + | ∃∃V,T,T0. ⇧[O,1] T0 ≡ T & T0 ➡ #i & + T1 = ⓓV. T + | ∃∃V,T. T ➡ #i & T1 = ⓣV. T. +#T1 #T2 * -T1 -T2 +[ #I #i #H destruct /2 width=1/ +| #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct +| #I #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct +| #V #T #T1 #T2 #HT1 #HT12 #i #H destruct /3 width=6/ +| #V #T1 #T2 #HT12 #i #H destruct /3 width=4/ +] +qed. + +lemma tpr_inv_lref2: ∀T1,i. T1 ➡ #i → + ∨∨ T1 = #i + | ∃∃V,T,T0. ⇧[O,1] T0 ≡ T & T0 ➡ #i & + T1 = ⓓV. T + | ∃∃V,T. T ➡ #i & T1 = ⓣV. T. +/2 width=3/ qed-. + +(* Basic_1: removed theorems 3: + pr0_subst0_back pr0_subst0_fwd pr0_subst0 + Basic_1: removed local theorems: 1: pr0_delta_tau +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_lift.ma new file mode 100644 index 000000000..be40639a8 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_lift.ma @@ -0,0 +1,121 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps_lift.ma". +include "Basic_2/reducibility/tpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +(* Relocation properties ****************************************************) + +(* Basic_1: was: pr0_lift *) +lemma tpr_lift: ∀T1,T2. T1 ➡ T2 → + ∀d,e,U1. ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → U1 ➡ U2. +#T1 #T2 #H elim H -T1 -T2 +[ * #i #d #e #U1 #HU1 #U2 #HU2 + lapply (lift_mono … HU1 … HU2) -HU1 #H destruct + [ lapply (lift_inv_sort1 … HU2) -HU2 #H destruct // + | lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct // + | lapply (lift_inv_gref1 … HU2) -HU2 #H destruct // + ] +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct + elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct /3 width=4/ +| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #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 #V3 #T3 #HV23 #HT23 #HX2 destruct /3 width=4/ +| #I #V1 #V2 #T1 #T2 #T0 #HV12 #HT12 #HT2 #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2 + elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct + elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct + elim (lift_total T2 (d + 1) e) #U2 #HTU2 + @tpr_delta + [4: @(tps_lift_le … HT2 … HTU2 HTU0 ?) /2 width=1/ |1: skip |2: /2 width=4/ |3: /2 width=4/ ] (**) (*/3. is too slow *) +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #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 + elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct + elim (lift_trans_ge … HV2 … HV3 ?) -V // /3 width=4/ +| #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #X #HX #T0 #HT20 + elim (lift_inv_bind1 … HX) -HX #V3 #T3 #_ #HT3 #HX destruct + elim (lift_trans_ge … HT1 … HT3 ?) -T // /3 width=6/ +| #V #T1 #T2 #_ #IHT12 #d #e #X #HX #T #HT2 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #_ #HT0 #HX destruct /3 width=4/ +] +qed. + +(* Basic_1: was: pr0_gen_lift *) +lemma tpr_inv_lift: ∀T1,T2. T1 ➡ T2 → + ∀d,e,U1. ⇧[d, e] U1 ≡ T1 → + ∃∃U2. ⇧[d, e] U2 ≡ T2 & U1 ➡ U2. +#T1 #T2 #H elim H -T1 -T2 +[ * #i #d #e #U1 #HU1 + [ lapply (lift_inv_sort2 … HU1) -HU1 #H destruct /2 width=3/ + | lapply (lift_inv_lref2 … HU1) -HU1 * * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … HU1) -HU1 #H destruct /2 width=3/ + ] +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct + elim (IHV12 … HV01) -V1 + elim (IHT12 … HT01) -T1 /3 width=5/ +| #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct + elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct + elim (IHV12 … HV01) -V1 + elim (IHT12 … HT01) -T1 /3 width=5/ +| #I #V1 #V2 #T1 #T2 #T0 #_ #_ #HT20 #IHV12 #IHT12 #d #e #X #HX + elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct + elim (IHV12 … HWV1) -V1 #W2 #HWV2 #HW12 + elim (IHT12 … HUT1) -T1 #U2 #HUT2 #HU12 + elim (tps_inv_lift1_le … HT20 … HUT2 ?) -T2 // [3: /2 width=5/ |2: skip ] #U0 #HU20 #HUT0 + @ex2_1_intro [2: /2 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /3 width=5/ is slow *) +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct + elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct + elim (IHV12 … HV01) -V1 #V3 #HV32 #HV03 + elim (IHW12 … HW01) -W1 #W3 #HW32 #HW03 + elim (IHT12 … HT01) -T1 #T3 #HT32 #HT03 + elim (lift_trans_le … HV32 … HV2 ?) -V2 // #V2 #HV32 #HV2 + @ex2_1_intro [2: /3 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /4 width=5/ is slow *) +| #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #X #HX + elim (lift_inv_bind2 … HX) -HX #V0 #T0 #_ #HT0 #H destruct + elim (lift_div_le … HT1 … HT0 ?) -T // #T #HT0 #HT1 + elim (IHT12 … HT1) -T1 /3 width=5/ +| #V #T1 #T2 #_ #IHT12 #d #e #X #HX + elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct + elim (IHT12 … HT01) -T1 /3 width=3/ +] +qed. + +(* Advanced inversion lemmas ************************************************) + +fact tpr_inv_abst1_aux: ∀U1,U2. U1 ➡ U2 → ∀V1,T1. U1 = ⓛV1. T1 → + ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2. +#U1 #U2 * -U1 -U2 +[ #I #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct +| #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct +| #I #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #V0 #T0 #H destruct + <(tps_inv_refl_SO2 … HT2 ? ? ?) -T /2 width=5/ +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct +| #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct +| #V #T1 #T2 #_ #V0 #T0 #H destruct +] +qed. + +(* Basic_1: was pr0_gen_abst *) +lemma tpr_inv_abst1: ∀V1,T1,U2. ⓛV1. T1 ➡ U2 → + ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛV2. T2. +/2 width=3/ qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_tpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_tpr.ma new file mode 100644 index 000000000..b1e3f3c1b --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_tpr.ma @@ -0,0 +1,287 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/tpr_tpss.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +(* Confluence lemmas ********************************************************) + +fact tpr_conf_atom_atom: ∀I. ∃∃X. ⓪{I} ➡ X & ⓪{I} ➡ X. +/2 width=3/ qed. + +fact tpr_conf_flat_flat: + ∀I,V0,V1,T0,T1,V2,T2. ( + ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → + ∃∃T0. ⓕ{I} V1. T1 ➡ T0 & ⓕ{I} V2. T2 ➡ T0. +#I #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 +elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=5/ +qed. + +fact tpr_conf_flat_beta: + ∀V0,V1,T1,V2,W0,U0,T2. ( + ∀X0:term. #[X0] < #[V0] + (#[W0] + #[U0] + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → + U0 ➡ T2 → ⓛW0. U0 ➡ T1 → + ∃∃X. ⓐV1. T1 ➡ X & ⓓV2. T2 ➡ X. +#V0 #V1 #T1 #V2 #W0 #U0 #T2 #IH #HV01 #HV02 #HT02 #H +elim (tpr_inv_abst1 … H) -H #W1 #U1 #HW01 #HU01 #H destruct +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HT02 … HU01) -HT02 -HU01 -IH /2 width=1/ /3 width=5/ +qed. + +(* basic-1: was: + pr0_cong_upsilon_refl pr0_cong_upsilon_zeta + pr0_cong_upsilon_cong pr0_cong_upsilon_delta +*) +fact tpr_conf_flat_theta: + ∀V0,V1,T1,V2,V,W0,W2,U0,U2. ( + ∀X0:term. #[X0] < #[V0] + (#[W0] + #[U0] + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → ⇧[O,1] V2 ≡ V → + W0 ➡ W2 → U0 ➡ U2 → ⓓW0. U0 ➡ T1 → + ∃∃X. ⓐV1. T1 ➡ X & ⓓW2. ⓐV. U2 ➡ X. +#V0 #V1 #T1 #V2 #V #W0 #W2 #U0 #U2 #IH #HV01 #HV02 #HV2 #HW02 #HU02 #H +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #VV #HVV1 #HVV2 +elim (lift_total VV 0 1) #VVV #HVV +lapply (tpr_lift … HVV2 … HV2 … HVV) #HVVV +elim (tpr_inv_abbr1 … H) -H * +(* case 1: delta *) +[ -HV2 -HVV2 #WW2 #UU2 #UU #HWW2 #HUU02 #HUU2 #H destruct + elim (IH … HW02 … HWW2) -HW02 -HWW2 /2 width=1/ #W #HW02 #HWW2 + elim (IH … HU02 … HUU02) -HU02 -HUU02 -IH /2 width=1/ #U #HU2 #HUUU2 + elim (tpr_tps_bind … HWW2 HUUU2 … HUU2) -UU2 #UUU #HUUU2 #HUUU1 + @ex2_1_intro + [2: @tpr_theta [6: @HVV |7: @HWW2 |8: @HUUU2 |1,2,3,4: skip | // ] + |1:skip + |3: @tpr_delta [3: @tpr_flat |1: skip ] /2 width=5/ + ] (**) (* /5 width=14/ is too slow *) +(* case 3: zeta *) +| -HW02 -HVV -HVVV #UU1 #HUU10 #HUUT1 + elim (tpr_inv_lift … HU02 … HUU10) -HU02 #UU #HUU2 #HUU1 + lapply (tw_lift … HUU10) -HUU10 #HUU10 + elim (IH … HUUT1 … HUU1) -HUUT1 -HUU1 -IH /2 width=1/ -HUU10 #U #HU2 #HUUU2 + @ex2_1_intro + [2: @tpr_flat + |1: skip + |3: @tpr_zeta [2: @lift_flat |1: skip |3: @tpr_flat ] + ] /2 width=5/ (**) (* /5 width=5/ is too slow *) +] +qed. + +fact tpr_conf_flat_cast: + ∀X2,V0,V1,T0,T1. ( + ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → T0 ➡ T1 → T0 ➡ X2 → + ∃∃X. ⓣV1. T1 ➡ X & X2 ➡ X. +#X2 #V0 #V1 #T0 #T1 #IH #_ #HT01 #HT02 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=3/ +qed. + +fact tpr_conf_beta_beta: + ∀W0:term. ∀V0,V1,T0,T1,V2,T2. ( + ∀X0:term. #[X0] < #[V0] + (#[W0] + #[T0] + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → + ∃∃X. ⓓV1. T1 ➡X & ⓓV2. T2 ➡ X. +#W0 #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ +elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ /3 width=5/ +qed. + +(* Basic_1: was: pr0_cong_delta pr0_delta_delta *) +fact tpr_conf_delta_delta: + ∀I1,V0,V1,T0,T1,TT1,V2,T2,TT2. ( + ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → + ⋆. ⓑ{I1} V1 ⊢ T1 [O, 1] ▶ TT1 → + ⋆. ⓑ{I1} V2 ⊢ T2 [O, 1] ▶ TT2 → + ∃∃X. ⓑ{I1} V1. TT1 ➡ X & ⓑ{I1} V2. TT2 ➡ X. +#I1 #V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2 +elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // #T #HT1 #HT2 +elim (tpr_tps_bind … HV1 HT1 … HTT1) -T1 #U1 #TTU1 #HTU1 +elim (tpr_tps_bind … HV2 HT2 … HTT2) -T2 #U2 #TTU2 #HTU2 +elim (tps_conf_eq … HTU1 … HTU2) -T #U #HU1 #HU2 +@ex2_1_intro [2,3: @tpr_delta |1: skip ] /width=10/ (**) (* /3 width=10/ is too slow *) +qed. + +fact tpr_conf_delta_zeta: + ∀X2,V0,V1,T0,T1,TT1,T2. ( + ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → T0 ➡ T1 → ⋆. ⓓV1 ⊢ T1 [O,1] ▶ TT1 → + T2 ➡ X2 → ⇧[O, 1] T2 ≡ T0 → + ∃∃X. ⓓV1. TT1 ➡ X & X2 ➡ X. +#X2 #V0 #V1 #T0 #T1 #TT1 #T2 #IH #_ #HT01 #HTT1 #HTX2 #HTT20 +elim (tpr_inv_lift … HT01 … HTT20) -HT01 #TT2 #HTT21 #HTT2 +lapply (tps_inv_lift1_eq … HTT1 … HTT21) -HTT1 #HTT1 destruct +lapply (tw_lift … HTT20) -HTT20 #HTT20 +elim (IH … HTX2 … HTT2) -HTX2 -HTT2 -IH // /3 width=3/ +qed. + +(* Basic_1: was: pr0_upsilon_upsilon *) +fact tpr_conf_theta_theta: + ∀VV1,V0,V1,W0,W1,T0,T1,V2,VV2,W2,T2. ( + ∀X0:term. #[X0] < #[V0] + (#[W0] + #[T0] + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → W0 ➡ W1 → W0 ➡ W2 → T0 ➡ T1 → T0 ➡ T2 → + ⇧[O, 1] V1 ≡ VV1 → ⇧[O, 1] V2 ≡ VV2 → + ∃∃X. ⓓW1. ⓐVV1. T1 ➡ X & ⓓW2. ⓐVV2. T2 ➡ X. +#VV1 #V0 #V1 #W0 #W1 #T0 #T1 #V2 #VV2 #W2 #T2 #IH #HV01 #HV02 #HW01 #HW02 #HT01 #HT02 #HVV1 #HVV2 +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HW01 … HW02) -HW01 -HW02 /2 width=1/ #W #HW1 #HW2 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ #T #HT1 #HT2 +elim (lift_total V 0 1) #VV #HVV +lapply (tpr_lift … HV1 … HVV1 … HVV) -V1 #HVV1 +lapply (tpr_lift … HV2 … HVV2 … HVV) -V2 -HVV #HVV2 +@ex2_1_intro [2,3: @tpr_bind |1:skip ] /2 width=5/ (**) (* /4 width=5/ is too slow *) +qed. + +fact tpr_conf_zeta_zeta: + ∀V0:term. ∀X2,TT0,T0,T1,T2. ( + ∀X0:term. #[X0] < #[V0] + #[TT0] + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + T0 ➡ T1 → T2 ➡ X2 → + ⇧[O, 1] T0 ≡ TT0 → ⇧[O, 1] T2 ≡ TT0 → + ∃∃X. T1 ➡ X & X2 ➡ X. +#V0 #X2 #TT0 #T0 #T1 #T2 #IH #HT01 #HTX2 #HTT0 #HTT20 +lapply (lift_inj … HTT0 … HTT20) -HTT0 #H destruct +lapply (tw_lift … HTT20) -HTT20 #HTT20 +elim (IH … HT01 … HTX2) -HT01 -HTX2 -IH // /2 width=3/ +qed. + +fact tpr_conf_tau_tau: + ∀V0,T0:term. ∀X2,T1. ( + ∀X0:term. #[X0] < #[V0] + #[T0] + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + T0 ➡ T1 → T0 ➡ X2 → + ∃∃X. T1 ➡ X & X2 ➡ X. +#V0 #T0 #X2 #T1 #IH #HT01 #HT02 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /2 width=3/ +qed. + +(* Confluence ***************************************************************) + +fact tpr_conf_aux: + ∀Y0:term. ( + ∀X0:term. #[X0] < #[Y0] → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + ∀X0,X1,X2. X0 ➡ X1 → X0 ➡ X2 → X0 = Y0 → + ∃∃X. X1 ➡ X & X2 ➡ X. +#Y0 #IH #X0 #X1 #X2 * -X0 -X1 +[ #I1 #H1 #H2 destruct + lapply (tpr_inv_atom1 … H1) -H1 +(* case 1: atom, atom *) + #H1 destruct // +| #I #V0 #V1 #T0 #T1 #HV01 #HT01 #H1 #H2 destruct + elim (tpr_inv_flat1 … H1) -H1 * +(* case 2: flat, flat *) + [ #V2 #T2 #HV02 #HT02 #H destruct + /3 width=7 by tpr_conf_flat_flat/ (**) (* /3 width=7/ is too slow *) +(* case 3: flat, beta *) + | #V2 #W #U0 #T2 #HV02 #HT02 #H1 #H2 #H3 destruct + /3 width=8 by tpr_conf_flat_beta/ (**) (* /3 width=8/ is too slow *) +(* case 4: flat, theta *) + | #V2 #V #W0 #W2 #U0 #U2 #HV02 #HW02 #HT02 #HV2 #H1 #H2 #H3 destruct + /3 width=11 by tpr_conf_flat_theta/ (**) (* /3 width=11/ is too slow *) +(* case 5: flat, tau *) + | #HT02 #H destruct + /3 width=6 by tpr_conf_flat_cast/ (**) (* /3 width=6/ is too slow *) + ] +| #V0 #V1 #W0 #T0 #T1 #HV01 #HT01 #H1 #H2 destruct + elim (tpr_inv_appl1 … H1) -H1 * +(* case 6: beta, flat (repeated) *) + [ #V2 #T2 #HV02 #HT02 #H destruct + @ex2_1_comm /3 width=8 by tpr_conf_flat_beta/ +(* case 7: beta, beta *) + | #V2 #WW0 #TT0 #T2 #HV02 #HT02 #H1 #H2 destruct + /3 width=8 by tpr_conf_beta_beta/ (**) (* /3 width=8/ is too slow *) +(* case 8, beta, theta (excluded) *) + | #V2 #VV2 #WW0 #W2 #TT0 #T2 #_ #_ #_ #_ #H destruct + ] +| #I1 #V0 #V1 #T0 #T1 #TT1 #HV01 #HT01 #HTT1 #H1 #H2 destruct + elim (tpr_inv_bind1 … H1) -H1 * +(* case 9: delta, delta *) + [ #V2 #T2 #TT2 #HV02 #HT02 #HTT2 #H destruct + /3 width=11 by tpr_conf_delta_delta/ (**) (* /3 width=11/ is too slow *) +(* case 10: delta, zata *) + | #T2 #HT20 #HTX2 #H destruct + /3 width=10 by tpr_conf_delta_zeta/ (**) (* /3 width=10/ is too slow *) + ] +| #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #HV01 #HVV1 #HW01 #HT01 #H1 #H2 destruct + elim (tpr_inv_appl1 … H1) -H1 * +(* case 11: theta, flat (repeated) *) + [ #V2 #T2 #HV02 #HT02 #H destruct + @ex2_1_comm /3 width=11 by tpr_conf_flat_theta/ +(* case 12: theta, beta (repeated) *) + | #V2 #WW0 #TT0 #T2 #_ #_ #H destruct +(* case 13: theta, theta *) + | #V2 #VV2 #WW0 #W2 #TT0 #T2 #V02 #HW02 #HT02 #HVV2 #H1 #H2 destruct + /3 width=14 by tpr_conf_theta_theta/ (**) (* /3 width=14/ is too slow *) + ] +| #V0 #TT0 #T0 #T1 #HTT0 #HT01 #H1 #H2 destruct + elim (tpr_inv_abbr1 … H1) -H1 * +(* case 14: zeta, delta (repeated) *) + [ #V2 #T2 #TT2 #HV02 #HT02 #HTT2 #H destruct + @ex2_1_comm /3 width=10 by tpr_conf_delta_zeta/ +(* case 15: zeta, zeta *) + | #T2 #HTT20 #HTX2 + /3 width=9 by tpr_conf_zeta_zeta/ (**) (* /3 width=9/ is too slow *) + ] +| #V0 #T0 #T1 #HT01 #H1 #H2 destruct + elim (tpr_inv_cast1 … H1) -H1 +(* case 16: tau, flat (repeated) *) + [ * #V2 #T2 #HV02 #HT02 #H destruct + @ex2_1_comm /3 width=6 by tpr_conf_flat_cast/ +(* case 17: tau, tau *) + | #HT02 + /3 width=5 by tpr_conf_tau_tau/ + ] +] +qed. + +(* Basic_1: was: pr0_confluence *) +theorem tpr_conf: ∀T0:term. ∀T1,T2. T0 ➡ T1 → T0 ➡ T2 → + ∃∃T. T1 ➡ T & T2 ➡ T. +#T @(tw_wf_ind … T) -T /3 width=6 by tpr_conf_aux/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_tpss.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_tpss.ma new file mode 100644 index 000000000..cf52701b5 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr_tpss.ma @@ -0,0 +1,94 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ltpss_ltpss.ma". +include "Basic_2/reducibility/ltpr_ldrop.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +(* Unfold properties ********************************************************) + +(* Basic_1: was: pr0_subst1 *) +lemma tpr_tps_ltpr: ∀T1,T2. T1 ➡ T2 → + ∀L1,d,e,U1. L1 ⊢ T1 [d, e] ▶ U1 → + ∀L2. L1 ➡ L2 → + ∃∃U2. U1 ➡ U2 & L2 ⊢ T2 [d, e] ▶* U2. +#T1 #T2 #H elim H -T1 -T2 +[ #I #L1 #d #e #X #H + elim (tps_inv_atom1 … H) -H + [ #H destruct /2 width=3/ + | * #K1 #V1 #i #Hdi #Hide #HLK1 #HVU1 #H #L2 #HL12 destruct + elim (ltpr_ldrop_conf … HLK1 … HL12) -L1 #X #HLK2 #H + elim (ltpr_inv_pair1 … H) -H #K2 #V2 #_ #HV12 #H destruct + elim (lift_total V2 0 (i+1)) #U2 #HVU2 + lapply (tpr_lift … HV12 … HVU1 … HVU2) -V1 #HU12 + @ex2_1_intro [2: @HU12 | skip | /3 width=4/ ] (**) (* /4 width=6/ is too slow *) + ] +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_flat1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct + elim (IHV12 … HVW1 … HL12) -V1 + elim (IHT12 … HTU1 … HL12) -T1 -HL12 /3 width=5/ +| #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct + elim (tps_inv_bind1 … HY) -HY #WW #TT1 #_ #HTT1 #H destruct + elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 + elim (IHT12 … HTT1 (L2. ⓛWW) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 + lapply (tpss_lsubs_conf … HTT2 (L2. ⓓVV2) ?) -HTT2 /3 width=5/ +| #I #V1 #V2 #T1 #T2 #U2 #HV12 #_ #HTU2 #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct + elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 + elim (IHT12 … HTT1 (L2. ⓑ{I} VV2) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 + elim (tpss_strip_neq … HTT2 … HTU2 ?) -T2 /2 width=1/ #T2 #HTT2 #HUT2 + lapply (tps_lsubs_conf … HTT2 (L2. ⓑ{I} V2) ?) -HTT2 /2 width=1/ #HTT2 + elim (ltpss_tps_conf … HTT2 (L2. ⓑ{I} VV2) (d + 1) e ?) -HTT2 /2 width=1/ #W2 #HTTW2 #HTW2 + lapply (tpss_lsubs_conf … HTTW2 (⋆. ⓑ{I} VV2) ?) -HTTW2 /2 width=1/ #HTTW2 + lapply (tpss_tps … HTTW2) -HTTW2 #HTTW2 + lapply (tpss_lsubs_conf … HTW2 (L2. ⓑ{I} VV2) ?) -HTW2 /2 width=1/ #HTW2 + lapply (tpss_trans_eq … HUT2 … HTW2) -T2 /3 width=5/ +| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct + elim (tps_inv_bind1 … HY) -HY #WW1 #TT1 #HWW1 #HTT1 #H destruct + elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 + elim (IHW12 … HWW1 … HL12) -W1 #WW2 #HWW12 #HWW2 + elim (IHT12 … HTT1 (L2. ⓓWW2) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 + elim (lift_total VV2 0 1) #VV #H2VV + lapply (tpss_lift_ge … HVV2 (L2. ⓓWW2) … HV2 … H2VV) -V2 /2 width=1/ #HVV + @ex2_1_intro [2: @tpr_theta |1: skip |3: @tpss_bind [2: @tpss_flat ] ] /width=11/ (**) (* /4 width=11/ is too slow *) +| #V1 #TT1 #T1 #T2 #HTT1 #_ #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_bind1 … H) -H #V2 #TT2 #HV12 #HTT12 #H destruct + elim (tps_inv_lift1_ge … HTT12 L1 … HTT1 ?) -TT1 /2 width=1/ #T2 #HT12 #HTT2 + elim (IHT12 … HT12 … HL12) -T1 -HL12 (aaa_inv_sort … H) -H // +| #I1 #L #K1 #V1 #B #i #HLK1 #_ #IHA1 #A2 #H + elim (aaa_inv_lref … H) -H #I2 #K2 #V2 #HLK2 #HA2 + lapply (ldrop_mono … HLK1 … HLK2) -L #H destruct /2 width=1/ +| #L #V #T #B1 #A1 #_ #_ #_ #IHA1 #A2 #H + elim (aaa_inv_abbr … H) -H /2 width=1/ +| #L #V1 #T1 #B1 #A1 #_ #_ #IHB1 #IHA1 #X #H + elim (aaa_inv_abst … H) -H #B2 #A2 #HB2 #HA2 #H destruct /3 width=1/ +| #L #V1 #T1 #B1 #A1 #_ #_ #_ #IHA1 #A2 #H + elim (aaa_inv_appl … H) -H #B2 #_ #HA2 + lapply (IHA1 … HA2) -L #H destruct // +| #L #V #T #A1 #_ #_ #_ #IHA1 #A2 #H + elim (aaa_inv_cast … H) -H /2 width=1/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/aaa_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/static/aaa_lift.ma new file mode 100644 index 000000000..edf08147e --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/aaa_lift.ma @@ -0,0 +1,47 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ldrop_ldrop.ma". +include "Basic_2/static/aaa.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties concerning basic relocation ***********************************) + +lemma aaa_lift: ∀L1,T1,A. L1 ⊢ T1 ÷ A → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → + ∀T2. ⇧[d, e] T1 ≡ T2 → L2 ⊢ T2 ÷ A. +#L1 #T1 #A #H elim H -L1 -T1 -A +[ #L1 #k #L2 #d #e #_ #T2 #H + >(lift_inv_sort1 … H) -H // +| #I #L1 #K1 #V1 #B #i #HLK1 #_ #IHB #L2 #d #e #HL21 #T2 #H + elim (lift_inv_lref1 … H) -H * #Hid #H destruct + [ 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 (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ + ] +| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H + elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /4 width=4/ +| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H + elim (lift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /4 width=4/ +| #L1 #V1 #T1 #B #A #_ #_ #IHB #IHA #L2 #d #e #HL21 #X #H + elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /3 width=4/ +| #L1 #V1 #T1 #A #_ #_ #IH1 #IH2 #L2 #d #e #HL21 #X #H + elim (lift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + /3 width=4/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/aaa_lifts.ma b/matita/matita/contribs/lambda_delta/basic_2/static/aaa_lifts.ma new file mode 100644 index 000000000..156e75c9d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/aaa_lifts.ma @@ -0,0 +1,30 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ldrops.ma". +include "Basic_2/static/aaa_lift.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties concerning generic relocation *********************************) + +lemma aaa_lifts: ∀L1,L2,T2,A,des. ⇩*[des] L2 ≡ L1 → ∀T1. ⇧*[des] T1 ≡ T2 → + L1 ⊢ T1 ÷ A → L2 ⊢ T2 ÷ A. +#L1 #L2 #T2 #A #des #H elim H -L1 -L2 -des +[ #L #T1 #H #HT1 + <(lifts_inv_nil … H) -H // +| #L1 #L #L2 #des #d #e #_ #HL2 #IHL1 #T1 #H #HT1 + elim (lifts_inv_cons … H) -H /3 width=9/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/lsuba.ma b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba.ma new file mode 100644 index 000000000..5d087c93c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba.ma @@ -0,0 +1,92 @@ +(**************************************************************************) +(* ___ *) +(* ||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.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) + +inductive lsuba: relation lenv ≝ +| lsuba_atom: lsuba (⋆) (⋆) +| lsuba_pair: ∀I,L1,L2,V. lsuba L1 L2 → lsuba (L1. ⓑ{I} V) (L2. ⓑ{I} V) +| lsuba_abbr: ∀L1,L2,V,W,A. L1 ⊢ V ÷ A → L2 ⊢ W ÷ A → + lsuba L1 L2 → lsuba (L1. ⓓV) (L2. ⓛW) +. + +interpretation + "local environment refinement (atomic arity assigment)" + 'CrSubEqA L1 L2 = (lsuba L1 L2). + +(* Basic inversion lemmas ***************************************************) + +fact lsuba_inv_atom1_aux: ∀L1,L2. L1 ÷⊑ L2 → L1 = ⋆ → L2 = ⋆. +#L1 #L2 * -L1 -L2 +[ // +| #I #L1 #L2 #V #_ #H destruct +| #L1 #L2 #V #W #A #_ #_ #_ #H destruct +] +qed. + +lemma lsuba_inv_atom1: ∀L2. ⋆ ÷⊑ L2 → L2 = ⋆. +/2 width=3/ qed-. + +fact lsuba_inv_pair1_aux: ∀L1,L2. L1 ÷⊑ L2 → ∀I,K1,V. L1 = K1. ⓑ{I} V → + (∃∃K2. K1 ÷⊑ K2 & L2 = K2. ⓑ{I} V) ∨ + ∃∃K2,W,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & + L2 = K2. ⓛW & I = Abbr. +#L1 #L2 * -L1 -L2 +[ #I #K1 #V #H destruct +| #J #L1 #L2 #V #HL12 #I #K1 #W #H destruct /3 width=3/ +| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K1 #V #H destruct /3 width=7/ +] +qed. + +lemma lsuba_inv_pair1: ∀I,K1,L2,V. K1. ⓑ{I} V ÷⊑ L2 → + (∃∃K2. K1 ÷⊑ K2 & L2 = K2. ⓑ{I} V) ∨ + ∃∃K2,W,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & + L2 = K2. ⓛW & I = Abbr. +/2 width=3/ qed-. + +fact lsuba_inv_atom2_aux: ∀L1,L2. L1 ÷⊑ L2 → L2 = ⋆ → L1 = ⋆. +#L1 #L2 * -L1 -L2 +[ // +| #I #L1 #L2 #V #_ #H destruct +| #L1 #L2 #V #W #A #_ #_ #_ #H destruct +] +qed. + +lemma lsubc_inv_atom2: ∀L1. L1 ÷⊑ ⋆ → L1 = ⋆. +/2 width=3/ qed-. + +fact lsuba_inv_pair2_aux: ∀L1,L2. L1 ÷⊑ L2 → ∀I,K2,W. L2 = K2. ⓑ{I} W → + (∃∃K1. K1 ÷⊑ K2 & L1 = K1. ⓑ{I} W) ∨ + ∃∃K1,V,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & + L1 = K1. ⓓV & I = Abst. +#L1 #L2 * -L1 -L2 +[ #I #K2 #W #H destruct +| #J #L1 #L2 #V #HL12 #I #K2 #W #H destruct /3 width=3/ +| #L1 #L2 #V1 #W2 #A #HV1 #HW2 #HL12 #I #K2 #W #H destruct /3 width=7/ +] +qed. + +lemma lsuba_inv_pair2: ∀I,L1,K2,W. L1 ÷⊑ K2. ⓑ{I} W → + (∃∃K1. K1 ÷⊑ K2 & L1 = K1. ⓑ{I} W) ∨ + ∃∃K1,V,A. K1 ⊢ V ÷ A & K2 ⊢ W ÷ A & K1 ÷⊑ K2 & + L1 = K1. ⓓV & I = Abst. +/2 width=3/ qed-. + +(* Basic properties *********************************************************) + +lemma lsuba_refl: ∀L. L ÷⊑ L. +#L elim L -L // /2 width=1/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_aaa.ma b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_aaa.ma new file mode 100644 index 000000000..a6a9c3fae --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_aaa.ma @@ -0,0 +1,54 @@ +(**************************************************************************) +(* ___ *) +(* ||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_aaa.ma". +include "Basic_2/static/lsuba_ldrop.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) + +(* Properties concerning atomic arity assignment ****************************) + +lemma lsuba_aaa_conf: ∀L1,V,A. L1 ⊢ V ÷ A → ∀L2. L1 ÷⊑ L2 → L2 ⊢ V ÷ A. +#L1 #V #A #H elim H -L1 -V -A +[ // +| #I #L1 #K1 #V1 #B #i #HLK1 #HV1B #IHV1 #L2 #HL12 + elim (lsuba_ldrop_O1_conf … HL12 … HLK1) -L1 #X #H #HLK2 + elim (lsuba_inv_pair1 … H) -H * #K2 + [ #HK12 #H destruct /3 width=5/ + | #V2 #A1 #HV1A1 #HV2 #_ #H1 #H2 destruct + >(aaa_mono … HV1B … HV1A1) -B -HV1A1 /2 width=5/ + ] +| /4 width=2/ +| /4 width=1/ +| /3 width=3/ +| /3 width=1/ +] +qed. + +lemma lsuba_aaa_trans: ∀L2,V,A. L2 ⊢ V ÷ A → ∀L1. L1 ÷⊑ L2 → L1 ⊢ V ÷ A. +#L2 #V #A #H elim H -L2 -V -A +[ // +| #I #L2 #K2 #V2 #B #i #HLK2 #HV2B #IHV2 #L1 #HL12 + elim (lsuba_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 + elim (lsuba_inv_pair2 … H) -H * #K1 + [ #HK12 #H destruct /3 width=5/ + | #V1 #A1 #HV1 #HV2A1 #_ #H1 #H2 destruct + >(aaa_mono … HV2B … HV2A1) -B -HV2A1 /2 width=5/ + ] +| /4 width=2/ +| /4 width=1/ +| /3 width=3/ +| /3 width=1/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_ldrop.ma new file mode 100644 index 000000000..522252ae1 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_ldrop.ma @@ -0,0 +1,63 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lsuba.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) + +(* Properties concerning basic local environment slicing ********************) + +(* Note: the constant 0 cannot be generalized *) + +lemma lsuba_ldrop_O1_conf: ∀L1,L2. L1 ÷⊑ L2 → ∀K1,e. ⇩[0, e] L1 ≡ K1 → + ∃∃K2. K1 ÷⊑ K2 & ⇩[0, e] L2 ≡ K2. +#L1 #L2 #H elim H -L1 -L2 +[ /2 width=3/ +| #I #L1 #L2 #V #_ #IHL12 #K1 #e #H + elim (ldrop_inv_O1 … H) -H * #He #HLK1 + [ destruct + elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H + <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + | elim (IHL12 … HLK1) -L1 /3 width=3/ + ] +| #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #K1 #e #H + elim (ldrop_inv_O1 … H) -H * #He #HLK1 + [ destruct + elim (IHL12 L1 0 ?) -IHL12 // #X #HL12 #H + <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + | elim (IHL12 … HLK1) -L1 /3 width=3/ + ] +] +qed-. + +lemma lsuba_ldrop_O1_trans: ∀L1,L2. L1 ÷⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → + ∃∃K1. K1 ÷⊑ K2 & ⇩[0, e] L1 ≡ K1. +#L1 #L2 #H elim H -L1 -L2 +[ /2 width=3/ +| #I #L1 #L2 #V #_ #IHL12 #K2 #e #H + elim (ldrop_inv_O1 … H) -H * #He #HLK2 + [ destruct + elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H + <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + | elim (IHL12 … HLK2) -L2 /3 width=3/ + ] +| #L1 #L2 #V #W #A #HV #HW #_ #IHL12 #K2 #e #H + elim (ldrop_inv_O1 … H) -H * #He #HLK2 + [ destruct + elim (IHL12 L2 0 ?) -IHL12 // #X #HL12 #H + <(ldrop_inv_refl … H) in HL12; -H /3 width=3/ + | elim (IHL12 … HLK2) -L2 /3 width=3/ + ] +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_lsuba.ma b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_lsuba.ma new file mode 100644 index 000000000..24da36ccd --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/lsuba_lsuba.ma @@ -0,0 +1,35 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lsuba_aaa.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) + +(* Main properties **********************************************************) + +theorem lsuba_trans: ∀L1,L. L1 ÷⊑ L → ∀L2. L ÷⊑ L2 → L1 ÷⊑ L2. +#L1 #L #H elim H -L1 -L +[ #X #H >(lsuba_inv_atom1 … H) -H // +| #I #L1 #L #V #HL1 #IHL1 #X #H + elim (lsuba_inv_pair1 … H) -H * #L2 + [ #HL2 #H destruct /3 width=1/ + | #V #A #HLV #HL2V #HL2 #H1 #H2 destruct /3 width=3/ + ] +| #L1 #L #V1 #W #A1 #HV1 #HW #HL1 #IHL1 #X #H + elim (lsuba_inv_pair1 … H) -H * #L2 + [ #HL2 #H destruct /3 width=5/ + | #V #A2 #_ #_ #_ #_ #H destruct + ] +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/static/sh.ma b/matita/matita/contribs/lambda_delta/basic_2/static/sh.ma new file mode 100644 index 000000000..7edbfd2af --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/static/sh.ma @@ -0,0 +1,23 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Ground_2/arith.ma". + +(* SORT HIERARCHY ***********************************************************) + +(* sort hierarchy specifications *) +record sh: Type[0] ≝ { + next: nat → nat; (* next sort in the hierarchy *) + next_lt: ∀k. k < next k (* strict monotonicity condition *) +}. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/gdrop.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/gdrop.ma new file mode 100644 index 000000000..eb4bd4865 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/gdrop.ma @@ -0,0 +1,80 @@ +(**************************************************************************) +(* ___ *) +(* ||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/genv.ma". + +(* GLOBAL ENVIRONMENT SLICING ***********************************************) + +inductive gdrop (e:nat): relation genv ≝ +| gdrop_gt: ∀G. |G| ≤ e → gdrop e G (⋆) +| gdrop_eq: ∀G. |G| = e + 1 → gdrop e G G +| gdrop_lt: ∀I,G1,G2,V. e < |G1| → gdrop e G1 G2 → gdrop e (G1. ⓑ{I} V) G2 +. + +interpretation "global slicing" + 'RDrop e G1 G2 = (gdrop e G1 G2). + +(* basic inversion lemmas ***************************************************) + +lemma gdrop_inv_gt: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| ≤ e → G2 = ⋆. +#G1 #G2 #e * -G1 -G2 // +[ #G #H >H -H >commutative_plus #H + lapply (le_plus_to_le_r … 0 H) -H #H + lapply (le_n_O_to_eq … H) -H #H destruct +| #I #G1 #G2 #V #H1 #_ #H2 + lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 normalize in ⊢ (? % ? → ?); >commutative_plus #H + lapply (lt_plus_to_lt_l … 0 H) -H #H + elim (lt_zero_false … H) +] +qed-. + +lemma gdrop_inv_eq: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| = e + 1 → G1 = G2. +#G1 #G2 #e * -G1 -G2 // +[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H + lapply (le_plus_to_le_r … 0 H) -H #H + lapply (le_n_O_to_eq … H) -H #H destruct +| #I #G1 #G2 #V #H1 #_ normalize #H2 + <(injective_plus_l … H2) in H1; -H2 #H + elim (lt_refl_false … H) +] +qed-. + +fact gdrop_inv_lt_aux: ∀I,G,G1,G2,V,e. ⇩[e] G ≡ G2 → G = G1. ⓑ{I} V → + e < |G1| → ⇩[e] G1 ≡ G2. +#I #G #G1 #G2 #V #e * -G -G2 +[ #G #H1 #H destruct #H2 + lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H + lapply (lt_plus_to_lt_l … 0 H) -H #H + elim (lt_zero_false … H) +| #G #H1 #H2 destruct >(injective_plus_l … H1) -H1 #H + elim (lt_refl_false … H) +| #J #G #G2 #W #_ #HG2 #H destruct // +] +qed. + +lemma gdrop_inv_lt: ∀I,G1,G2,V,e. + ⇩[e] G1. ⓑ{I} V ≡ G2 → e < |G1| → ⇩[e] G1 ≡ G2. +/2 width=5/ qed-. + +(* Basic properties *********************************************************) + +lemma gdrop_total: ∀e,G1. ∃G2. ⇩[e] G1 ≡ G2. +#e #G1 elim G1 -G1 /3 width=2/ +#I #V #G1 * #G2 #HG12 +elim (lt_or_eq_or_gt e (|G1|)) #He +[ /3 width=2/ +| destruct /3 width=2/ +| @ex_intro [2: @gdrop_gt normalize /2 width=1/ | skip ] (**) (* explicit constructor *) +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/gdrop_gdrop.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/gdrop_gdrop.ma new file mode 100644 index 000000000..640c36c78 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/gdrop_gdrop.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/substitution/gdrop.ma". + +(* GLOBAL ENVIRONMENT SLICING ***********************************************) + +(* Main properties **********************************************************) + +theorem gdrop_mono: ∀G,G1,e. ⇩[e] G ≡ G1 → ∀G2. ⇩[e] G ≡ G2 → G1 = G2. +#G #G1 #e #H elim H -G -G1 +[ #G #He #G2 #H + >(gdrop_inv_gt … H He) -H -He // +| #G #He #G2 #H + >(gdrop_inv_eq … H He) -H -He // +| #I #G #G1 #V #He #_ #IHG1 #G2 #H + lapply (gdrop_inv_lt … H He) -H -He /2 width=1/ +] +qed-. + +lemma gdrop_dec: ∀G1,G2,e. Decidable (⇩[e] G1 ≡ G2). +#G1 #G2 #e +elim (gdrop_total e G1) #G #HG1 +elim (genv_eq_dec G G2) #HG2 +[ destruct /2 width=1/ +| @or_intror #HG12 + lapply (gdrop_mono … HG1 … HG12) -HG1 -HG12 /2 width=1/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma new file mode 100644 index 000000000..c00819a74 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop.ma @@ -0,0 +1,227 @@ +(**************************************************************************) +(* ___ *) +(* ||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_weight.ma". +include "Basic_2/grammar/lsubs.ma". +include "Basic_2/substitution/lift.ma". + +(* LOCAL ENVIRONMENT SLICING ************************************************) + +(* Basic_1: includes: drop_skip_bind *) +inductive ldrop: nat → nat → relation lenv ≝ +| ldrop_atom : ∀d,e. ldrop d e (⋆) (⋆) +| ldrop_pair : ∀L,I,V. ldrop 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) +| ldrop_ldrop: ∀L1,L2,I,V,e. ldrop 0 e L1 L2 → ldrop 0 (e + 1) (L1. ⓑ{I} V) L2 +| ldrop_skip : ∀L1,L2,I,V1,V2,d,e. + ldrop d e L1 L2 → ⇧[d,e] V2 ≡ V1 → + ldrop (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) +. + +interpretation "local slicing" 'RDrop d e L1 L2 = (ldrop d e L1 L2). + +(* Basic inversion lemmas ***************************************************) + +fact ldrop_inv_refl_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| // +| #L1 #L2 #I #V #e #_ #_ >commutative_plus normalize #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +(* Basic_1: was: drop_gen_refl *) +lemma ldrop_inv_refl: ∀L1,L2. ⇩[0, 0] L1 ≡ L2 → L1 = L2. +/2 width=5/ qed-. + +fact ldrop_inv_atom1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → L1 = ⋆ → + L2 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V #e #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +(* Basic_1: was: drop_gen_sort *) +lemma ldrop_inv_atom1: ∀d,e,L2. ⇩[d, e] ⋆ ≡ L2 → L2 = ⋆. +/2 width=5/ qed-. + +fact ldrop_inv_O1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → + ∀K,I,V. L1 = K. ⓑ{I} V → + (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ + (0 < e ∧ ⇩[d, e - 1] K ≡ L2). +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #K #I #V #H destruct +| #L #I #V #_ #K #J #W #HX destruct /3 width=1/ +| #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct /3 width=1/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ldrop_inv_O1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 → + (e = 0 ∧ L2 = K. ⓑ{I} V) ∨ + (0 < e ∧ ⇩[0, e - 1] K ≡ L2). +/2 width=3/ qed-. + +(* Basic_1: was: drop_gen_drop *) +lemma ldrop_inv_ldrop1: ∀e,K,I,V,L2. + ⇩[0, e] K. ⓑ{I} V ≡ L2 → 0 < e → ⇩[0, e - 1] K ≡ L2. +#e #K #I #V #L2 #H #He +elim (ldrop_inv_O1 … H) -H * // #H destruct +elim (lt_refl_false … He) +qed-. + +fact ldrop_inv_skip1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → + ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & + ⇧[d - 1, e] V2 ≡ V1 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K #V #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) +| #X #L2 #Y #Z #V2 #d #e #HL12 #HV12 #_ #I #L1 #V1 #H destruct /2 width=5/ +] +qed. + +(* Basic_1: was: drop_gen_skip_l *) +lemma ldrop_inv_skip1: ∀d,e,I,K1,V1,L2. ⇩[d, e] K1. ⓑ{I} V1 ≡ L2 → 0 < d → + ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 & + ⇧[d - 1, e] V2 ≡ V1 & + L2 = K2. ⓑ{I} V2. +/2 width=3/ qed-. + +fact ldrop_inv_skip2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d → + ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & + ⇧[d - 1, e] V2 ≡ V1 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K #V #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H) +| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct /2 width=5/ +] +qed. + +(* Basic_1: was: drop_gen_skip_r *) +lemma ldrop_inv_skip2: ∀d,e,I,L1,K2,V2. ⇩[d, e] L1 ≡ K2. ⓑ{I} V2 → 0 < d → + ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & ⇧[d - 1, e] V2 ≡ V1 & + L1 = K1. ⓑ{I} V1. +/2 width=3/ qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was by definition: drop_refl *) +lemma ldrop_refl: ∀L. ⇩[0, 0] L ≡ L. +#L elim L -L // +qed. + +lemma ldrop_ldrop_lt: ∀L1,L2,I,V,e. + ⇩[0, e - 1] L1 ≡ L2 → 0 < e → ⇩[0, e] L1. ⓑ{I} V ≡ L2. +#L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) // /2 width=1/ +qed. + +lemma ldrop_lsubs_ldrop1_abbr: ∀L1,L2,d,e. L1 [d, e] ≼ L2 → + ∀K1,V,i. ⇩[0, i] L1 ≡ K1. ⓓV → + d ≤ i → i < d + e → + ∃∃K2. K1 [0, d + e - i - 1] ≼ K2 & + ⇩[0, i] L2 ≡ K2. ⓓV. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e +[ #d #e #K1 #V #i #H + lapply (ldrop_inv_atom1 … H) -H #H destruct +| #L1 #L2 #K1 #V #i #_ #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #V #e #HL12 #IHL12 #K1 #W #i #H #_ #Hie + elim (ldrop_inv_O1 … H) -H * #Hi #HLK1 + [ -IHL12 -Hie destruct + minus_minus_comm >arith_b1 // /4 width=3/ + ] +| #L1 #L2 #I #V1 #V2 #e #_ #IHL12 #K1 #W #i #H #_ #Hie + elim (ldrop_inv_O1 … H) -H * #Hi #HLK1 + [ -IHL12 -Hie -Hi destruct + | elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hie >minus_minus_comm >arith_b1 // /3 width=3/ + ] +| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide + elim (le_inv_plus_l … Hdi) #Hdim #Hi + lapply (ldrop_inv_ldrop1 … H ?) -H // #HLK1 + elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hdi -Hide >minus_minus_comm >arith_b1 // /3 width=3/ +] +qed. + +(* Basic forvard lemmas *****************************************************) + +(* Basic_1: was: drop_S *) +lemma ldrop_fwd_ldrop2: ∀L1,I2,K2,V2,e. ⇩[O, e] L1 ≡ K2. ⓑ{I2} V2 → + ⇩[O, e + 1] L1 ≡ K2. +#L1 elim L1 -L1 +[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct +| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H + elim (ldrop_inv_O1 … H) -H * #He #H + [ -IHL1 destruct /2 width=1/ + | @ldrop_ldrop >(plus_minus_m_m e 1) // /2 width=3/ + ] +] +qed-. + +lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → #[L2] ≤ #[L1]. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize +[ /2 width=3/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 + >(tw_lift … HV21) -HV21 /2 width=1/ +] +qed-. + +lemma ldrop_fwd_ldrop2_length: ∀L1,I2,K2,V2,e. + ⇩[0, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|. +#L1 elim L1 -L1 +[ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct +| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H + elim (ldrop_inv_O1 … H) -H * #He #H + [ -IHL1 destruct // + | lapply (IHL1 … H) -IHL1 -H #HeK1 whd in ⊢ (? ? %); /2 width=1/ + ] +] +qed-. + +lemma ldrop_fwd_O1_length: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → |L2| = |L1| - e. +#L1 elim L1 -L1 +[ #L2 #e #H >(ldrop_inv_atom1 … H) -H // +| #K1 #I1 #V1 #IHL1 #L2 #e #H + elim (ldrop_inv_O1 … H) -H * #He #H + [ -IHL1 destruct // + | lapply (IHL1 … H) -IHL1 -H #H >H -H normalize + >minus_le_minus_minus_comm // + ] +] +qed-. + +(* Basic_1: removed theorems 49: + drop_skip_flat + cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf + drop_clear drop_clear_O drop_clear_S + clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r + clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle + getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans + getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt + getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev + drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge + getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O + getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le + getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_ldrop.ma new file mode 100644 index 000000000..90f724ad3 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_ldrop.ma @@ -0,0 +1,126 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lift_lift.ma". +include "Basic_2/substitution/ldrop.ma". + +(* DROPPING *****************************************************************) + +(* Main properties **********************************************************) + +(* Basic_1: was: drop_mono *) +theorem ldrop_mono: ∀d,e,L,L1. ⇩[d, e] L ≡ L1 → + ∀L2. ⇩[d, e] L ≡ L2 → L1 = L2. +#d #e #L #L1 #H elim H -d -e -L -L1 +[ #d #e #L2 #H + >(ldrop_inv_atom1 … H) -L2 // +| #K #I #V #L2 #HL12 + <(ldrop_inv_refl … HL12) -L2 // +| #L #K #I #V #e #_ #IHLK #L2 #H + lapply (ldrop_inv_ldrop1 … H ?) -H // /2 width=1/ +| #L #K1 #I #T #V1 #d #e #_ #HVT1 #IHLK1 #X #H + elim (ldrop_inv_skip1 … H ?) -H // (lift_inj … HVT1 … HVT2) -HVT1 -HVT2 + >(IHLK1 … HLK2) -IHLK1 -HLK2 // +] +qed-. + +(* Basic_1: was: drop_conf_ge *) +theorem ldrop_conf_ge: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 → + ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 + e1 ≤ e2 → + ⇩[0, e2 - e1] L1 ≡ L2. +#d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 +[ #d #e #e2 #L2 #H + >(ldrop_inv_atom1 … H) -L2 // +| // +| #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H /2 width=2/ #HL2 + minus_minus_comm /3 width=1/ +| #L #K #I #V1 #V2 #d #e #_ #_ #IHLK #e2 #L2 #H #Hdee2 + lapply (transitive_le 1 … Hdee2) // #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // -He2 #HL2 + lapply (transitive_le (1 + e) … Hdee2) // #Hee2 + @ldrop_ldrop_lt >minus_minus_comm /3 width=1/ (**) (* explicit constructor *) +] +qed. + +(* Basic_1: was: drop_conf_lt *) +theorem ldrop_conf_lt: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 → + ∀e2,K2,I,V2. ⇩[0, e2] L ≡ K2. ⓑ{I} V2 → + e2 < d1 → let d ≝ d1 - e2 - 1 in + ∃∃K1,V1. ⇩[0, e2] L1 ≡ K1. ⓑ{I} V1 & + ⇩[d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2. +#d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 +[ #d #e #e2 #K2 #I #V2 #H + lapply (ldrop_inv_atom1 … H) -H #H destruct +| #L #I #V #e2 #K2 #J #V2 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d + elim (ldrop_inv_O1 … H) -H * + [ -IHL12 -He2d #H1 #H2 destruct /2 width=5/ + | -HL12 -HV12 #He #HLK + elim (IHL12 … HLK ?) -IHL12 -HLK [ (ldrop_inv_atom1 … H) -L2 /2 width=3/ +| #K #I #V #e2 #L2 #HL2 #H + lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ +| #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H + lapply (le_n_O_to_eq … H) -H #H destruct + elim (IHL12 … HL2 ?) -IHL12 -HL2 // #L0 #H #HL0 + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d + elim (ldrop_inv_O1 … H) -H * + [ -He2d -IHL12 #H1 #H2 destruct /3 width=5/ + | -HL12 -HV12 #He2 #HL2 + elim (IHL12 … HL2 ?) -L2 [ >minus_le_minus_minus_comm // /3 width=3/ | /2 width=1/ ] + ] +] +qed. + +(* Basic_1: was: drop_trans_ge *) +theorem ldrop_trans_ge: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L → + ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 ≤ e2 → ⇩[0, e1 + e2] L1 ≡ L2. +#d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L +[ #d #e #e2 #L2 #H + >(ldrop_inv_atom1 … H) -H -L2 // +| // +| /3 width=1/ +| #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2 + lapply (lt_to_le_to_lt 0 … Hde2) // #He2 + lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HL2 + @ldrop_ldrop_lt // >le_plus_minus // @IHL12 /2 width=1/ (**) (* explicit constructor *) +] +qed. + +theorem ldrop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L. + ⇩[d1, e1] L1 ≡ L → ⇩[0, e2] L ≡ L2 → d1 ≤ e2 → + ⇩[0, e2 + e1] L1 ≡ L2. +#e1 #e1 #e2 >commutative_plus /2 width=5/ +qed. + +(* Basic_1: was: drop_conf_rev *) +axiom ldrop_div: ∀e1,L1,L. ⇩[0, e1] L1 ≡ L → ∀e2,L2. ⇩[0, e2] L2 ≡ L → + ∃∃L0. ⇩[0, e1] L0 ≡ L2 & ⇩[e1, e2] L0 ≡ L1. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/lift.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/lift.ma new file mode 100644 index 000000000..cb8aac3a0 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/lift.ma @@ -0,0 +1,376 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_weight.ma". +include "Basic_2/grammar/term_simple.ma". + +(* BASIC TERM RELOCATION ****************************************************) + +(* Basic_1: includes: + lift_sort lift_lref_lt lift_lref_ge lift_bind lift_flat +*) +inductive lift: nat → nat → relation term ≝ +| lift_sort : ∀k,d,e. lift d e (⋆k) (⋆k) +| lift_lref_lt: ∀i,d,e. i < d → lift d e (#i) (#i) +| lift_lref_ge: ∀i,d,e. d ≤ i → lift d e (#i) (#(i + e)) +| lift_gref : ∀p,d,e. lift d e (§p) (§p) +| lift_bind : ∀I,V1,V2,T1,T2,d,e. + lift d e V1 V2 → lift (d + 1) e T1 T2 → + lift d e (ⓑ{I} V1. T1) (ⓑ{I} V2. T2) +| lift_flat : ∀I,V1,V2,T1,T2,d,e. + lift d e V1 V2 → lift d e T1 T2 → + lift d e (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +. + +interpretation "relocation" 'RLift d e T1 T2 = (lift d e T1 T2). + +(* Basic inversion lemmas ***************************************************) + +fact lift_inv_refl_O2_aux: ∀d,e,T1,T2. ⇧[d, e] T1 ≡ T2 → e = 0 → T1 = T2. +#d #e #T1 #T2 #H elim H -d -e -T1 -T2 // /3 width=1/ +qed. + +lemma lift_inv_refl_O2: ∀d,T1,T2. ⇧[d, 0] T1 ≡ T2 → T1 = T2. +/2 width=4/ qed-. + +fact lift_inv_sort1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k. +#d #e #T1 #T2 * -d -e -T1 -T2 // +[ #i #d #e #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +] +qed. + +lemma lift_inv_sort1: ∀d,e,T2,k. ⇧[d,e] ⋆k ≡ T2 → T2 = ⋆k. +/2 width=5/ qed-. + +fact lift_inv_lref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T1 = #i → + (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). +#d #e #T1 #T2 * -d -e -T1 -T2 +[ #k #d #e #i #H destruct +| #j #d #e #Hj #i #Hi destruct /3 width=1/ +| #j #d #e #Hj #i #Hi destruct /3 width=1/ +| #p #d #e #i #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct +] +qed. + +lemma lift_inv_lref1: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → + (i < d ∧ T2 = #i) ∨ (d ≤ i ∧ T2 = #(i + e)). +/2 width=3/ qed-. + +lemma lift_inv_lref1_lt: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → i < d → T2 = #i. +#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // +#Hdi #_ #Hid lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd +elim (lt_refl_false … Hdd) +qed-. + +lemma lift_inv_lref1_ge: ∀d,e,T2,i. ⇧[d,e] #i ≡ T2 → d ≤ i → T2 = #(i + e). +#d #e #T2 #i #H elim (lift_inv_lref1 … H) -H * // +#Hid #_ #Hdi lapply (le_to_lt_to_lt … Hdi Hid) -Hdi -Hid #Hdd +elim (lt_refl_false … Hdd) +qed-. + +fact lift_inv_gref1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀p. T1 = §p → T2 = §p. +#d #e #T1 #T2 * -d -e -T1 -T2 // +[ #i #d #e #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +] +qed. + +lemma lift_inv_gref1: ∀d,e,T2,p. ⇧[d,e] §p ≡ T2 → T2 = §p. +/2 width=5/ qed-. + +fact lift_inv_bind1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → + ∀I,V1,U1. T1 = ⓑ{I} V1.U1 → + ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 & + T2 = ⓑ{I} V2. U2. +#d #e #T1 #T2 * -d -e -T1 -T2 +[ #k #d #e #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #p #d #e #I #V1 #U1 #H destruct +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct +] +qed. + +lemma lift_inv_bind1: ∀d,e,T2,I,V1,U1. ⇧[d,e] ⓑ{I} V1. U1 ≡ T2 → + ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d+1,e] U1 ≡ U2 & + T2 = ⓑ{I} V2. U2. +/2 width=3/ qed-. + +fact lift_inv_flat1_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → + ∀I,V1,U1. T1 = ⓕ{I} V1.U1 → + ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d,e] U1 ≡ U2 & + T2 = ⓕ{I} V2. U2. +#d #e #T1 #T2 * -d -e -T1 -T2 +[ #k #d #e #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #i #d #e #_ #I #V1 #U1 #H destruct +| #p #d #e #I #V1 #U1 #H destruct +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct +| #J #W1 #W2 #T1 #T2 #d #e #HW #HT #I #V1 #U1 #H destruct /2 width=5/ +] +qed. + +lemma lift_inv_flat1: ∀d,e,T2,I,V1,U1. ⇧[d,e] ⓕ{I} V1. U1 ≡ T2 → + ∃∃V2,U2. ⇧[d,e] V1 ≡ V2 & ⇧[d,e] U1 ≡ U2 & + T2 = ⓕ{I} V2. U2. +/2 width=3/ qed-. + +fact lift_inv_sort2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k. +#d #e #T1 #T2 * -d -e -T1 -T2 // +[ #i #d #e #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #k #H destruct +] +qed. + +(* Basic_1: was: lift_gen_sort *) +lemma lift_inv_sort2: ∀d,e,T1,k. ⇧[d,e] T1 ≡ ⋆k → T1 = ⋆k. +/2 width=5/ qed-. + +fact lift_inv_lref2_aux: ∀d,e,T1,T2. ⇧[d,e] T1 ≡ T2 → ∀i. T2 = #i → + (i < d ∧ T1 = #i) ∨ (d + e ≤ i ∧ T1 = #(i - e)). +#d #e #T1 #T2 * -d -e -T1 -T2 +[ #k #d #e #i #H destruct +| #j #d #e #Hj #i #Hi destruct /3 width=1/ +| #j #d #e #Hj #i #Hi destruct (plus_minus_m_m i e) in ⊢ (? ? ? ? %); /2 width=2/ /3 width=2/ +qed. + +lemma lift_lref_ge_minus_eq: ∀d,e,i,j. d + e ≤ i → j = i - e → ⇧[d, e] #j ≡ #i. +/2 width=1/ qed-. + +(* Basic_1: was: lift_r *) +lemma lift_refl: ∀T,d. ⇧[d, 0] T ≡ T. +#T elim T -T +[ * #i // #d elim (lt_or_ge i d) /2 width=1/ +| * /2 width=1/ +] +qed. + +lemma lift_total: ∀T1,d,e. ∃T2. ⇧[d,e] T1 ≡ T2. +#T1 elim T1 -T1 +[ * #i /2 width=2/ #d #e elim (lt_or_ge i d) /3 width=2/ +| * #I #V1 #T1 #IHV1 #IHT1 #d #e + elim (IHV1 d e) -IHV1 #V2 #HV12 + [ elim (IHT1 (d+1) e) -IHT1 /3 width=2/ + | elim (IHT1 d e) -IHT1 /3 width=2/ + ] +] +qed. + +(* Basic_1: was: lift_free (right to left) *) +lemma lift_split: ∀d1,e2,T1,T2. ⇧[d1, e2] T1 ≡ T2 → + ∀d2,e1. d1 ≤ d2 → d2 ≤ d1 + e1 → e1 ≤ e2 → + ∃∃T. ⇧[d1, e1] T1 ≡ T & ⇧[d2, e2 - e1] T ≡ T2. +#d1 #e2 #T1 #T2 #H elim H -d1 -e2 -T1 -T2 +[ /3 width=3/ +| #i #d1 #e2 #Hid1 #d2 #e1 #Hd12 #_ #_ + lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4 width=3/ +| #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12 + lapply (transitive_le … (i+e1) Hd21 ?) /2 width=1/ -Hd21 #Hd21 + >(plus_minus_m_m e2 e1 ?) // /3 width=3/ +| /3 width=3/ +| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 + elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b + elim (IHT (d2+1) … ? ? He12) /2 width=1/ /3 width=5/ +| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12 + elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b + elim (IHT d2 … ? ? He12) // /3 width=5/ +] +qed. + +(* Basic_1: was only: dnf_dec2 dnf_dec *) +lemma is_lift_dec: ∀T2,d,e. Decidable (∃T1. ⇧[d,e] T1 ≡ T2). +#T1 elim T1 -T1 +[ * [1,3: /3 width=2/ ] #i #d #e + elim (lt_dec i d) #Hid + [ /4 width=2/ + | lapply (false_lt_to_le … Hid) -Hid #Hid + elim (lt_dec i (d + e)) #Hide + [ @or_intror * #T1 #H + elim (lift_inv_lref2_be … H Hid Hide) + | lapply (false_lt_to_le … Hide) -Hide /4 width=2/ + ] + ] +| * #I #V2 #T2 #IHV2 #IHT2 #d #e + [ elim (IHV2 d e) -IHV2 + [ * #V1 #HV12 elim (IHT2 (d+1) e) -IHT2 + [ * #T1 #HT12 @or_introl /3 width=2/ + | -V1 #HT2 @or_intror * #X #H + elim (lift_inv_bind2 … H) -H /3 width=2/ + ] + | -IHT2 #HV2 @or_intror * #X #H + elim (lift_inv_bind2 … H) -H /3 width=2/ + ] + | elim (IHV2 d e) -IHV2 + [ * #V1 #HV12 elim (IHT2 d e) -IHT2 + [ * #T1 #HT12 /4 width=2/ + | -V1 #HT2 @or_intror * #X #H + elim (lift_inv_flat2 … H) -H /3 width=2/ + ] + | -IHT2 #HV2 @or_intror * #X #H + elim (lift_inv_flat2 … H) -H /3 width=2/ + ] + ] +] +qed. + +(* Basic_1: removed theorems 7: + lift_head lift_gen_head + lift_weight_map lift_weight lift_weight_add lift_weight_add_O + lift_tlt_dx +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/lift_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/lift_lift.ma new file mode 100644 index 000000000..30bf8886e --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/lift_lift.ma @@ -0,0 +1,200 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lift.ma". + +(* BASIC TERM RELOCATION ****************************************************) + +(* Main properies ***********************************************************) + +(* Basic_1: was: lift_inj *) +theorem lift_inj: ∀d,e,T1,U. ⇧[d,e] T1 ≡ U → ∀T2. ⇧[d,e] T2 ≡ U → T1 = T2. +#d #e #T1 #U #H elim H -d -e -T1 -U +[ #k #d #e #X #HX + lapply (lift_inv_sort2 … HX) -HX // +| #i #d #e #Hid #X #HX + lapply (lift_inv_lref2_lt … HX ?) -HX // +| #i #d #e #Hdi #X #HX + lapply (lift_inv_lref2_ge … HX ?) -HX // /2 width=1/ +| #p #d #e #X #HX + lapply (lift_inv_gref2 … HX) -HX // +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ +] +qed-. + +(* Basic_1: was: lift_gen_lift *) +theorem lift_div_le: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ⇧[d2 + e1, e2] T2 ≡ T → + d1 ≤ d2 → + ∃∃T0. ⇧[d1, e1] T0 ≡ T2 & ⇧[d2, e2] T0 ≡ T1. +#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T +[ #k #d1 #e1 #d2 #e2 #T2 #Hk #Hd12 + lapply (lift_inv_sort2 … Hk) -Hk #Hk destruct /3 width=3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12 + lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 + lapply (lift_inv_lref2_lt … Hi ?) -Hi /2 width=3/ /3 width=3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #Hi #Hd12 + elim (lift_inv_lref2 … Hi) -Hi * #Hid2 #H destruct + [ -Hd12 lapply (lt_plus_to_lt_l … Hid2) -Hid2 #Hid2 /3 width=3/ + | -Hid1 >plus_plus_comm_23 in Hid2; #H lapply (le_plus_to_le_r … H) -H #H + elim (le_inv_plus_l … H) -H #Hide2 #He2i + lapply (transitive_le … Hd12 Hide2) -Hd12 #Hd12 + >le_plus_minus_comm // >(plus_minus_m_m i e2) in ⊢ (? ? ? %); // -He2i + /4 width=3/ + ] +| #p #d1 #e1 #d2 #e2 #T2 #Hk #Hd12 + lapply (lift_inv_gref2 … Hk) -Hk #Hk destruct /3 width=3/ +| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 + lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct + elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1 + >plus_plus_comm_23 in HU2; #HU2 elim (IHU … HU2 ?) /2 width=1/ /3 width=5/ +| #I #W1 #W #U1 #U #d1 #e1 #_ #_ #IHW #IHU #d2 #e2 #T2 #H #Hd12 + lapply (lift_inv_flat2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct + elim (IHW … HW2 ?) // -IHW -HW2 #W0 #HW2 #HW1 + elim (IHU … HU2 ?) // /3 width=5/ +] +qed. + +(* Note: apparently this was missing in Basic_1 *) +theorem lift_div_be: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → + ∀e,e2,T2. ⇧[d1 + e, e2] T2 ≡ T → + e ≤ e1 → e1 ≤ e + e2 → + ∃∃T0. ⇧[d1, e] T0 ≡ T2 & ⇧[d1, e + e2 - e1] T0 ≡ T1. +#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T +[ #k #d1 #e1 #e #e2 #T2 #H >(lift_inv_sort2 … H) -H /2 width=3/ +| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2 + >(lift_inv_lref2_lt … H) -H [ /3 width=3/ | /2 width=3/ ] +| #i #d1 #e1 #Hid1 #e #e2 #T2 #H #He1 #He1e2 + elim (lt_or_ge (i+e1) (d1+e+e2)) #Hie1d1e2 + [ elim (lift_inv_lref2_be … H ? ?) -H // /2 width=1/ + | >(lift_inv_lref2_ge … H ?) -H // + lapply (le_plus_to_minus … Hie1d1e2) #Hd1e21i + elim (le_inv_plus_l … Hie1d1e2) -Hie1d1e2 #Hd1e12 #He2ie1 + @ex2_1_intro [2: /2 width=1/ | skip ] -Hd1e12 + @lift_lref_ge_minus_eq [ >plus_minus_commutative // | /2 width=1/ ] + ] +| #p #d1 #e1 #e #e2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3/ +| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2 + elim (lift_inv_bind2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct + elim (IHV1 … HV2 ? ?) -V // >plus_plus_comm_23 in HT2; #HT2 + elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/ +| #I #V1 #V #T1 #T #d1 #e1 #_ #_ #IHV1 #IHT1 #e #e2 #X #H #He1 #He1e2 + elim (lift_inv_flat2 … H) -H #V2 #T2 #HV2 #HT2 #H destruct + elim (IHV1 … HV2 ? ?) -V // + elim (IHT1 … HT2 ? ?) -T // -He1 -He1e2 /3 width=5/ +] +qed. + +theorem lift_mono: ∀d,e,T,U1. ⇧[d,e] T ≡ U1 → ∀U2. ⇧[d,e] T ≡ U2 → U1 = U2. +#d #e #T #U1 #H elim H -d -e -T -U1 +[ #k #d #e #X #HX + lapply (lift_inv_sort1 … HX) -HX // +| #i #d #e #Hid #X #HX + lapply (lift_inv_lref1_lt … HX ?) -HX // +| #i #d #e #Hdi #X #HX + lapply (lift_inv_lref1_ge … HX ?) -HX // +| #p #d #e #X #HX + lapply (lift_inv_gref1 … HX) -HX // +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ +| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1/ +] +qed-. + +(* Basic_1: was: lift_free (left to right) *) +theorem lift_trans_be: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ⇧[d2, e2] T ≡ T2 → + d1 ≤ d2 → d2 ≤ d1 + e1 → ⇧[d1, e1 + e2] T1 ≡ T2. +#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T +[ #k #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_ + >(lift_inv_sort1 … HT2) -HT2 // +| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #Hd12 #_ + lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 + lapply (lift_inv_lref1_lt … HT2 Hid2) /2 width=1/ +| #i #d1 #e1 #Hid1 #d2 #e2 #T2 #HT2 #_ #Hd21 + lapply (lift_inv_lref1_ge … HT2 ?) -HT2 + [ @(transitive_le … Hd21 ?) -Hd21 /2 width=1/ + | -Hd21 /2 width=1/ + ] +| #p #d1 #e1 #d2 #e2 #T2 #HT2 #_ #_ + >(lift_inv_gref1 … HT2) -HT2 // +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 + elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct + lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10 + lapply (IHT12 … HT20 ? ?) /2 width=1/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd12 #Hd21 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct + lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10 + lapply (IHT12 … HT20 ? ?) // /2 width=1/ +] +qed. + +(* Basic_1: was: lift_d (right to left) *) +theorem lift_trans_le: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ⇧[d2, e2] T ≡ T2 → d2 ≤ d1 → + ∃∃T0. ⇧[d2, e2] T1 ≡ T0 & ⇧[d1 + e2, e1] T0 ≡ T2. +#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T +[ #k #d1 #e1 #d2 #e2 #X #HX #_ + >(lift_inv_sort1 … HX) -HX /2 width=3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ + lapply (lt_to_le_to_lt … (d1+e2) Hid1 ?) // #Hie2 + elim (lift_inv_lref1 … HX) -HX * #Hid2 #HX destruct /3 width=3/ /4 width=3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hd21 + lapply (transitive_le … Hd21 Hid1) -Hd21 #Hid2 + lapply (lift_inv_lref1_ge … HX ?) -HX /2 width=3/ #HX destruct + >plus_plus_comm_23 /4 width=3/ +| #p #d1 #e1 #d2 #e2 #X #HX #_ + >(lift_inv_gref1 … HX) -HX /2 width=3/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 + elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct + elim (IHV12 … HV20 ?) -IHV12 -HV20 // + elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ /3 width=5/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hd21 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct + elim (IHV12 … HV20 ?) -IHV12 -HV20 // + elim (IHT12 … HT20 ?) -IHT12 -HT20 // /3 width=5/ +] +qed. + +(* Basic_1: was: lift_d (left to right) *) +theorem lift_trans_ge: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T → + ∀d2,e2,T2. ⇧[d2, e2] T ≡ T2 → d1 + e1 ≤ d2 → + ∃∃T0. ⇧[d2 - e1, e2] T1 ≡ T0 & ⇧[d1, e1] T0 ≡ T2. +#d1 #e1 #T1 #T #H elim H -d1 -e1 -T1 -T +[ #k #d1 #e1 #d2 #e2 #X #HX #_ + >(lift_inv_sort1 … HX) -HX /2 width=3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #Hded + lapply (lt_to_le_to_lt … (d1+e1) Hid1 ?) // #Hid1e + lapply (lt_to_le_to_lt … (d2-e1) Hid1 ?) /2 width=1/ #Hid2e + lapply (lt_to_le_to_lt … Hid1e Hded) -Hid1e -Hded #Hid2 + lapply (lift_inv_lref1_lt … HX ?) -HX // #HX destruct /3 width=3/ +| #i #d1 #e1 #Hid1 #d2 #e2 #X #HX #_ + elim (lift_inv_lref1 … HX) -HX * #Hied #HX destruct /4 width=3/ +| #p #d1 #e1 #d2 #e2 #X #HX #_ + >(lift_inv_gref1 … HX) -HX /2 width=3/ +| #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded + elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct + elim (IHV12 … HV20 ?) -IHV12 -HV20 // + elim (IHT12 … HT20 ?) -IHT12 -HT20 /2 width=1/ #T + (liftv_inv_nil1 … H) -H // +| #Ts #U1s #T #U1 #HTU1 #_ #IHTU1s #X #H destruct + elim (liftv_inv_cons1 … H) -H #U2 #U2s #HTU2 #HTU2s #H destruct + >(lift_mono … HTU1 … HTU2) -T /3 width=1/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/lift_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/lift_vector.ma new file mode 100644 index 000000000..aa0a30f95 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/lift_vector.ma @@ -0,0 +1,62 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_vector.ma". +include "Basic_2/substitution/lift.ma". + +(* BASIC TERM VECTOR RELOCATION *********************************************) + +inductive liftv (d,e:nat) : relation (list term) ≝ +| liftv_nil : liftv d e ◊ ◊ +| liftv_cons: ∀T1s,T2s,T1,T2. + ⇧[d, e] T1 ≡ T2 → liftv d e T1s T2s → + liftv d e (T1 :: T1s) (T2 :: T2s) +. + +interpretation "relocation (vector)" 'RLift d e T1s T2s = (liftv d e T1s T2s). + +(* Basic inversion lemmas ***************************************************) + +fact liftv_inv_nil1_aux: ∀T1s,T2s,d,e. ⇧[d, e] T1s ≡ T2s → T1s = ◊ → T2s = ◊. +#T1s #T2s #d #e * -T1s -T2s // +#T1s #T2s #T1 #T2 #_ #_ #H destruct +qed. + +lemma liftv_inv_nil1: ∀T2s,d,e. ⇧[d, e] ◊ ≡ T2s → T2s = ◊. +/2 width=5/ qed-. + +fact liftv_inv_cons1_aux: ∀T1s,T2s,d,e. ⇧[d, e] T1s ≡ T2s → + ∀U1,U1s. T1s = U1 :: U1s → + ∃∃U2,U2s. ⇧[d, e] U1 ≡ U2 & ⇧[d, e] U1s ≡ U2s & + T2s = U2 :: U2s. +#T1s #T2s #d #e * -T1s -T2s +[ #U1 #U1s #H destruct +| #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5/ +] +qed. + +lemma liftv_inv_cons1: ∀U1,U1s,T2s,d,e. ⇧[d, e] U1 :: U1s ≡ T2s → + ∃∃U2,U2s. ⇧[d, e] U1 ≡ U2 & ⇧[d, e] U1s ≡ U2s & + T2s = U2 :: U2s. +/2 width=3/ qed-. + +(* Basic properties *********************************************************) + +lemma liftv_total: ∀d,e. ∀T1s:list term. ∃T2s. ⇧[d, e] T1s ≡ T2s. +#d #e #T1s elim T1s -T1s +[ /2 width=2/ +| #T1 #T1s * #T2s #HT12s + elim (lift_total T1 d e) /3 width=2/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/ltps.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/ltps.ma new file mode 100644 index 000000000..4f30025d3 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ltps.ma @@ -0,0 +1,191 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps.ma". + +(* PARALLEL SUBSTITUTION ON LOCAL ENVIRONMENTS ******************************) + +(* Basic_1: includes: csubst1_bind *) +inductive ltps: nat → nat → relation lenv ≝ +| ltps_atom: ∀d,e. ltps d e (⋆) (⋆) +| ltps_pair: ∀L,I,V. ltps 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) +| ltps_tps2: ∀L1,L2,I,V1,V2,e. + ltps 0 e L1 L2 → L2 ⊢ V1 [0, e] ▶ V2 → + ltps 0 (e + 1) (L1. ⓑ{I} V1) L2. ⓑ{I} V2 +| ltps_tps1: ∀L1,L2,I,V1,V2,d,e. + ltps d e L1 L2 → L2 ⊢ V1 [d, e] ▶ V2 → + ltps (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) +. + +interpretation "parallel substritution (local environment)" + 'PSubst L1 d e L2 = (ltps d e L1 L2). + +(* Basic properties *********************************************************) + +lemma ltps_tps2_lt: ∀L1,L2,I,V1,V2,e. + L1 [0, e - 1] ▶ L2 → L2 ⊢ V1 [0, e - 1] ▶ V2 → + 0 < e → L1. ⓑ{I} V1 [0, e] ▶ L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He +>(plus_minus_m_m e 1) /2 width=1/ +qed. + +lemma ltps_tps1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 [d - 1, e] ▶ L2 → L2 ⊢ V1 [d - 1, e] ▶ V2 → + 0 < d → L1. ⓑ{I} V1 [d, e] ▶ L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd +>(plus_minus_m_m d 1) /2 width=1/ +qed. + +(* Basic_1: was by definition: csubst1_refl *) +lemma ltps_refl: ∀L,d,e. L [d, e] ▶ L. +#L elim L -L // +#L #I #V #IHL * /2 width=1/ * /2 width=1/ +qed. + +lemma ltps_weak_all: ∀L1,L2,d,e. L1 [d, e] ▶ L2 → L1 [0, |L2|] ▶ L2. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e +// /3 width=2/ /3 width=3/ +qed. + +(* Basic forward lemmas *****************************************************) + +lemma ltps_fwd_length: ∀L1,L2,d,e. L1 [d, e] ▶ L2 → |L1| = |L2|. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e +normalize // +qed-. + +(* Basic inversion lemmas ***************************************************) + +fact ltps_inv_refl_O2_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → e = 0 → L1 = L2. +#d #e #L1 #L2 #H elim H -d -e -L1 -L2 // +[ #L1 #L2 #I #V1 #V2 #e #_ #_ #_ >commutative_plus normalize #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV12 #IHL12 #He destruct + >(IHL12 ?) -IHL12 // >(tps_inv_refl_O2 … HV12) // +] +qed. + +lemma ltps_inv_refl_O2: ∀d,L1,L2. L1 [d, 0] ▶ L2 → L1 = L2. +/2 width=4/ qed-. + +fact ltps_inv_atom1_aux: ∀d,e,L1,L2. + L1 [d, e] ▶ L2 → L1 = ⋆ → L2 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma ltps_inv_atom1: ∀d,e,L2. ⋆ [d, e] ▶ L2 → L2 = ⋆. +/2 width=5/ qed-. + +fact ltps_inv_tps21_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → d = 0 → 0 < e → + ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 [0, e - 1] ▶ K2 & + K2 ⊢ V1 [0, e - 1] ▶ V2 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K1 #J #W1 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ltps_inv_tps21: ∀e,K1,I,V1,L2. K1. ⓑ{I} V1 [0, e] ▶ L2 → 0 < e → + ∃∃K2,V2. K1 [0, e - 1] ▶ K2 & K2 ⊢ V1 [0, e - 1] ▶ V2 & + L2 = K2. ⓑ{I} V2. +/2 width=5/ qed-. + +fact ltps_inv_tps11_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → 0 < d → + ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 [d - 1, e] ▶ K2 & + K2 ⊢ V1 [d - 1, e] ▶ V2 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K1 #V1 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K1 #W1 #H destruct /2 width=5/ +] +qed. + +lemma ltps_inv_tps11: ∀d,e,I,K1,V1,L2. K1. ⓑ{I} V1 [d, e] ▶ L2 → 0 < d → + ∃∃K2,V2. K1 [d - 1, e] ▶ K2 & + K2 ⊢ V1 [d - 1, e] ▶ V2 & + L2 = K2. ⓑ{I} V2. +/2 width=3/ qed-. + +fact ltps_inv_atom2_aux: ∀d,e,L1,L2. + L1 [d, e] ▶ L2 → L2 = ⋆ → L1 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma ltps_inv_atom2: ∀d,e,L1. L1 [d, e] ▶ ⋆ → L1 = ⋆. +/2 width=5/ qed-. + +fact ltps_inv_tps22_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → d = 0 → 0 < e → + ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & + K2 ⊢ V1 [0, e - 1] ▶ V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ltps_inv_tps22: ∀e,L1,K2,I,V2. L1 [0, e] ▶ K2. ⓑ{I} V2 → 0 < e → + ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & K2 ⊢ V1 [0, e - 1] ▶ V2 & + L1 = K1. ⓑ{I} V1. +/2 width=5/ qed-. + +fact ltps_inv_tps12_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → 0 < d → + ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & + K2 ⊢ V1 [d - 1, e] ▶ V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K2 #V2 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ +] +qed. + +lemma ltps_inv_tps12: ∀L1,K2,I,V2,d,e. L1 [d, e] ▶ K2. ⓑ{I} V2 → 0 < d → + ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & + K2 ⊢ V1 [d - 1, e] ▶ V2 & + L1 = K1. ⓑ{I} V1. +/2 width=3/ qed-. + +(* Basic_1: removed theorems 27: + csubst0_clear_O csubst0_drop_lt csubst0_drop_gt csubst0_drop_eq + csubst0_clear_O_back csubst0_clear_S csubst0_clear_trans + csubst0_drop_gt_back csubst0_drop_eq_back csubst0_drop_lt_back + csubst0_gen_sort csubst0_gen_head csubst0_getl_ge csubst0_getl_lt + csubst0_gen_S_bind_2 csubst0_getl_ge_back csubst0_getl_lt_back + csubst0_snd_bind csubst0_fst_bind csubst0_both_bind + csubst1_head csubst1_flat csubst1_gen_head + csubst1_getl_ge csubst1_getl_lt csubst1_getl_ge_back getl_csubst1 + +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/ltps_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/ltps_ldrop.ma new file mode 100644 index 000000000..bf4e1f711 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ltps_ldrop.ma @@ -0,0 +1,131 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ltps.ma". + +(* PARALLEL SUBSTITUTION ON LOCAL ENVIRONMENTS ******************************) + +lemma ltps_ldrop_conf_ge: ∀L0,L1,d1,e1. L0 [d1, e1] ▶ L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → + d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. +#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // +| // +| normalize #K0 #K1 #I #V0 #V1 #e1 #_ #_ #IHK01 #L2 #e2 #H #He12 + elim (le_inv_plus_l … He12) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ +| #K0 #K1 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ +] +qed. + +lemma ltps_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶ L0 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → + d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. +#L1 #L0 #d1 #e1 #H elim H -L1 -L0 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // +| // +| normalize #K1 #K0 #I #V1 #V0 #e1 #_ #_ #IHK10 #L2 #e2 #H #He12 + elim (le_inv_plus_l … He12) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK10 … HK0L2 ?) -K0 /2 width=1/ +| #K0 #K1 #I #V1 #V0 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK10 … HK0L2 ?) -IHK10 -HK0L2 /2 width=1/ +] +qed. + +lemma ltps_ldrop_conf_be: ∀L0,L1,d1,e1. L0 [d1, e1] ▶ L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → + ∃∃L. L2 [0, d1 + e1 - e2] ▶ L & ⇩[0, e2] L1 ≡ L. +#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ +| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 + lapply (le_n_O_to_eq … He2) -He2 #H destruct + lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ +| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #_ #He21 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK01 -He21 destruct plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 #He2de1 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + (ldrop_inv_atom1 … H) -H /2 width=3/ +| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 + lapply (le_n_O_to_eq … He2) -He2 #H destruct + lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ +| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #_ #He21 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK10 -He21 destruct plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 #He2de1 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + (ldrop_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #_ #L2 #e2 #H #He2 + lapply (le_n_O_to_eq … He2) -He2 #He2 destruct + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ +| #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #He2d1 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK01 -He2d1 destruct (ldrop_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #_ #L2 #e2 #H #He2 + lapply (le_n_O_to_eq … He2) -He2 #He2 destruct + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ +| #K1 #K0 #I #V1 #V0 #d1 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #He2d1 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK10 -He2d1 destruct minus_plus minus_plus >commutative_plus /2 width=1/ + | lapply (ltps_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /3 width=4/ + ] + ] +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 + elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2 + elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 /3 width=5/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 + elim (IHVW2 … HL01) -IHVW2 + elim (IHTU2 … HL01) -IHTU2 -HL01 /3 width=5/ +] +qed. + +lemma ltps_tps_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 → + ∀L1,d1,e1. L1 [d1, e1] ▶ L0 → d1 + e1 ≤ d2 → + L1 ⊢ T2 [d2, e2] ▶ U2. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ // +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 #Hde1d2 + lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2 + lapply (ltps_ldrop_trans_ge … HL10 … HLK0 ?) -L0 // /2 width=4/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 #Hde1d2 + @tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltps_tps1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *) +| /3 width=4/ +] +qed. + +(* Basic_1: was: subst1_subst1 *) +lemma ltps_tps_trans: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 → + ∀L1,d1,e1. L1 [d1, e1] ▶ L0 → + ∃∃T. L1 ⊢ T2 [d2, e2] ▶ T & + L0 ⊢ T [d1, e1] ▶ U2. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ /2 width=3/ +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 + elim (lt_or_ge i2 d1) #Hi2d1 + [ elim (ltps_ldrop_trans_le … HL10 … HLK0 ?) -HL10 /2 width=2/ #X #H #HLK1 + elim (ltps_inv_tps12 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK0) -HLK0 #H + elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 + lapply (tps_lift_ge … HV01 … H HVW1 HVW0 ?) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ + | lapply (ltps_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/ + ] + ] +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 + elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2 + elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 /3 width=5/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 + elim (IHVW2 … HL10) -IHVW2 + elim (IHTU2 … HL10) -IHTU2 -HL10 /3 width=5/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma new file mode 100644 index 000000000..11edc90b1 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma @@ -0,0 +1,239 @@ +(**************************************************************************) +(* ___ *) +(* ||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/substitution/ldrop.ma". + +(* PARALLEL SUBSTITUTION ON TERMS *******************************************) + +inductive tps: nat → nat → lenv → relation term ≝ +| tps_atom : ∀L,I,d,e. tps d e L (⓪{I}) (⓪{I}) +| tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV → ⇧[0, i + 1] V ≡ W → tps d e L (#i) W +| tps_bind : ∀L,I,V1,V2,T1,T2,d,e. + tps d e L V1 V2 → tps (d + 1) e (L. ⓑ{I} V2) T1 T2 → + tps d e L (ⓑ{I} V1. T1) (ⓑ{I} V2. T2) +| tps_flat : ∀L,I,V1,V2,T1,T2,d,e. + tps d e L V1 V2 → tps d e L T1 T2 → + tps d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +. + +interpretation "parallel substritution (term)" + 'PSubst L T1 d e T2 = (tps d e L T1 T2). + +(* Basic properties *********************************************************) + +lemma tps_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 → + ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ▶ T2. +#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e +[ // +| #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 + elim (ldrop_lsubs_ldrop1_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /2 width=4/ +| /4 width=1/ +| /3 width=1/ +] +qed. + +lemma tps_refl: ∀T,L,d,e. L ⊢ T [d, e] ▶ T. +#T elim T -T // +#I elim I -I /2 width=1/ +qed. + +(* Basic_1: was: subst1_ex *) +lemma tps_full: ∀K,V,T1,L,d. ⇩[0, d] L ≡ (K. ⓓV) → + ∃∃T2,T. L ⊢ T1 [d, 1] ▶ T2 & ⇧[d, 1] T ≡ T2. +#K #V #T1 elim T1 -T1 +[ * #i #L #d #HLK /2 width=4/ + elim (lt_or_eq_or_gt i d) #Hid /3 width=4/ + destruct + elim (lift_total V 0 (i+1)) #W #HVW + elim (lift_split … HVW i i ? ? ?) // /3 width=4/ +| * #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=8/ + | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8/ + ] +] +qed. + +lemma tps_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ▶ T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + L ⊢ T1 [d2, e2] ▶ T2. +#L #T1 #T2 #d1 #e1 #H elim H -L -T1 -T2 -d1 -e1 +[ // +| #L #K #V #W #i #d1 #e1 #Hid1 #Hide1 #HLK #HVW #d2 #e2 #Hd12 #Hde12 + lapply (transitive_le … Hd12 … Hid1) -Hd12 -Hid1 #Hid2 + lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 /2 width=4/ +| /4 width=3/ +| /4 width=1/ +] +qed. + +lemma tps_weak_top: ∀L,T1,T2,d,e. + L ⊢ T1 [d, e] ▶ T2 → L ⊢ T1 [d, |L| - d] ▶ T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ // +| #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW + lapply (ldrop_fwd_ldrop2_length … HLK) #Hi + lapply (le_to_lt_to_lt … Hdi Hi) /3 width=4/ +| normalize /2 width=1/ +| /2 width=1/ +] +qed. + +lemma tps_weak_all: ∀L,T1,T2,d,e. + L ⊢ T1 [d, e] ▶ T2 → L ⊢ T1 [0, |L|] ▶ T2. +#L #T1 #T2 #d #e #HT12 +lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 +lapply (tps_weak_top … HT12) // +qed. + +lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶ T2 → ∀i. d ≤ i → i ≤ d + e → + ∃∃T. L ⊢ T1 [d, i - d] ▶ T & L ⊢ T [i, d + e - i] ▶ T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ /2 width=3/ +| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde + elim (lt_or_ge i j) + [ -Hide -Hjde + >(plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=4/ + | -Hdi -Hdj #Hid + generalize in match Hide; -Hide (**) (* rewriting in the premises, rewrites in the goal too *) + >(plus_minus_m_m … Hjde) in ⊢ (% → ?); -Hjde /4 width=4/ + ] +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 + elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ + -Hdi -Hide >arith_c1x #T #HT1 #HT2 + lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // + -Hdi -Hide /3 width=5/ +] +qed. + +(* Basic inversion lemmas ***************************************************) + +fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶ T2 → ∀I. T1 = ⓪{I} → + T2 = ⓪{I} ∨ + ∃∃K,V,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV & + ⇧[O, i + 1] V ≡ T2 & + I = LRef i. +#L #T1 #T2 #d #e * -L -T1 -T2 -d -e +[ #L #I #d #e #J #H destruct /2 width=1/ +| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/ +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +] +qed. + +lemma tps_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} [d, e] ▶ T2 → + T2 = ⓪{I} ∨ + ∃∃K,V,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV & + ⇧[O, i + 1] V ≡ T2 & + I = LRef i. +/2 width=3/ qed-. + + +(* Basic_1: was: subst1_gen_sort *) +lemma tps_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k [d, e] ▶ T2 → T2 = ⋆k. +#L #T2 #k #d #e #H +elim (tps_inv_atom1 … H) -H // +* #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +(* Basic_1: was: subst1_gen_lref *) +lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i [d, e] ▶ T2 → + T2 = #i ∨ + ∃∃K,V. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV & + ⇧[O, i + 1] V ≡ T2. +#L #T2 #i #d #e #H +elim (tps_inv_atom1 … H) -H /2 width=1/ +* #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=4/ +qed-. + +lemma tps_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p [d, e] ▶ T2 → T2 = §p. +#L #T2 #p #d #e #H +elim (tps_inv_atom1 … H) -H // +* #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ▶ U2 → + ∀I,V1,T1. U1 = ⓑ{I} V1. T1 → + ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & + L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶ T2 & + U2 = ⓑ{I} V2. T2. +#d #e #L #U1 #U2 * -d -e -L -U1 -U2 +[ #L #k #d #e #I #V1 #T1 #H destruct +| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct +| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ +| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct +] +qed. + +lemma tps_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 [d, e] ▶ U2 → + ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & + L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶ T2 & + U2 = ⓑ{I} V2. T2. +/2 width=3/ qed-. + +fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ▶ U2 → + ∀I,V1,T1. U1 = ⓕ{I} V1. T1 → + ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & L ⊢ T1 [d, e] ▶ T2 & + U2 = ⓕ{I} V2. T2. +#d #e #L #U1 #U2 * -d -e -L -U1 -U2 +[ #L #k #d #e #I #V1 #T1 #H destruct +| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct +| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct +| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ +] +qed. + +lemma tps_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 [d, e] ▶ U2 → + ∃∃V2,T2. L ⊢ V1 [d, e] ▶ V2 & L ⊢ T1 [d, e] ▶ T2 & + U2 = ⓕ{I} V2. T2. +/2 width=3/ qed-. + +fact tps_inv_refl_O2_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶ T2 → e = 0 → T1 = T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ // +| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct + lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi -Hide (le_to_le_to_eq … Hdi ?) /2 width=1/ -d #K #V #HLK + lapply (ldrop_mono … HLK0 … HLK) #H destruct +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK + >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 K V) -IHT12 // /2 width=1/ +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK + >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 … HLK) -IHT12 // +] +qed. + +lemma tps_inv_refl_SO2: ∀L,T1,T2,d. L ⊢ T1 [d, 1] ▶ T2 → + ∀K,V. ⇩[0, d] L ≡ K. ⓛV → T1 = T2. +/2 width=8/ qed-. + +(* Relocation properties ****************************************************) + +(* Basic_1: was: subst1_lift_lt *) +lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶ T2 → + ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + dt + et ≤ d → + L ⊢ U1 [dt, et] ▶ U2. +#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et +[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdetd + lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid + lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct + elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=4/ +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + @tps_bind [ /2 width=6/ | @IHT12 /2 width=6/ ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ +] +qed. + +lemma tps_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶ T2 → + ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + dt ≤ d → d ≤ dt + et → + L ⊢ U1 [dt, et + e] ▶ U2. +#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et +[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdtd #_ + elim (lift_inv_lref1 … H) -H * #Hid #H destruct + [ -Hdtd + lapply (lt_to_le_to_lt … (dt+et+e) Hidet ?) // -Hidet #Hidete + elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=4/ + | -Hdti + lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti + lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ + ] +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + @tps_bind [ /2 width=6/ | @IHT12 [3,4: // | skip |5,6: /2 width=1/ | /2 width=1/ ] + ] (**) (* /3 width=6/ is too slow, simplification like tps_lift_le is too slow *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ +] +qed. + +(* Basic_1: was: subst1_lift_ge *) +lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶ T2 → + ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + d ≤ dt → + L ⊢ U1 [dt + e, et] ▶ U2. +#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et +[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt + lapply (transitive_le … Hddt … Hdti) -Hddt #Hid + lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct + lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + @tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_gen_lift_lt *) +lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶ U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. K ⊢ T1 [dt, et] ▶ T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et +[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ] +| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd + lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid + lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct + elim (ldrop_conf_lt … HLK … HLKV ?) -L // #L #U #HKL #_ #HUV + elim (lift_trans_le … HUV … HVW ?) -V // >minus_plus minus_plus plus_minus // commutative_plus >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) + ] +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet + elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct + elim (IHV12 … HLK … HWV1 ? ?) -V1 // #W2 #HW12 #HWV2 + elim (IHU12 … HTU1 ? ?) -U1 [5: @ldrop_skip // |2: skip |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1/ |4: /2 width=1/ ] (**) (* 29s without the rewrites *) + /3 width=5/ +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet + elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct + elim (IHV12 … HLK … HWV1 ? ?) -V1 // + elim (IHU12 … HLK … HTU1 ? ?) -U1 -HLK // /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_gen_lift_ge *) +lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶ U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d + e ≤ dt → + ∃∃T2. K ⊢ T1 [dt - e, et] ▶ T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et +[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ] +| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt + lapply (transitive_le … Hdedt … Hdti) #Hdei + elim (le_inv_plus_l … Hdedt) -Hdedt #_ #Hedt + elim (le_inv_plus_l … Hdei) #Hdie #Hei + lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct + lapply (ldrop_conf_ge … HLK … HLKV ?) -L // #HKV + elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW [4: // |3: /2 width=1/ |2: /3 width=1/ ] -Hdei -Hdie + #V0 #HV10 >plus_minus // plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd + elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct + elim (le_inv_plus_l … Hdetd) #_ #Hedt + elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2 + elim (IHU12 … HTU1 ?) -U1 [4: @ldrop_skip // |2: skip |3: /2 width=1/ ] + IHV12 // >IHT12 // +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct + >IHV12 // >IHT12 // +] +qed. +(* + Theorem subst0_gen_lift_rev_ge: (t1,v,u2,i,h,d:?) + (subst0 i v t1 (lift h d u2)) -> + (le (plus d h) i) -> + (EX u1 | (subst0 (minus i h) v u1 u2) & + t1 = (lift h d u1) + ). + + + Theorem subst0_gen_lift_rev_lelt: (t1,v,u2,i,h,d:?) + (subst0 i v t1 (lift h d u2)) -> + (le d i) -> (lt i (plus d h)) -> + (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). +*) +lemma tps_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶ U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ▶ T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (tps_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 +lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct +elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -U -L // (lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3/ + ] +| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (tps_lsubs_conf … HT02 (L. ⓑ{I} V1) ?) -HT02 /2 width=1/ #HT02 + elim (IHV01 … HV02) -V0 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 #T #HT1 #HT2 + lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubs_conf … HT2 (L. ⓑ{I} V) ?) -HT2 /3 width=5/ +| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -V0 + elim (IHT01 … HT02) -T0 /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_confluence_neq *) +theorem tps_conf_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 [d1, e1] ▶ T1 → + ∀L2,T2,d2,e2. L2 ⊢ T0 [d2, e2] ▶ T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. L2 ⊢ T1 [d2, e2] ▶ T & L1 ⊢ T2 [d1, e1] ▶ T. +#L1 #T0 #T1 #d1 #e1 #H elim H -L1 -T0 -T1 -d1 -e1 +[ /2 width=3/ +| #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 + elim (tps_inv_lref1 … H1) -H1 + [ #H destruct /4 width=4/ + | -HLK1 -HVT1 * #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded + [ -Hd1 -Hde2 + lapply (transitive_le … Hded Hd2) -Hded -Hd2 #H + lapply (lt_to_le_to_lt … Hde1 H) -Hde1 -H #H + elim (lt_refl_false … H) + | -Hd2 -Hde1 + lapply (transitive_le … Hded Hd1) -Hded -Hd1 #H + lapply (lt_to_le_to_lt … Hde2 H) -Hde2 -H #H + elim (lt_refl_false … H) + ] + ] +| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 #V #HV1 #HV2 + elim (IHT01 … HT02 ?) -T0 + [ -H #T #HT1 #HT2 + lapply (tps_lsubs_conf … HT1 (L2. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubs_conf … HT2 (L1. ⓑ{I} V) ?) -HT2 /2 width=1/ /3 width=5/ + | -HV1 -HV2 >plus_plus_comm_23 >plus_plus_comm_23 in ⊢ (? ? %); elim H -H #H + [ @or_introl | @or_intror ] /2 by monotonic_le_plus_l/ (**) (* /3 / is too slow *) + ] +| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 + elim (IHT01 … HT02 H) -T0 -H /3 width=5/ +] +qed. + +(* Note: the constant 1 comes from tps_subst *) +(* Basic_1: was: subst1_trans *) +theorem tps_trans_ge: ∀L,T1,T0,d,e. L ⊢ T1 [d, e] ▶ T0 → + ∀T2. L ⊢ T0 [d, 1] ▶ T2 → 1 ≤ e → + L ⊢ T1 [d, e] ▶ T2. +#L #T1 #T0 #d #e #H elim H -L -T1 -T0 -d -e +[ #L #I #d #e #T2 #H #He + elim (tps_inv_atom1 … H) -H + [ #H destruct // + | * #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct + lapply (lt_to_le_to_lt … (d + e) Hide2 ?) /2 width=4/ + ] +| #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He + lapply (tps_weak … HVT2 0 (i +1) ? ?) -HVT2 /2 width=1/ #HVT2 + <(tps_inv_lift1_eq … HVT2 … HVW) -HVT2 /2 width=4/ +| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (tps_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct + lapply (tps_lsubs_conf … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 + lapply (IHT10 … HT02 He) -T0 #HT12 + lapply (tps_lsubs_conf … HT12 (L. ⓑ{I} V2) ?) -HT12 /2 width=1/ /3 width=1/ +| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (tps_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1/ +] +qed. + +theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ▶ T0 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶ T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 [d2, e2] ▶ T & L ⊢ T [d1, e1] ▶ T2. +#L #T1 #T0 #d1 #e1 #H elim H -L -T1 -T0 -d1 -e1 +[ /2 width=3/ +| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 + lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1 + lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2 normalize /2 width=1/ -Hde2i1 #HWT2 + <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /4 width=4/ +| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (tps_lsubs_conf … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 + elim (IHV10 … HV02 ?) -IHV10 -HV02 // #V + elim (IHT10 … HT02 ?) -T0 /2 width=1/ #T #HT1 #HT2 + lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubs_conf … HT2 (L. ⓑ{I} V2) ?) -HT2 /2 width=1/ /3 width=6/ +| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV10 … HV02 ?) -V0 // + elim (IHT10 … HT02 ?) -T0 // /3 width=6/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/delift.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/delift.ma new file mode 100644 index 000000000..ec2d6c373 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/delift.ma @@ -0,0 +1,84 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/tpss.ma". + +(* DELIFT ON TERMS **********************************************************) + +definition delift: nat → nat → lenv → relation term ≝ + λd,e,L,T1,T2. ∃∃T. L ⊢ T1 [d, e] ▶* T & ⇧[d, e] T2 ≡ T. + +interpretation "delift (term)" + 'TSubst L T1 d e T2 = (delift d e L T1 T2). + +(* Basic properties *********************************************************) + +lemma delift_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≡ T2 → + ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ≡ T2. +#L1 #T1 #T2 #d #e * /3 width=3/ +qed. + +lemma delift_bind: ∀I,L,V1,V2,T1,T2,d,e. + L ⊢ V1 [d, e] ≡ V2 → L. ⓑ{I} V2 ⊢ T1 [d+1, e] ≡ T2 → + L ⊢ ⓑ{I} V1. T1 [d, e] ≡ ⓑ{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * #T #HT1 #HT2 +lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ /3 width=5/ +qed. + +lemma delift_flat: ∀I,L,V1,V2,T1,T2,d,e. + L ⊢ V1 [d, e] ≡ V2 → L ⊢ T1 [d, e] ≡ T2 → + L ⊢ ⓕ{I} V1. T1 [d, e] ≡ ⓕ{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * /3 width=5/ +qed. + +(* Basic forward lemmas *****************************************************) + +lemma delift_fwd_sort1: ∀L,U2,d,e,k. L ⊢ ⋆k [d, e] ≡ U2 → U2 = ⋆k. +#L #U2 #d #e #k * #U #HU +>(tpss_inv_sort1 … HU) -HU #HU2 +>(lift_inv_sort2 … HU2) -HU2 // +qed-. + +lemma delift_fwd_gref1: ∀L,U2,d,e,p. L ⊢ §p [d, e] ≡ U2 → U2 = §p. +#L #U #d #e #p * #U #HU +>(tpss_inv_gref1 … HU) -HU #HU2 +>(lift_inv_gref2 … HU2) -HU2 // +qed-. + +lemma delift_fwd_bind1: ∀I,L,V1,T1,U2,d,e. L ⊢ ⓑ{I} V1. T1 [d, e] ≡ U2 → + ∃∃V2,T2. L ⊢ V1 [d, e] ≡ V2 & + L. ⓑ{I} V2 ⊢ T1 [d+1, e] ≡ T2 & + U2 = ⓑ{I} V2. T2. +#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 +elim (tpss_inv_bind1 … HU) -HU #V #T #HV1 #HT1 #X destruct +elim (lift_inv_bind2 … HU2) -HU2 #V2 #T2 #HV2 #HT2 +lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ +qed-. + +lemma delift_fwd_flat1: ∀I,L,V1,T1,U2,d,e. L ⊢ ⓕ{I} V1. T1 [d, e] ≡ U2 → + ∃∃V2,T2. L ⊢ V1 [d, e] ≡ V2 & + L ⊢ T1 [d, e] ≡ T2 & + U2 = ⓕ{I} V2. T2. +#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 +elim (tpss_inv_flat1 … HU) -HU #V #T #HV1 #HT1 #X destruct +elim (lift_inv_flat2 … HU2) -HU2 /3 width=5/ +qed-. + +(* Basic Inversion lemmas ***************************************************) + +lemma delift_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ≡ T2 → T1 = T2. +#L #T1 #T2 #d * #T #HT1 +>(tpss_inv_refl_O2 … HT1) -HT1 #HT2 +>(lift_inv_refl_O2 … HT2) -HT2 // +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/delift_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/delift_lift.ma new file mode 100644 index 000000000..a02d3db2c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/delift_lift.ma @@ -0,0 +1,54 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/tpss_lift.ma". +include "Basic_2/unfold/delift.ma". + +(* DELIFT ON TERMS **********************************************************) + +(* Advanced forward lemmas **************************************************) + +lemma delift_fwd_lref1_lt: ∀L,U2,i,d,e. L ⊢ #i [d, e] ≡ U2 → i < d → U2 = #i. +#L #U2 #i #d #e * #U #HU #HU2 #Hid +elim (tpss_inv_lref1 … HU) -HU +[ #H destruct >(lift_inv_lref2_lt … HU2) // +| * #K #V1 #V2 #Hdi + lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi + elim (lt_refl_false … Hi) +] +qed-. + +lemma delift_fwd_lref1_be: ∀L,U2,d,e,i. L ⊢ #i [d, e] ≡ U2 → + d ≤ i → i < d + e → + ∃∃K,V1,V2. ⇩[0, i] L ≡ K. ⓓV1 & + K ⊢ V1 [0, d + e - i - 1] ≡ V2 & + ⇧[0, d] V2 ≡ U2. +#L #U2 #d #e #i * #U #HU #HU2 #Hdi #Hide +elim (tpss_inv_lref1 … HU) -HU +[ #H destruct elim (lift_inv_lref2_be … HU2 ? ?) // +| * #K #V1 #V #_ #_ #HLK #HV1 #HVU + elim (lift_div_be … HVU … HU2 ? ?) -U // /2 width=1/ /3 width=6/ +] +qed-. + +lemma delift_fwd_lref1_ge: ∀L,U2,i,d,e. L ⊢ #i [d, e] ≡ U2 → + d + e ≤ i → U2 = #(i - e). +#L #U2 #i #d #e * #U #HU #HU2 #Hdei +elim (tpss_inv_lref1 … HU) -HU +[ #H destruct >(lift_inv_lref2_ge … HU2) // +| * #K #V1 #V2 #_ #Hide + lapply (lt_to_le_to_lt … Hide Hdei) -Hide -Hdei #Hi + elim (lt_refl_false … Hi) +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2.ma new file mode 100644 index 000000000..0896b6ba3 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2.ma @@ -0,0 +1,73 @@ +(**************************************************************************) +(* ___ *) +(* ||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/term_vector.ma". + +(* GENERIC RELOCATION WITH PAIRS ********************************************) + +inductive at: list2 nat nat → relation nat ≝ +| at_nil: ∀i. at ⟠ i i +| at_lt : ∀des,d,e,i1,i2. i1 < d → + at des i1 i2 → at ({d, e} :: des) i1 i2 +| at_ge : ∀des,d,e,i1,i2. d ≤ i1 → + at des (i1 + e) i2 → at ({d, e} :: des) i1 i2 +. + +interpretation "application (generic relocation with pairs)" + 'RAt i1 des i2 = (at des i1 i2). + +(* Basic inversion lemmas ***************************************************) + +fact at_inv_nil_aux: ∀des,i1,i2. @[i1] des ≡ i2 → des = ⟠ → i1 = i2. +#des #i1 #i2 * -des -i1 -i2 +[ // +| #des #d #e #i1 #i2 #_ #_ #H destruct +| #des #d #e #i1 #i2 #_ #_ #H destruct +] +qed. + +lemma at_inv_nil: ∀i1,i2. @[i1] ⟠ ≡ i2 → i1 = i2. +/2 width=3/ qed-. + +fact at_inv_cons_aux: ∀des,i1,i2. @[i1] des ≡ i2 → + ∀d,e,des0. des = {d, e} :: des0 → + i1 < d ∧ @[i1] des0 ≡ i2 ∨ + d ≤ i1 ∧ @[i1 + e] des0 ≡ i2. +#des #i1 #i2 * -des -i1 -i2 +[ #i #d #e #des #H destruct +| #des1 #d1 #e1 #i1 #i2 #Hid1 #Hi12 #d2 #e2 #des2 #H destruct /3 width=1/ +| #des1 #d1 #e1 #i1 #i2 #Hdi1 #Hi12 #d2 #e2 #des2 #H destruct /3 width=1/ +] +qed. + +lemma at_inv_cons: ∀des,d,e,i1,i2. @[i1] {d, e} :: des ≡ i2 → + i1 < d ∧ @[i1] des ≡ i2 ∨ + d ≤ i1 ∧ @[i1 + e] des ≡ i2. +/2 width=3/ qed-. + +lemma at_inv_cons_lt: ∀des,d,e,i1,i2. @[i1] {d, e} :: des ≡ i2 → + i1 < d → @[i1] des ≡ i2. +#des #d #e #i1 #e2 #H +elim (at_inv_cons … H) -H * // #Hdi1 #_ #Hi1d +lapply (le_to_lt_to_lt … Hdi1 Hi1d) -Hdi1 -Hi1d #Hd +elim (lt_refl_false … Hd) +qed-. + +lemma at_inv_cons_ge: ∀des,d,e,i1,i2. @[i1] {d, e} :: des ≡ i2 → + d ≤ i1 → @[i1 + e] des ≡ i2. +#des #d #e #i1 #e2 #H +elim (at_inv_cons … H) -H * // #Hi1d #_ #Hdi1 +lapply (le_to_lt_to_lt … Hdi1 Hi1d) -Hdi1 -Hi1d #Hd +elim (lt_refl_false … Hd) +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_gr2.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_gr2.ma new file mode 100644 index 000000000..438605e85 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_gr2.ma @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/gr2.ma". + +(* GENERIC RELOCATION WITH PAIRS ********************************************) + +(* Main properties **********************************************************) + +theorem at_mono: ∀des,i,i1. @[i] des ≡ i1 → ∀i2. @[i] des ≡ i2 → i1 = i2. +#des #i #i1 #H elim H -des -i -i1 +[ #i #x #H <(at_inv_nil … H) -x // +| #des #d #e #i #i1 #Hid #_ #IHi1 #x #H + lapply (at_inv_cons_lt … H Hid) -H -Hid /2 width=1/ +| #des #d #e #i #i1 #Hdi #_ #IHi1 #x #H + lapply (at_inv_cons_ge … H Hdi) -H -Hdi /2 width=1/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_minus.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_minus.ma new file mode 100644 index 000000000..5e3144c93 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_minus.ma @@ -0,0 +1,76 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/gr2.ma". + +(* GENERIC RELOCATION WITH PAIRS ********************************************) + +inductive minuss: nat → relation (list2 nat nat) ≝ +| minuss_nil: ∀i. minuss i ⟠ ⟠ +| minuss_lt : ∀des1,des2,d,e,i. i < d → minuss i des1 des2 → + minuss i ({d, e} :: des1) ({d - i, e} :: des2) +| minuss_ge : ∀des1,des2,d,e,i. d ≤ i → minuss (e + i) des1 des2 → + minuss i ({d, e} :: des1) des2 +. + +interpretation "minus (generic relocation with pairs)" + 'RMinus des1 i des2 = (minuss i des1 des2). + +(* Basic inversion lemmas ***************************************************) + +fact minuss_inv_nil1_aux: ∀des1,des2,i. des1 ▭ i ≡ des2 → des1 = ⟠ → des2 = ⟠. +#des1 #des2 #i * -des1 -des2 -i +[ // +| #des1 #des2 #d #e #i #_ #_ #H destruct +| #des1 #des2 #d #e #i #_ #_ #H destruct +] +qed. + +lemma minuss_inv_nil1: ∀des2,i. ⟠ ▭ i ≡ des2 → des2 = ⟠. +/2 width=4/ qed-. + +fact minuss_inv_cons1_aux: ∀des1,des2,i. des1 ▭ i ≡ des2 → + ∀d,e,des. des1 = {d, e} :: des → + d ≤ i ∧ des ▭ e + i ≡ des2 ∨ + ∃∃des0. i < d & des ▭ i ≡ des0 & + des2 = {d - i, e} :: des0. +#des1 #des2 #i * -des1 -des2 -i +[ #i #d #e #des #H destruct +| #des1 #des #d1 #e1 #i1 #Hid1 #Hdes #d2 #e2 #des2 #H destruct /3 width=3/ +| #des1 #des #d1 #e1 #i1 #Hdi1 #Hdes #d2 #e2 #des2 #H destruct /3 width=1/ +] +qed. + +lemma minuss_inv_cons1: ∀des1,des2,d,e,i. {d, e} :: des1 ▭ i ≡ des2 → + d ≤ i ∧ des1 ▭ e + i ≡ des2 ∨ + ∃∃des. i < d & des1 ▭ i ≡ des & + des2 = {d - i, e} :: des. +/2 width=3/ qed-. + +lemma minuss_inv_cons1_ge: ∀des1,des2,d,e,i. {d, e} :: des1 ▭ i ≡ des2 → + d ≤ i → des1 ▭ e + i ≡ des2. +#des1 #des2 #d #e #i #H +elim (minuss_inv_cons1 … H) -H * // #des #Hid #_ #_ #Hdi +lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi +elim (lt_refl_false … Hi) +qed-. + +lemma minuss_inv_cons1_lt: ∀des1,des2,d,e,i. {d, e} :: des1 ▭ i ≡ des2 → + i < d → + ∃∃des. des1 ▭ i ≡ des & des2 = {d - i, e} :: des. +#des1 #des2 #d #e #i #H +elim (minuss_inv_cons1 … H) -H * /2 width=3/ #Hdi #_ #Hid +lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi +elim (lt_refl_false … Hi) +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_plus.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_plus.ma new file mode 100644 index 000000000..938f955b9 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/gr2_plus.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/unfold/gr2.ma". + +(* GENERIC RELOCATION WITH PAIRS ********************************************) + +let rec pluss (des:list2 nat nat) (i:nat) on des ≝ match des with +[ nil2 ⇒ ⟠ +| cons2 d e des ⇒ {d + i, e} :: pluss des i +]. + +interpretation "plus (generic relocation with pairs)" + 'plus x y = (pluss x y). + +(* Basic inversion lemmas ***************************************************) + +lemma pluss_inv_nil2: ∀i,des. des + i = ⟠ → des = ⟠. +#i * // normalize +#d #e #des #H destruct +qed. + +lemma pluss_inv_cons2: ∀i,d,e,des2,des. des + i = {d, e} :: des2 → + ∃∃des1. des1 + i = des2 & des = {d - i, e} :: des1. +#i #d #e #des2 * normalize +[ #H destruct +| #d1 #e1 #des1 #H destruct /2 width=3/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops.ma new file mode 100644 index 000000000..28b9a8c7e --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops.ma @@ -0,0 +1,90 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ldrop.ma". +include "Basic_2/unfold/gr2_minus.ma". +include "Basic_2/unfold/lifts.ma". + +(* GENERIC LOCAL ENVIRONMENT SLICING ****************************************) + +inductive ldrops: list2 nat nat → relation lenv ≝ +| ldrops_nil : ∀L. ldrops ⟠ L L +| ldrops_cons: ∀L1,L,L2,des,d,e. + ldrops des L1 L → ⇩[d,e] L ≡ L2 → ldrops ({d, e} :: des) L1 L2 +. + +interpretation "generic local environment slicing" + 'RDropStar des T1 T2 = (ldrops des T1 T2). + +(* Basic inversion lemmas ***************************************************) + +fact ldrops_inv_nil_aux: ∀L1,L2,des. ⇩*[des] L1 ≡ L2 → des = ⟠ → L1 = L2. +#L1 #L2 #des * -L1 -L2 -des // +#L1 #L #L2 #d #e #des #_ #_ #H destruct +qed. + +(* Basic_1: was: drop1_gen_pnil *) +lemma ldrops_inv_nil: ∀L1,L2. ⇩*[⟠] L1 ≡ L2 → L1 = L2. +/2 width=3/ qed-. + +fact ldrops_inv_cons_aux: ∀L1,L2,des. ⇩*[des] L1 ≡ L2 → + ∀d,e,tl. des = {d, e} :: tl → + ∃∃L. ⇩*[tl] L1 ≡ L & ⇩[d, e] L ≡ L2. +#L1 #L2 #des * -L1 -L2 -des +[ #L #d #e #tl #H destruct +| #L1 #L #L2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct + /2 width=3/ +qed. + +(* Basic_1: was: drop1_gen_pcons *) +lemma ldrops_inv_cons: ∀L1,L2,d,e,des. ⇩*[{d, e} :: des] L1 ≡ L2 → + ∃∃L. ⇩*[des] L1 ≡ L & ⇩[d, e] L ≡ L2. +/2 width=3/ qed-. + +lemma ldrops_inv_skip2: ∀I,des,i,des2. des ▭ i ≡ des2 → + ∀L1,K2,V2. ⇩*[des2] L1 ≡ K2. ⓑ{I} V2 → + ∃∃K1,V1,des1. des + 1 ▭ i + 1 ≡ des1 + 1 & + ⇩*[des1] K1 ≡ K2 & + ⇧*[des1] V2 ≡ V1 & + L1 = K1. ⓑ{I} V1. +#I #des #i #des2 #H elim H -des -i -des2 +[ #i #L1 #K2 #V2 #H + >(ldrops_inv_nil … H) -L1 /2 width=7/ +| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H + elim (ldrops_inv_cons … H) -H #L #HL1 #H + elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ #K #V >minus_plus #HK2 #HV2 #H destruct + elim (IHdes2 … HL1) -IHdes2 -HL1 #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct + @(ex4_3_intro … K1 V1 … ) // [3,4: /2 width=7/ | skip ] + normalize >plus_minus // @minuss_lt // /2 width=1/ (**) (* explicit constructors, /3 width=1/ is a bit slow *) +| #des #des2 #d #e #i #Hid #_ #IHdes2 #L1 #K2 #V2 #H + elim (IHdes2 … H) -IHdes2 -H #K1 #V1 #des1 #Hdes1 #HK1 #HV1 #X destruct + /4 width=7/ +] +qed-. + +(* Basic properties *********************************************************) + +(* Basic_1: was: drop1_skip_bind *) +lemma ldrops_skip: ∀L1,L2,des. ⇩*[des] L1 ≡ L2 → ∀V1,V2. ⇧*[des] V2 ≡ V1 → + ∀I. ⇩*[des + 1] L1. ⓑ{I} V1 ≡ L2. ⓑ{I} V2. +#L1 #L2 #des #H elim H -L1 -L2 -des +[ #L #V1 #V2 #HV12 #I + >(lifts_inv_nil … HV12) -HV12 // +| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #V1 #V2 #H #I + elim (lifts_inv_cons … H) -H /3 width=5/ +]. +qed. + +(* Basic_1: removed theorems 1: drop1_getl_trans +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops_ldrop.ma new file mode 100644 index 000000000..de8a1fef6 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops_ldrop.ma @@ -0,0 +1,35 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ldrop_ldrop.ma". +include "Basic_2/unfold/ldrops.ma". + +(* GENERIC LOCAL ENVIRONMENT SLICING ****************************************) + +(* Properties concerning basic local environment slicing ********************) + +lemma ldrops_ldrop_trans: ∀L1,L,des. ⇩*[des] L1 ≡ L → ∀L2,i. ⇩[0, i] L ≡ L2 → + ∃∃L0,des0,i0. ⇩[0, i0] L1 ≡ L0 & ⇩*[des0] L0 ≡ L2 & + @[i] des ≡ i0 & des ▭ i ≡ des0. +#L1 #L #des #H elim H -L1 -L -des +[ /2 width=7/ +| #L1 #L3 #L #des3 #d #e #_ #HL3 #IHL13 #L2 #i #HL2 + elim (lt_or_ge i d) #Hid + [ elim (ldrop_trans_le … HL3 … HL2 ?) -L /2 width=2/ #L #HL3 #HL2 + elim (IHL13 … HL3) -L3 /3 width=7/ + | lapply (ldrop_trans_ge … HL3 … HL2 ?) -L // #HL32 + elim (IHL13 … HL32) -L3 /3 width=7/ + ] +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops_ldrops.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops_ldrops.ma new file mode 100644 index 000000000..1bb40cb5d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ldrops_ldrops.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/unfold/ldrops_ldrop.ma". + +(* GENERIC LOCAL ENVIRONMENT SLICING ****************************************) + +(* Main properties **********************************************************) + +(* Basic_1: was: drop1_trans *) +theorem ldrops_trans: ∀L,L2,des2. ⇩*[des2] L ≡ L2 → ∀L1,des1. ⇩*[des1] L1 ≡ L → + ⇩*[des2 @ des1] L1 ≡ L2. +#L #L2 #des2 #H elim H -L -L2 -des2 // /3 width=3/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts.ma new file mode 100644 index 000000000..6a3c647a0 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts.ma @@ -0,0 +1,150 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lift.ma". +include "Basic_2/unfold/gr2_plus.ma". + +(* GENERIC TERM RELOCATION **************************************************) + +inductive lifts: list2 nat nat → relation term ≝ +| lifts_nil : ∀T. lifts ⟠ T T +| lifts_cons: ∀T1,T,T2,des,d,e. + ⇧[d,e] T1 ≡ T → lifts des T T2 → lifts ({d, e} :: des) T1 T2 +. + +interpretation "generic relocation (term)" + 'RLiftStar des T1 T2 = (lifts des T1 T2). + +(* Basic inversion lemmas ***************************************************) + +fact lifts_inv_nil_aux: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → des = ⟠ → T1 = T2. +#T1 #T2 #des * -T1 -T2 -des // +#T1 #T #T2 #d #e #des #_ #_ #H destruct +qed. + +lemma lifts_inv_nil: ∀T1,T2. ⇧*[⟠] T1 ≡ T2 → T1 = T2. +/2 width=3/ qed-. + +fact lifts_inv_cons_aux: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → + ∀d,e,tl. des = {d, e} :: tl → + ∃∃T. ⇧[d, e] T1 ≡ T & ⇧*[tl] T ≡ T2. +#T1 #T2 #des * -T1 -T2 -des +[ #T #d #e #tl #H destruct +| #T1 #T #T2 #des #d #e #HT1 #HT2 #hd #he #tl #H destruct + /2 width=3/ +qed. + +lemma lifts_inv_cons: ∀T1,T2,d,e,des. ⇧*[{d, e} :: des] T1 ≡ T2 → + ∃∃T. ⇧[d, e] T1 ≡ T & ⇧*[des] T ≡ T2. +/2 width=3/ qed-. + +(* Basic_1: was: lift1_sort *) +lemma lifts_inv_sort1: ∀T2,k,des. ⇧*[des] ⋆k ≡ T2 → T2 = ⋆k. +#T2 #k #des elim des -des +[ #H <(lifts_inv_nil … H) -H // +| #d #e #des #IH #H + elim (lifts_inv_cons … H) -H #X #H + >(lift_inv_sort1 … H) -H /2 width=1/ +] +qed-. + +(* Basic_1: was: lift1_lref *) +lemma lifts_inv_lref1: ∀T2,des,i1. ⇧*[des] #i1 ≡ T2 → + ∃∃i2. @[i1] des ≡ i2 & T2 = #i2. +#T2 #des elim des -des +[ #i1 #H <(lifts_inv_nil … H) -H /2 width=3/ +| #d #e #des #IH #i1 #H + elim (lifts_inv_cons … H) -H #X #H1 #H2 + elim (lift_inv_lref1 … H1) -H1 * #Hdi1 #H destruct + elim (IH … H2) -IH -H2 /3 width=3/ +] +qed-. + +lemma lifts_inv_gref1: ∀T2,p,des. ⇧*[des] §p ≡ T2 → T2 = §p. +#T2 #p #des elim des -des +[ #H <(lifts_inv_nil … H) -H // +| #d #e #des #IH #H + elim (lifts_inv_cons … H) -H #X #H + >(lift_inv_gref1 … H) -H /2 width=1/ +] +qed-. + +(* Basic_1: was: lift1_bind *) +lemma lifts_inv_bind1: ∀I,T2,des,V1,U1. ⇧*[des] ⓑ{I} V1. U1 ≡ T2 → + ∃∃V2,U2. ⇧*[des] V1 ≡ V2 & ⇧*[des + 1] U1 ≡ U2 & + T2 = ⓑ{I} V2. U2. +#I #T2 #des elim des -des +[ #V1 #U1 #H + <(lifts_inv_nil … H) -H /2 width=5/ +| #d #e #des #IHdes #V1 #U1 #H + elim (lifts_inv_cons … H) -H #X #H #HT2 + elim (lift_inv_bind1 … H) -H #V #U #HV1 #HU1 #H destruct + elim (IHdes … HT2) -IHdes -HT2 #V2 #U2 #HV2 #HU2 #H destruct + /3 width=5/ +] +qed-. + +(* Basic_1: was: lift1_flat *) +lemma lifts_inv_flat1: ∀I,T2,des,V1,U1. ⇧*[des] ⓕ{I} V1. U1 ≡ T2 → + ∃∃V2,U2. ⇧*[des] V1 ≡ V2 & ⇧*[des] U1 ≡ U2 & + T2 = ⓕ{I} V2. U2. +#I #T2 #des elim des -des +[ #V1 #U1 #H + <(lifts_inv_nil … H) -H /2 width=5/ +| #d #e #des #IHdes #V1 #U1 #H + elim (lifts_inv_cons … H) -H #X #H #HT2 + elim (lift_inv_flat1 … H) -H #V #U #HV1 #HU1 #H destruct + elim (IHdes … HT2) -IHdes -HT2 #V2 #U2 #HV2 #HU2 #H destruct + /3 width=5/ +] +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma lifts_simple_dx: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → 𝐒[T1] → 𝐒[T2]. +#T1 #T2 #des #H elim H -T1 -T2 -des // /3 width=5 by lift_simple_dx/ +qed-. + +lemma lifts_simple_sn: ∀T1,T2,des. ⇧*[des] T1 ≡ T2 → 𝐒[T2] → 𝐒[T1]. +#T1 #T2 #des #H elim H -T1 -T2 -des // /3 width=5 by lift_simple_sn/ +qed-. + +(* Basic properties *********************************************************) + +lemma lifts_bind: ∀I,T2,V1,V2,des. ⇧*[des] V1 ≡ V2 → + ∀T1. ⇧*[des + 1] T1 ≡ T2 → + ⇧*[des] ⓑ{I} V1. T1 ≡ ⓑ{I} V2. T2. +#I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des +[ #V #T1 #H >(lifts_inv_nil … H) -H // +| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H + elim (lifts_inv_cons … H) -H /3 width=3/ +] +qed. + +lemma lifts_flat: ∀I,T2,V1,V2,des. ⇧*[des] V1 ≡ V2 → + ∀T1. ⇧*[des] T1 ≡ T2 → + ⇧*[des] ⓕ{I} V1. T1 ≡ ⓕ{I} V2. T2. +#I #T2 #V1 #V2 #des #H elim H -V1 -V2 -des +[ #V #T1 #H >(lifts_inv_nil … H) -H // +| #V1 #V #V2 #des #d #e #HV1 #_ #IHV #T1 #H + elim (lifts_inv_cons … H) -H /3 width=3/ +] +qed. + +lemma lifts_total: ∀des,T1. ∃T2. ⇧*[des] T1 ≡ T2. +#des elim des -des /2 width=2/ +#d #e #des #IH #T1 +elim (lift_total T1 d e) #T #HT1 +elim (IH T) -IH /3 width=4/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_lift.ma new file mode 100644 index 000000000..3f63d2eb8 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_lift.ma @@ -0,0 +1,59 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lift_lift.ma". +include "Basic_2/unfold/gr2_minus.ma". +include "Basic_2/unfold/lifts.ma". + +(* GENERIC TERM RELOCATION **************************************************) + +(* Properties concerning basic term relocation ******************************) + +(* Basic_1: was: lift1_xhg (right to left) *) +lemma lifts_lift_trans_le: ∀T1,T,des. ⇧*[des] T1 ≡ T → ∀T2. ⇧[0, 1] T ≡ T2 → + ∃∃T0. ⇧[0, 1] T1 ≡ T0 & ⇧*[des + 1] T0 ≡ T2. +#T1 #T #des #H elim H -T1 -T -des +[ /2 width=3/ +| #T1 #T3 #T #des #d #e #HT13 #_ #IHT13 #T2 #HT2 + elim (IHT13 … HT2) -T #T #HT3 #HT2 + elim (lift_trans_le … HT13 … HT3 ?) -T3 // /3 width=5/ +] +qed-. + +(* Basic_1: was: lift1_free (right to left) *) +lemma lifts_lift_trans: ∀des,i,i0. @[i] des ≡ i0 → + ∀des0. des + 1 ▭ i + 1 ≡ des0 + 1 → + ∀T1,T0. ⇧*[des0] T1 ≡ T0 → + ∀T2. ⇧[O, i0 + 1] T0 ≡ T2 → + ∃∃T. ⇧[0, i + 1] T1 ≡ T & ⇧*[des] T ≡ T2. +#des elim des -des normalize +[ #i #x #H1 #des0 #H2 #T1 #T0 #HT10 #T2 + <(at_inv_nil … H1) -x #HT02 + lapply (minuss_inv_nil1 … H2) -H2 #H + >(pluss_inv_nil2 … H) in HT10; -des0 #H + >(lifts_inv_nil … H) -T1 /2 width=3/ +| #d #e #des #IHdes #i #i0 #H1 #des0 #H2 #T1 #T0 #HT10 #T2 #HT02 + elim (at_inv_cons … H1) -H1 * #Hid #Hi0 + [ elim (minuss_inv_cons1_lt … H2 ?) -H2 [2: /2 width=1/ ] #des1 #Hdes1 minus_plus #HT1 #HT0 + elim (IHdes … Hi0 … Hdes1 … HT0 … HT02) -IHdes -Hi0 -Hdes1 -T0 #T0 #HT0 #HT02 + elim (lift_trans_le … HT1 … HT0 ?) -T /2 width=1/ #T #HT1 commutative_plus in Hi0; #Hi0 + lapply (minuss_inv_cons1_ge … H2 ?) -H2 [ /2 width=1/ ] (liftv_inv_nil1 … H) -T1s /2 width=3/ +| #T1s #Ts #T1 #T #HT1 #_ #IHT1s #X #H + elim (liftv_inv_cons1 … H) -H #T2 #T2s #HT2 #HT2s #H destruct + elim (IHT1s … HT2s) -Ts #Ts #HT1s #HT2s + elim (lifts_lift_trans_le … HT1 … HT2) -T /3 width=5/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_lifts.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_lifts.ma new file mode 100644 index 000000000..6062f89b1 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_lifts.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/unfold/lifts_lift.ma". + +(* GENERIC RELOCATION *******************************************************) + +(* Main properties **********************************************************) + +(* Basic_1: was: lift1_lift1 (left to right) *) +theorem lifts_trans: ∀T1,T,des1. ⇧*[des1] T1 ≡ T → ∀T2:term. ∀des2. ⇧*[des2] T ≡ T2 → + ⇧*[des1 @ des2] T1 ≡ T2. +#T1 #T #des1 #H elim H -T1 -T -des1 // /3 width=3/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_vector.ma new file mode 100644 index 000000000..2bd579d01 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/lifts_vector.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/substitution/lift_vector.ma". +include "Basic_2/unfold/lifts.ma". + +(* GENERIC TERM VECTOR RELOCATION *******************************************) + +inductive liftsv (des:list2 nat nat) : relation (list term) ≝ +| liftsv_nil : liftsv des ◊ ◊ +| liftsv_cons: ∀T1s,T2s,T1,T2. + ⇧*[des] T1 ≡ T2 → liftsv des T1s T2s → + liftsv des (T1 :: T1s) (T2 :: T2s) +. + +interpretation "generic relocation (vector)" + 'RLiftStar des T1s T2s = (liftsv des T1s T2s). + +(* Basic inversion lemmas ***************************************************) + +lemma lifts_inv_applv1: ∀V1s,U1,T2,des. ⇧*[des] Ⓐ V1s. U1 ≡ T2 → + ∃∃V2s,U2. ⇧*[des] V1s ≡ V2s & ⇧*[des] U1 ≡ U2 & + T2 = Ⓐ V2s. U2. +#V1s elim V1s -V1s normalize +[ #T1 #T2 #des #HT12 + @(ex3_2_intro) [3,4: // |1,2: skip | // ] (**) (* explicit constructor *) +| #V1 #V1s #IHV1s #T1 #X #des #H + elim (lifts_inv_flat1 … H) -H #V2 #Y #HV12 #HY #H destruct + elim (IHV1s … HY) -IHV1s -HY #V2s #T2 #HV12s #HT12 #H destruct + @(ex3_2_intro) [4: // |3: /2 width=2/ |1,2: skip | // ] (**) (* explicit constructor *) +] +qed-. + +(* Basic properties *********************************************************) + +lemma lifts_applv: ∀V1s,V2s,des. ⇧*[des] V1s ≡ V2s → + ∀T1,T2. ⇧*[des] T1 ≡ T2 → + ⇧*[des] Ⓐ V1s. T1 ≡ Ⓐ V2s. T2. +#V1s #V2s #des #H elim H -V1s -V2s // /3 width=1/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma new file mode 100644 index 000000000..207bfc60f --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma @@ -0,0 +1,115 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ltps.ma". +include "Basic_2/unfold/tpss.ma". + +(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) + +definition ltpss: nat → nat → relation lenv ≝ + λd,e. TC … (ltps d e). + +interpretation "partial unfold (local environment)" + 'PSubstStar L1 d e L2 = (ltpss d e L1 L2). + +(* Basic eliminators ********************************************************) + +lemma ltpss_ind: ∀d,e,L1. ∀R:predicate lenv. R L1 → + (∀L,L2. L1 [d, e] ▶* L → L [d, e] ▶ L2 → R L → R L2) → + ∀L2. L1 [d, e] ▶* L2 → R L2. +#d #e #L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // +qed-. + +(* Basic properties *********************************************************) + +lemma ltpss_strap: ∀L1,L,L2,d,e. + L1 [d, e] ▶ L → L [d, e] ▶* L2 → L1 [d, e] ▶* L2. +/2 width=3/ qed. + +lemma ltpss_refl: ∀L,d,e. L [d, e] ▶* L. +/2 width=1/ qed. + +lemma ltpss_weak_all: ∀L1,L2,d,e. L1 [d, e] ▶* L2 → L1 [0, |L2|] ▶* L2. +#L1 #L2 #d #e #H @(ltpss_ind … H) -L2 // +#L #L2 #_ #HL2 +>(ltps_fwd_length … HL2) /3 width=5/ +qed. + +(* Basic forward lemmas *****************************************************) + +lemma ltpss_fwd_length: ∀L1,L2,d,e. L1 [d, e] ▶* L2 → |L1| = |L2|. +#L1 #L2 #d #e #H @(ltpss_ind … H) -L2 // +#L #L2 #_ #HL2 #IHL12 >IHL12 -IHL12 +/2 width=3 by ltps_fwd_length/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +lemma ltpss_inv_refl_O2: ∀d,L1,L2. L1 [d, 0] ▶* L2 → L1 = L2. +#d #L1 #L2 #H @(ltpss_ind … H) -L2 // +#L #L2 #_ #HL2 #IHL <(ltps_inv_refl_O2 … HL2) -HL2 // +qed-. + +lemma ltpss_inv_atom1: ∀d,e,L2. ⋆ [d, e] ▶* L2 → L2 = ⋆. +#d #e #L2 #H @(ltpss_ind … H) -L2 // +#L #L2 #_ #HL2 #IHL destruct +>(ltps_inv_atom1 … HL2) -HL2 // +qed-. + +fact ltpss_inv_atom2_aux: ∀d,e,L1,L2. L1 [d, e] ▶* L2 → L2 = ⋆ → L1 = ⋆. +#d #e #L1 #L2 #H @(ltpss_ind … H) -L2 // +#L2 #L #_ #HL2 #IHL2 #H destruct +lapply (ltps_inv_atom2 … HL2) -HL2 /2 width=1/ +qed. + +lemma ltpss_inv_atom2: ∀d,e,L1. L1 [d, e] ▶* ⋆ → L1 = ⋆. +/2 width=5/ qed-. +(* +fact ltps_inv_tps22_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → d = 0 → 0 < e → + ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & + K2 ⊢ V1 [0, e - 1] ▶ V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d e L1 L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H elim (plus_S_eq_O_false … H) +] +qed. + +lemma ltps_inv_tps22: ∀e,L1,K2,I,V2. L1 [0, e] ▶ K2. ⓑ{I} V2 → 0 < e → + ∃∃K1,V1. K1 [0, e - 1] ▶ K2 & K2 ⊢ V1 [0, e - 1] ▶ V2 & + L1 = K1. ⓑ{I} V1. +/2 width=5/ qed. + +fact ltps_inv_tps12_aux: ∀d,e,L1,L2. L1 [d, e] ▶ L2 → 0 < d → + ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & + K2 ⊢ V1 [d - 1, e] ▶ V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d e L1 L2 +[ #d #e #_ #I #K2 #V2 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ +] +qed. + +lemma ltps_inv_tps12: ∀L1,K2,I,V2,d,e. L1 [d, e] ▶ K2. ⓑ{I} V2 → 0 < d → + ∃∃K1,V1. K1 [d - 1, e] ▶ K2 & + K2 ⊢ V1 [d - 1, e] ▶ V2 & + L1 = K1. ⓑ{I} V1. +/2 width=1/ qed. +*) diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ldrop.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ldrop.ma new file mode 100644 index 000000000..b52cedae7 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ldrop.ma @@ -0,0 +1,74 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ltps_ldrop.ma". +include "Basic_2/unfold/ltpss.ma". + +(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) + +lemma ltpss_ldrop_conf_ge: ∀L0,L1,d1,e1. L0 [d1, e1] ▶* L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → + d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. +#L0 #L1 #d1 #e1 #H @(ltpss_ind … H) -L1 // /3 width=6/ +qed. + +lemma ltpss_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → + d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. +#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 // /3 width=6/ +qed. + +lemma ltpss_ldrop_conf_be: ∀L0,L1,d1,e1. L0 [d1, e1] ▶* L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → + ∃∃L. L2 [0, d1 + e1 - e2] ▶* L & ⇩[0, e2] L1 ≡ L. +#L0 #L1 #d1 #e1 #H @(ltpss_ind … H) -L1 +[ /2 width=3/ +| #L #L1 #_ #HL1 #IHL #L2 #e2 #HL02 #Hd1e2 #He2de1 + elim (IHL … HL02 Hd1e2 He2de1) -L0 #L0 #HL20 #HL0 + elim (ltps_ldrop_conf_be … HL1 … HL0 Hd1e2 He2de1) -L /3 width=3/ +] +qed. + +lemma ltpss_ldrop_trans_be: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → + ∃∃L. L [0, d1 + e1 - e2] ▶* L2 & ⇩[0, e2] L1 ≡ L. +#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 +[ /2 width=3/ +| #L #L0 #_ #HL0 #IHL #L2 #e2 #HL02 #Hd1e2 #He2de1 + elim (ltps_ldrop_trans_be … HL0 … HL02 Hd1e2 He2de1) -L0 #L0 #HL02 #HL0 + elim (IHL … HL0 Hd1e2 He2de1) -L /3 width=3/ +] +qed. + +lemma ltpss_ldrop_conf_le: ∀L0,L1,d1,e1. L0 [d1, e1] ▶* L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 → + ∃∃L. L2 [d1 - e2, e1] ▶* L & ⇩[0, e2] L1 ≡ L. +#L0 #L1 #d1 #e1 #H @(ltpss_ind … H) -L1 +[ /2 width=3/ +| #L #L1 #_ #HL1 #IHL #L2 #e2 #HL02 #He2d1 + elim (IHL … HL02 He2d1) -L0 #L0 #HL20 #HL0 + elim (ltps_ldrop_conf_le … HL1 … HL0 He2d1) -L /3 width=3/ +] +qed. + +lemma ltpss_ldrop_trans_le: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 → + ∃∃L. L [d1 - e2, e1] ▶* L2 & ⇩[0, e2] L1 ≡ L. +#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 +[ /2 width=3/ +| #L #L0 #_ #HL0 #IHL #L2 #e2 #HL02 #He2d1 + elim (ltps_ldrop_trans_le … HL0 … HL02 He2d1) -L0 #L0 #HL02 #HL0 + elim (IHL … HL0 He2d1) -L /3 width=3/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ltpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ltpss.ma new file mode 100644 index 000000000..7091e0803 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ltpss.ma @@ -0,0 +1,55 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/ltpss_tpss.ma". + +(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) + +(* Main properties **********************************************************) + +theorem ltpss_trans_eq: ∀L1,L,L2,d,e. + L1 [d, e] ▶* L → L [d, e] ▶* L2 → L1 [d, e] ▶* L2. +/2 width=3/ qed. + +lemma ltpss_tpss2: ∀L1,L2,I,V1,V2,e. + L1 [0, e] ▶* L2 → L2 ⊢ V1 [0, e] ▶* V2 → + L1. ⓑ{I} V1 [0, e + 1] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #e #HL12 #H @(tpss_ind … H) -V2 +[ /2 width=1/ +| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2 width=1/ +] +qed. + +lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e. + L1 [0, e - 1] ▶* L2 → L2 ⊢ V1 [0, e - 1] ▶* V2 → + 0 < e → L1. ⓑ{I} V1 [0, e] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He +>(plus_minus_m_m e 1) // /2 width=1/ +qed. + +lemma ltpss_tpss1: ∀L1,L2,I,V1,V2,d,e. + L1 [d, e] ▶* L2 → L2 ⊢ V1 [d, e] ▶* V2 → + L1. ⓑ{I} V1 [d + 1, e] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #H @(tpss_ind … H) -V2 +[ /2 width=1/ +| #V #V2 #_ #HV2 #IHV @(ltpss_trans_eq … IHV) /2 width=1/ +] +qed. + +lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 [d - 1, e] ▶* L2 → L2 ⊢ V1 [d - 1, e] ▶* V2 → + 0 < d → L1. ⓑ{I} V1 [d, e] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd +>(plus_minus_m_m d 1) // /2 width=1/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_tpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_tpss.ma new file mode 100644 index 000000000..24f1a595e --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_tpss.ma @@ -0,0 +1,169 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic_2/unfold/tpss_ltps.ma". +include "Basic_2/unfold/ltpss.ma". + +(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) + +(* Properties concerning partial unfold on terms ****************************) + +lemma ltpss_tpss_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → + ∀T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶* U2 → + d1 + e1 ≤ d2 → L1 ⊢ T2 [d2, e2] ▶* U2. +#L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 // +#L #L0 #_ #HL0 #IHL #T2 #U2 #d2 #e2 #HTU2 #Hde1d2 +lapply (ltps_tpss_trans_ge … HL0 HTU2) -HL0 -HTU2 /2 width=1/ +qed. + +lemma ltpss_tps_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 → + ∀T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 → + d1 + e1 ≤ d2 → L1 ⊢ T2 [d2, e2] ▶* U2. +#L1 #L0 #d1 #e1 #HL10 #T2 #U2 #d2 #e2 #HTU2 #Hde1d2 +@(ltpss_tpss_trans_ge … HL10 … Hde1d2) /2 width=1/ (**) (* /3 width=6/ is too slow *) +qed. + +lemma ltpss_tpss_trans_eq: ∀L0,L1,d,e. L0 [d, e] ▶* L1 → + ∀T2,U2. L1 ⊢ T2 [d, e] ▶* U2 → L0 ⊢ T2 [d, e] ▶* U2. +#L0 #L1 #d #e #H @(ltpss_ind … H) -L1 +[ /2 width=1/ +| #L #L1 #_ #HL1 #IHL #T2 #U2 #HTU2 + lapply (ltps_tpss_trans_eq … HL1 HTU2) -HL1 -HTU2 /2 width=1/ +] +qed. + +lemma ltpss_tps_trans_eq: ∀L0,L1,d,e. L0 [d, e] ▶* L1 → + ∀T2,U2. L1 ⊢ T2 [d, e] ▶ U2 → L0 ⊢ T2 [d, e] ▶* U2. +/3 width=3/ qed. + +lemma ltpss_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. d1 + e1 ≤ d2 → + L0 ⊢ T2 [d2, e2] ▶* U2 → L0 [d1, e1] ▶* L1 → + L1 ⊢ T2 [d2, e2] ▶* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HTU2 #H @(ltpss_ind … H) -L1 +[ // +| -HTU2 #L #L1 #_ #HL1 #IHL + lapply (ltps_tpss_conf_ge … HL1 IHL) -HL1 -IHL // +] +qed. + +lemma ltpss_tps_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. d1 + e1 ≤ d2 → + L0 ⊢ T2 [d2, e2] ▶ U2 → L0 [d1, e1] ▶* L1 → + L1 ⊢ T2 [d2, e2] ▶* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HTU2 #HL01 +@(ltpss_tpss_conf_ge … Hde1d2 … HL01) /2 width=1/ (**) (* /3 width=6/ is too slow *) +qed. + +lemma ltpss_tpss_conf_eq: ∀L0,L1,T2,U2,d,e. + L0 ⊢ T2 [d, e] ▶* U2 → L0 [d, e] ▶* L1 → + ∃∃T. L1 ⊢ T2 [d, e] ▶* T & L1 ⊢ U2 [d, e] ▶* T. +#L0 #L1 #T2 #U2 #d #e #HTU2 #H @(ltpss_ind … H) -L1 +[ /2 width=3/ +| -HTU2 #L #L1 #_ #HL1 * #W2 #HTW2 #HUW2 + elim (ltps_tpss_conf … HL1 HTW2) -HTW2 #T #HT2 #HW2T + elim (ltps_tpss_conf … HL1 HUW2) -HL1 -HUW2 #U #HU2 #HW2U + elim (tpss_conf_eq … HW2T … HW2U) -HW2T -HW2U #V #HTV #HUV + lapply (tpss_trans_eq … HT2 … HTV) -T + lapply (tpss_trans_eq … HU2 … HUV) -U /2 width=3/ +] +qed. + +lemma ltpss_tps_conf_eq: ∀L0,L1,T2,U2,d,e. + L0 ⊢ T2 [d, e] ▶ U2 → L0 [d, e] ▶* L1 → + ∃∃T. L1 ⊢ T2 [d, e] ▶* T & L1 ⊢ U2 [d, e] ▶* T. +/3 width=3/ qed. + +lemma ltpss_tpss_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶* T2 → + ∀L2,ds,es. L1 [ds, es] ▶* L2 → + ∃∃T. L2 ⊢ T1 [d, e] ▶* T & L1 ⊢ T2 [ds, es] ▶* T. +#L1 #T1 #T2 #d #e #HT12 #L2 #ds #es #H @(ltpss_ind … H) -L2 +[ /3 width=3/ +| #L #L2 #HL1 #HL2 * #T #HT1 #HT2 + elim (ltps_tpss_conf … HL2 HT1) -HT1 #T0 #HT10 #HT0 + lapply (ltps_tpss_trans_eq … HL2 … HT0) -HL2 -HT0 #HT0 + lapply (ltpss_tpss_trans_eq … HL1 … HT0) -HL1 -HT0 #HT0 + lapply (tpss_trans_eq … HT2 … HT0) -T /2 width=3/ +] +qed. + +lemma ltpss_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 → + ∀L2,ds,es. L1 [ds, es] ▶* L2 → + ∃∃T. L2 ⊢ T1 [d, e] ▶* T & L1 ⊢ T2 [ds, es] ▶* T. +/3 width=1/ qed. + +(* Advanced properties ******************************************************) + +lemma ltpss_tps2: ∀L1,L2,I,e. + L1 [0, e] ▶* L2 → ∀V1,V2. L2 ⊢ V1 [0, e] ▶ V2 → + L1. ⓑ{I} V1 [0, e + 1] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #e #H @(ltpss_ind … H) -L2 +[ /3 width=1/ +| #L #L2 #_ #HL2 #IHL #V1 #V2 #HV12 + elim (ltps_tps_trans … HV12 … HL2) -HV12 #V #HV1 #HV2 + lapply (IHL … HV1) -IHL -HV1 #HL1 + @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +] +qed. + +lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e. + L1 [0, e - 1] ▶* L2 → L2 ⊢ V1 [0, e - 1] ▶ V2 → + 0 < e → L1. ⓑ{I} V1 [0, e] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He +>(plus_minus_m_m e 1) // /2 width=1/ +qed. + +lemma ltpss_tps1: ∀L1,L2,I,d,e. L1 [d, e] ▶* L2 → + ∀V1,V2. L2 ⊢ V1 [d, e] ▶ V2 → + L1. ⓑ{I} V1 [d + 1, e] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #d #e #H @(ltpss_ind … H) -L2 +[ /3 width=1/ +| #L #L2 #_ #HL2 #IHL #V1 #V2 #HV12 + elim (ltps_tps_trans … HV12 … HL2) -HV12 #V #HV1 #HV2 + lapply (IHL … HV1) -IHL -HV1 #HL1 + @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +] +qed. + +lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 [d - 1, e] ▶* L2 → L2 ⊢ V1 [d - 1, e] ▶ V2 → + 0 < d → L1. ⓑ{I} V1 [d, e] ▶* L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd +>(plus_minus_m_m d 1) // /2 width=1/ +qed. + +(* Advanced forward lemmas **************************************************) + +lemma ltpss_fwd_tpss21: ∀e,K1,I,V1,L2. 0 < e → K1. ⓑ{I} V1 [0, e] ▶* L2 → + ∃∃K2,V2. K1 [0, e - 1] ▶* K2 & K1 ⊢ V1 [0, e - 1] ▶* V2 & + L2 = K2. ⓑ{I} V2. +#e #K1 #I #V1 #L2 #He #H @(ltpss_ind … H) -L2 +[ /2 width=5/ +| #L #L2 #_ #HL2 * #K #V #HK1 #HV1 #H destruct + elim (ltps_inv_tps21 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H + lapply (ltps_tps_trans_eq … HV2 … HK2) -HV2 #HV2 + lapply (ltpss_tpss_trans_eq … HK1 … HV2) -HV2 #HV2 /3 width=5/ +] +qed-. + +lemma ltpss_fwd_tpss11: ∀d,e,I,K1,V1,L2. 0 < d → K1. ⓑ{I} V1 [d, e] ▶* L2 → + ∃∃K2,V2. K1 [d - 1, e] ▶* K2 & + K1 ⊢ V1 [d - 1, e] ▶* V2 & + L2 = K2. ⓑ{I} V2. +#d #e #K1 #I #V1 #L2 #Hd #H @(ltpss_ind … H) -L2 +[ /2 width=5/ +| #L #L2 #_ #HL2 * #K #V #HK1 #HV1 #H destruct + elim (ltps_inv_tps11 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H + lapply (ltps_tps_trans_eq … HV2 … HK2) -HV2 #HV2 + lapply (ltpss_tpss_trans_eq … HK1 … HV2) -HV2 #HV2 /3 width=5/ +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma new file mode 100644 index 000000000..16a0d3e4d --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma @@ -0,0 +1,146 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +definition tpss: nat → nat → lenv → relation term ≝ + λd,e,L. TC … (tps d e L). + +interpretation "partial unfold (term)" + 'PSubstStar L T1 d e T2 = (tpss d e L T1 T2). + +(* Basic eliminators ********************************************************) + +lemma tpss_ind: ∀d,e,L,T1. ∀R:predicate term. R T1 → + (∀T,T2. L ⊢ T1 [d, e] ▶* T → L ⊢ T [d, e] ▶ T2 → R T → R T2) → + ∀T2. L ⊢ T1 [d, e] ▶* T2 → R T2. +#d #e #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // +qed-. + +(* Basic properties *********************************************************) + +lemma tpss_strap: ∀L,T1,T,T2,d,e. + L ⊢ T1 [d, e] ▶ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 [d, e] ▶* T2. +/2 width=3/ qed. + +lemma tpss_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶* T2 → + ∀L2. L1 [d, e] ≼ L2 → L2 ⊢ T1 [d, e] ▶* T2. +/3 width=3/ qed. + +lemma tpss_refl: ∀d,e,L,T. L ⊢ T [d, e] ▶* T. +/2 width=1/ qed. + +lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 [d, e] ▶* V2 → + ∀I,T1,T2. L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶* T2 → + L ⊢ ⓑ{I} V1. T1 [d, e] ▶* ⓑ{I} V2. T2. +#L #V1 #V2 #d #e #HV12 elim HV12 -V2 +[ #V2 #HV12 #I #T1 #T2 #HT12 elim HT12 -T2 + [ /3 width=5/ + | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) + ] +| #V #V2 #_ #HV12 #IHV #I #T1 #T2 #HT12 + lapply (tpss_lsubs_conf … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12 + lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +] +qed. + +lemma tpss_flat: ∀L,I,V1,V2,T1,T2,d,e. + L ⊢ V1 [d, e] ▶ * V2 → L ⊢ T1 [d, e] ▶* T2 → + L ⊢ ⓕ{I} V1. T1 [d, e] ▶* ⓕ{I} V2. T2. +#L #I #V1 #V2 #T1 #T2 #d #e #HV12 elim HV12 -V2 +[ #V2 #HV12 #HT12 elim HT12 -T2 + [ /3 width=1/ + | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) + ] +| #V #V2 #_ #HV12 #IHV #HT12 + lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +] +qed. + +lemma tpss_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ▶* T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + L ⊢ T1 [d2, e2] ▶* T2. +#L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT12 #IHT + lapply (tps_weak … HT12 … Hd21 Hde12) -HT12 -Hd21 -Hde12 /2 width=3/ +] +qed. + +lemma tpss_weak_top: ∀L,T1,T2,d,e. + L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 [d, |L| - d] ▶* T2. +#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT12 #IHT + lapply (tps_weak_top … HT12) -HT12 /2 width=3/ +] +qed. + +lemma tpss_weak_all: ∀L,T1,T2,d,e. + L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 [0, |L|] ▶* T2. +#L #T1 #T2 #d #e #HT12 +lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 +lapply (tpss_weak_top … HT12) // +qed. + +(* Basic inversion lemmas ***************************************************) + +(* Note: this can be derived from tpss_inv_atom1 *) +lemma tpss_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k [d, e] ▶* T2 → T2 = ⋆k. +#L #T2 #k #d #e #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT2 #IHT destruct + >(tps_inv_sort1 … HT2) -HT2 // +] +qed-. + +(* Note: this can be derived from tpss_inv_atom1 *) +lemma tpss_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p [d, e] ▶* T2 → T2 = §p. +#L #T2 #p #d #e #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT2 #IHT destruct + >(tps_inv_gref1 … HT2) -HT2 // +] +qed-. + +lemma tpss_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 [d, e] ▶* U2 → + ∃∃V2,T2. L ⊢ V1 [d, e] ▶* V2 & + L. ⓑ{I} V2 ⊢ T1 [d + 1, e] ▶* T2 & + U2 = ⓑ{I} V2. T2. +#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 +[ /2 width=5/ +| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct + elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H + lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ +] +qed-. + +lemma tpss_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 [d, e] ▶* U2 → + ∃∃V2,T2. L ⊢ V1 [d, e] ▶* V2 & L ⊢ T1 [d, e] ▶* T2 & + U2 = ⓕ{I} V2. T2. +#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 +[ /2 width=5/ +| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct + elim (tps_inv_flat1 … HU2) -HU2 /3 width=5/ +] +qed-. + +lemma tpss_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ▶* T2 → T1 = T2. +#L #T1 #T2 #d #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT2 #IHT <(tps_inv_refl_O2 … HT2) -HT2 // +] +qed-. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_lift.ma new file mode 100644 index 000000000..748efd827 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_lift.ma @@ -0,0 +1,166 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps_lift.ma". +include "Basic_2/unfold/tpss.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +(* Advanced properties ******************************************************) + +lemma tpss_subst: ∀L,K,V,U1,i,d,e. + d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV → K ⊢ V [0, d + e - i - 1] ▶* U1 → + ∀U2. ⇧[0, i + 1] U1 ≡ U2 → L ⊢ #i [d, e] ▶* U2. +#L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(tpss_ind … H) -U1 +[ /3 width=4/ +| #U #U1 #_ #HU1 #IHU #U2 #HU12 + elim (lift_total U 0 (i+1)) #U0 #HU0 + lapply (IHU … HU0) -IHU #H + lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK + lapply (tps_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // normalize #HU02 + lapply (tps_weak … HU02 d e ? ?) -HU02 [ >minus_plus >commutative_plus /2 width=1/ | /2 width=1/ | /2 width=3/ ] +] +qed. + +(* Advanced inverion lemmas *************************************************) + +lemma tpss_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} [d, e] ▶* T2 → + T2 = ⓪{I} ∨ + ∃∃K,V1,V2,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV1 & + K ⊢ V1 [0, d + e - i - 1] ▶* V2 & + ⇧[O, i + 1] V2 ≡ T2 & + I = LRef i. +#L #T2 #I #d #e #H @(tpss_ind … H) -T2 +[ /2 width=1/ +| #T #T2 #_ #HT2 * + [ #H destruct + elim (tps_inv_atom1 … HT2) -HT2 [ /2 width=1/ | * /3 width=10/ ] + | * #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI + lapply (ldrop_fwd_ldrop2 … HLK) #H + elim (tps_inv_lift1_up … HT2 … H … HVT ? ? ?) normalize -HT2 -H -HVT [2,3,4: /2 width=1/ ] #V2 (lift_mono … HTU1 … H) -H // +| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (IHT … HTU) -IHT #HU1 + lapply (tps_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ +] +qed. + +lemma tpss_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶* T2 → + ∀L,U1,d,e. dt ≤ d → d ≤ dt + et → + ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → + ∀U2. ⇧[d, e] T2 ≡ U2 → L ⊢ U1 [dt, et + e] ▶* U2. +#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hdtd #Hddet #HLK #HTU1 @(tpss_ind … H) -T2 +[ #U2 #H >(lift_mono … HTU1 … H) -H // +| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (IHT … HTU) -IHT #HU1 + lapply (tps_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ +] +qed. + +lemma tpss_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ▶* T2 → + ∀L,U1,d,e. d ≤ dt → ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → + L ⊢ U1 [dt + e, et] ▶* U2. +#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hddt #HLK #HTU1 @(tpss_ind … H) -T2 +[ #U2 #H >(lift_mono … HTU1 … H) -H // +| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (IHT … HTU) -IHT #HU1 + lapply (tps_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ +] +qed. + +lemma tpss_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. K ⊢ T1 [dt, et] ▶* T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdetd @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_le … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_be: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 [dt, et - e] ▶* T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_be … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d + e ≤ dt → + ∃∃T2. K ⊢ T1 [dt - e, et] ▶* T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdedt @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_ge … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_eq: ∀L,U1,U2,d,e. + L ⊢ U1 [d, e] ▶* U2 → ∀T1. ⇧[d, e] T1 ≡ U1 → U1 = U2. +#L #U1 #U2 #d #e #H #T1 #HTU1 @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU destruct +<(tps_inv_lift1_eq … HU2 … HTU1) -HU2 -HTU1 // +qed. + +lemma tpss_inv_lift1_be_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → dt + et ≤ d + e → + ∃∃T2. K ⊢ T1 [dt, d - dt] ▶* T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_ltps.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_ltps.ma new file mode 100644 index 000000000..58b63cd01 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_ltps.ma @@ -0,0 +1,95 @@ +(**************************************************************************) +(* ___ *) +(* ||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/ltps_tps.ma". +include "Basic_2/unfold/tpss_tpss.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +(* Properties concerning parallel substitution on local environments ********) + +lemma ltps_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. + d1 + e1 ≤ d2 → L0 [d1, e1] ▶ L1 → + L0 ⊢ T2 [d2, e2] ▶* U2 → L1 ⊢ T2 [d2, e2] ▶* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL01 #H @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltps_tps_conf_ge … HU2 … HL01 ?) -HU2 -HL01 // /2 width=3/ +qed. + +lemma ltps_tpss_conf: ∀L0,L1,T2,U2,d1,e1,d2,e2. + L0 [d1, e1] ▶ L1 → L0 ⊢ T2 [d2, e2] ▶* U2 → + ∃∃T. L1 ⊢ T2 [d2, e2] ▶* T & L1 ⊢ U2 [d1, e1] ▶* T. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #HL01 #H @(tpss_ind … H) -U2 +[ /3 width=3/ +| #U #U2 #_ #HU2 * #T #HT2 #HUT + elim (ltps_tps_conf … HU2 … HL01) -HU2 -HL01 #W #HUW #HU2W + elim (tpss_strip_eq … HUT … HUW) -U + /3 width=5 by ex2_1_intro, step, tpss_strap/ (**) (* just /3 width=5/ is too slow *) +] +qed. + +lemma ltps_tpss_trans_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. + d1 + e1 ≤ d2 → L1 [d1, e1] ▶ L0 → + L0 ⊢ T2 [d2, e2] ▶* U2 → L1 ⊢ T2 [d2, e2] ▶* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL10 #H @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltps_tps_trans_ge … HU2 … HL10 ?) -HU2 -HL10 // /2 width=3/ +qed. + +lemma ltps_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 → + L1 [d1, e1] ▶ L0 → L0 ⊢ T2 [d2, e2] ▶* U2 → + ∃∃T. L1 ⊢ T2 [d2, e2] ▶* T & L0 ⊢ T [d1, e1] ▶* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde2d1 #HL10 #H @(tpss_ind … H) -U2 +[ /3 width=3/ +| #U #U2 #_ #HU2 * #T #HT2 #HTU + elim (tpss_strap1_down … HTU … HU2 ?) -U // #U #HTU #HU2 + elim (ltps_tps_trans … HTU … HL10) -HTU -HL10 #W #HTW #HWU + @(ex2_1_intro … W) /2 width=3/ (**) (* /3 width=5/ does not work as in ltps_tpss_conf *) +] +qed. + +fact ltps_tps_trans_eq_aux: ∀Y1,X2,L1,T2,U2,d,e. + L1 ⊢ T2 [d, e] ▶ U2 → ∀L0. L0 [d, e] ▶ L1 → + Y1 = L1 → X2 = T2 → L0 ⊢ T2 [d, e] ▶* U2. +#Y1 #X2 @(cw_wf_ind … Y1 X2) -Y1 -X2 #Y1 #X2 #IH +#L1 #T2 #U2 #d #e * -L1 -T2 -U2 -d -e +[ // +| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L0 #HL10 #H1 #H2 destruct + lapply (ldrop_fwd_lw … HLK1) normalize #H1 + elim (ltps_ldrop_trans_be … HL10 … HLK1 ? ?) -HL10 -HLK1 // /2 width=2/ #X #H #HLK0 + elim (ltps_inv_tps22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct + lapply (tps_fwd_tw … HV01) #H2 + lapply (transitive_le (#[K1] + #[V0]) … H1) -H1 /2 width=1/ -H2 #H + lapply (IH … HV01 … HK01 ? ?) -IH -HV01 -HK01 + [1,3: // |2,4: skip | normalize /2 width=1/ | /3 width=6/ ] +| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct + lapply (tps_lsubs_conf … HT12 (L. ⓑ{I} V1) ?) -HT12 /2 width=1/ #HT12 + lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=2/ ] #HV12 + lapply (IH … HT12 (L0. ⓑ{I} V1) ? ? ?) -IH -HT12 [1,3,5: /2 width=2/ |2,4: skip | normalize // ] -HL0 #HT12 + lapply (tpss_lsubs_conf … HT12 (L0. ⓑ{I} V2) ?) -HT12 /2 width=1/ +| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct + lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=3/ ] + lapply (IH … HT12 … HL0 ? ?) -IH -HT12 [1,3,5: normalize // |2,4: skip ] -HL0 /2 width=1/ +] +qed. + +lemma ltps_tps_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 [d, e] ▶ U2 → + ∀L0. L0 [d, e] ▶ L1 → L0 ⊢ T2 [d, e] ▶* U2. +/2 width=5/ qed. + +lemma ltps_tpss_trans_eq: ∀L0,L1,T2,U2,d,e. L0 [d, e] ▶ L1 → + L1 ⊢ T2 [d, e] ▶* U2 → L0 ⊢ T2 [d, e] ▶* U2. +#L0 #L1 #T2 #U2 #d #e #HL01 #H @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU @(tpss_trans_eq … IHU) /2 width=3/ +qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_tpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_tpss.ma new file mode 100644 index 000000000..981402a24 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss_tpss.ma @@ -0,0 +1,93 @@ +(**************************************************************************) +(* ___ *) +(* ||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/tps_tps.ma". +include "Basic_2/unfold/tpss_lift.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +(* Advanced properties ******************************************************) + +lemma tpss_tps: ∀L,T1,T2,d. L ⊢ T1 [d, 1] ▶* T2 → L ⊢ T1 [d, 1] ▶ T2. +#L #T1 #T2 #d #H @(tpss_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 +lapply (tps_trans_ge … IHT1 … HT2 ?) // +qed. + +lemma tpss_strip_eq: ∀L,T0,T1,d1,e1. L ⊢ T0 [d1, e1] ▶* T1 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶ T2 → + ∃∃T. L ⊢ T1 [d2, e2] ▶ T & L ⊢ T2 [d1, e1] ▶* T. +/3 width=3/ qed. + +lemma tpss_strip_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 [d1, e1] ▶* T1 → + ∀L2,T2,d2,e2. L2 ⊢ T0 [d2, e2] ▶ T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. L2 ⊢ T1 [d2, e2] ▶ T & L1 ⊢ T2 [d1, e1] ▶* T. +/3 width=3/ qed. + +lemma tpss_strap1_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ▶* T0 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶ T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 [d2, e2] ▶ T & L ⊢ T [d1, e1] ▶* T2. +/3 width=3/ qed. + +lemma tpss_strap2_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ▶ T0 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ▶* T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 [d2, e2] ▶* T & L ⊢ T [d1, e1] ▶ T2. +/3 width=3/ qed. + +lemma tpss_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶* T2 → + ∀i. d ≤ i → i ≤ d + e → + ∃∃T. L ⊢ T1 [d, i - d] ▶* T & L ⊢ T [i, d + e - i] ▶* T2. +#L #T1 #T2 #d #e #H #i #Hdi #Hide @(tpss_ind … H) -T2 +[ /2 width=3/ +| #T #T2 #_ #HT12 * #T3 #HT13 #HT3 + elim (tps_split_up … HT12 … Hdi Hide) -HT12 -Hide #T0 #HT0 #HT02 + elim (tpss_strap1_down … HT3 … HT0 ?) -T [2: >commutative_plus /2 width=1/ ] + /3 width=7 by ex2_1_intro, step/ (**) (* just /3 width=7/ is too slow *) +] +qed. + +lemma tpss_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ▶* U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ▶* T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (tpss_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (tpss_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 +lapply (tpss_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct +elim (tpss_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 -HLK -HTU1 // minus_minus_comm >minus_le_minus_minus_comm // +qed. + +lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b. +#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1/ +qed. + +lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. +/3 by monotonic_le_minus_l, le_to_le_to_eq, le_n/ qed. + +lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b → + a1 - c1 + a2 - (b - c1) = a1 + a2 - b. +#a1 #a2 #b #c1 #H1 #H2 >plus_minus // /2 width=1/ +qed. + +(* inversion & forward lemmas ***********************************************) + +axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). + +axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). + +lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m. +#m #n elim (lt_or_ge m n) /2 width=1/ +#H elim H -m /2 width=1/ +#m #Hm * #H /2 width=1/ /3 width=1/ +qed-. + +lemma lt_refl_false: ∀n. n < n → False. +#n #H elim (lt_to_not_eq … H) -H /2 width=1/ +qed-. + +lemma lt_zero_false: ∀n. n < 0 → False. +#n #H elim (lt_to_not_le … H) -H /2 width=1/ +qed-. + +lemma false_lt_to_le: ∀x,y. (x < y → False) → y ≤ x. +#x #y #H elim (decidable_lt x y) /2 width=1/ +#Hxy elim (H Hxy) +qed-. + +(* +lemma pippo: ∀x,y,z. x < z → y < z - x → x + y < z. +/3 width=2/ + +lemma le_or_ge: ∀m,n. m ≤ n ∨ n ≤ m. +#m #n elim (lt_or_ge m n) /2 width=1/ /3 width=2/ +qed-. +*) diff --git a/matita/matita/contribs/lambda_delta/ground_2/list.ma b/matita/matita/contribs/lambda_delta/ground_2/list.ma new file mode 100644 index 000000000..0a6e69bbe --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/list.ma @@ -0,0 +1,55 @@ +(**************************************************************************) +(* ___ *) +(* ||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/arith.ma". + +(* LISTS ********************************************************************) + +inductive list (A:Type[0]) : Type[0] := + | nil : list A + | cons: A → list A → list A. + +interpretation "nil (list)" 'Nil = (nil ?). + +interpretation "cons (list)" 'Cons hd tl = (cons ? hd tl). + +let rec all A (R:predicate A) (l:list A) on l ≝ + match l with + [ nil ⇒ True + | cons hd tl ⇒ R hd ∧ all A R tl + ]. + +inductive list2 (A1,A2:Type[0]) : Type[0] := + | nil2 : list2 A1 A2 + | cons2: A1 → A2 → list2 A1 A2 → list2 A1 A2. + +interpretation "nil (list of pairs)" 'Nil2 = (nil2 ? ?). + +interpretation "cons (list of pairs)" 'Cons hd1 hd2 tl = (cons2 ? ? hd1 hd2 tl). + +let rec append2 (A1,A2:Type[0]) (l1,l2:list2 A1 A2) on l1 ≝ match l1 with +[ nil2 ⇒ l2 +| cons2 a1 a2 tl ⇒ {a1, a2} :: append2 A1 A2 tl l2 +]. + +interpretation "append (list of pairs)" + 'Append l1 l2 = (append2 ? ? l1 l2). + +let rec length2 (A1,A2:Type[0]) (l:list2 A1 A2) on l ≝ match l with +[ nil2 ⇒ 0 +| cons2 _ _ l ⇒ length2 A1 A2 l + 1 +]. + +interpretation "length (list of pairs)" + 'card l = (length2 ? ? l). diff --git a/matita/matita/contribs/lambda_delta/ground_2/notation.ma b/matita/matita/contribs/lambda_delta/ground_2/notation.ma new file mode 100644 index 000000000..71fbbe0c2 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/notation.ma @@ -0,0 +1,37 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +(* GENERAL NOTATION USED BY THE FORMAL SYSTEM λδ ****************************) + +(* Lists ********************************************************************) + +notation "hvbox( ◊ )" + non associative with precedence 90 + for @{'Nil}. + +notation "hvbox( hd :: break tl )" + right associative with precedence 47 + for @{'Cons $hd $tl}. + +notation "hvbox( l1 @ break l2 )" + right associative with precedence 47 + for @{'Append $l1 $l2 }. + +notation "hvbox( ⟠ )" + non associative with precedence 90 + for @{'Nil2}. + +notation "hvbox( { hd1 , break hd2 } :: break tl )" + non associative with precedence 47 + for @{'Cons $hd1 $hd2 $tl}. diff --git a/matita/matita/contribs/lambda_delta/ground_2/star.ma b/matita/matita/contribs/lambda_delta/ground_2/star.ma new file mode 100644 index 000000000..ed3580642 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/star.ma @@ -0,0 +1,112 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "basics/star.ma". +include "Ground_2/xoa_props.ma". +include "Ground_2/notation.ma". + +(* PROPERTIES OF RELATIONS **************************************************) + +definition Decidable: Prop → Prop ≝ λR. R ∨ (R → False). + +definition confluent2: ∀A. ∀R1,R2: relation A. Prop ≝ λA,R1,R2. + ∀a0,a1. R1 a0 a1 → ∀a2. R2 a0 a2 → + ∃∃a. R2 a1 a & R1 a2 a. + +definition transitive2: ∀A. ∀R1,R2: relation A. Prop ≝ λA,R1,R2. + ∀a1,a0. R1 a1 a0 → ∀a2. R2 a0 a2 → + ∃∃a. R2 a1 a & R1 a a2. + +lemma TC_strip1: ∀A,R1,R2. confluent2 A R1 R2 → + ∀a0,a1. TC … R1 a0 a1 → ∀a2. R2 a0 a2 → + ∃∃a. R2 a1 a & TC … R1 a2 a. +#A #R1 #R2 #HR12 #a0 #a1 #H elim H -a1 +[ #a1 #Ha01 #a2 #Ha02 + elim (HR12 … Ha01 … Ha02) -HR12 -a0 /3 width=3/ +| #a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 + elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20 + elim (HR12 … Ha1 … Ha0) -HR12 -a /4 width=3/ +] +qed. + +lemma TC_strip2: ∀A,R1,R2. confluent2 A R1 R2 → + ∀a0,a2. TC … R2 a0 a2 → ∀a1. R1 a0 a1 → + ∃∃a. TC … R2 a1 a & R1 a2 a. +#A #R1 #R2 #HR12 #a0 #a2 #H elim H -a2 +[ #a2 #Ha02 #a1 #Ha01 + elim (HR12 … Ha01 … Ha02) -HR12 -a0 /3 width=3/ +| #a #a2 #_ #Ha2 #IHa0 #a1 #Ha01 + elim (IHa0 … Ha01) -a0 #a0 #Ha10 #Ha0 + elim (HR12 … Ha0 … Ha2) -HR12 -a /4 width=3/ +] +qed. + +lemma TC_confluent2: ∀A,R1,R2. + confluent2 A R1 R2 → confluent2 A (TC … R1) (TC … R2). +#A #R1 #R2 #HR12 #a0 #a1 #H elim H -a1 +[ #a1 #Ha01 #a2 #Ha02 + elim (TC_strip2 … HR12 … Ha02 … Ha01) -HR12 -a0 /3 width=3/ +| #a #a1 #_ #Ha1 #IHa0 #a2 #Ha02 + elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20 + elim (TC_strip2 … HR12 … Ha0 … Ha1) -HR12 -a /4 width=3/ +] +qed. + +lemma TC_strap1: ∀A,R1,R2. transitive2 A R1 R2 → + ∀a1,a0. TC … R1 a1 a0 → ∀a2. R2 a0 a2 → + ∃∃a. R2 a1 a & TC … R1 a a2. +#A #R1 #R2 #HR12 #a1 #a0 #H elim H -a0 +[ #a0 #Ha10 #a2 #Ha02 + elim (HR12 … Ha10 … Ha02) -HR12 -a0 /3 width=3/ +| #a #a0 #_ #Ha0 #IHa #a2 #Ha02 + elim (HR12 … Ha0 … Ha02) -HR12 -a0 #a0 #Ha0 #Ha02 + elim (IHa … Ha0) -a /4 width=3/ +] +qed. + +lemma TC_strap2: ∀A,R1,R2. transitive2 A R1 R2 → + ∀a0,a2. TC … R2 a0 a2 → ∀a1. R1 a1 a0 → + ∃∃a. TC … R2 a1 a & R1 a a2. +#A #R1 #R2 #HR12 #a0 #a2 #H elim H -a2 +[ #a2 #Ha02 #a1 #Ha10 + elim (HR12 … Ha10 … Ha02) -HR12 -a0 /3 width=3/ +| #a #a2 #_ #Ha02 #IHa #a1 #Ha10 + elim (IHa … Ha10) -a0 #a0 #Ha10 #Ha0 + elim (HR12 … Ha0 … Ha02) -HR12 -a /4 width=3/ +] +qed. + +lemma TC_transitive2: ∀A,R1,R2. + transitive2 A R1 R2 → transitive2 A (TC … R1) (TC … R2). +#A #R1 #R2 #HR12 #a1 #a0 #H elim H -a0 +[ #a0 #Ha10 #a2 #Ha02 + elim (TC_strap2 … HR12 … Ha02 … Ha10) -HR12 -a0 /3 width=3/ +| #a #a0 #_ #Ha0 #IHa #a2 #Ha02 + elim (TC_strap2 … HR12 … Ha02 … Ha0) -HR12 -a0 #a0 #Ha0 #Ha02 + elim (IHa … Ha0) -a /4 width=3/ +] +qed. + +definition NF: ∀A. relation A → relation A → predicate A ≝ + λA,R,S,a1. ∀a2. R a1 a2 → S a1 a2. + +inductive SN (A) (R,S:relation A): predicate A ≝ +| SN_intro: ∀a1. (∀a2. R a1 a2 → (S a1 a2 → False) → SN A R S a2) → SN A R S a1 +. + +lemma NF_to_SN: ∀A,R,S,a. NF A R S a → SN A R S a. +#A #R #S #a1 #Ha1 +@SN_intro #a2 #HRa12 #HSa12 +elim (HSa12 ?) -HSa12 /2 width=1/ +qed. diff --git a/matita/matita/contribs/lambda_delta/ground_2/tri.ma b/matita/matita/contribs/lambda_delta/ground_2/tri.ma new file mode 100644 index 000000000..005806c11 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/tri.ma @@ -0,0 +1,44 @@ +(**************************************************************************) +(* ___ *) +(* ||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/arith.ma". + +(* TRICOTOMY FUNCTION *******************************************************) + +let rec tri (A:Type[0]) m n a b c on m : A ≝ + match m with + [ O ⇒ match n with [ O ⇒ b | S n ⇒ a ] + | S m ⇒ match n with [ O ⇒ c | S n ⇒ tri A m n a b c ] + ]. + +(* Basic properties *********************************************************) + +lemma tri_lt: ∀A,a,b,c,n,m. m < n → tri A m n a b c = a. +#A #a #b #c #n elim n -n +[ #m #H elim (lt_zero_false … H) +| #n #IH #m elim m -m // /3 width=1/ +] +qed. + +lemma tri_eq: ∀A,a,b,c,m. tri A m m a b c = b. +#A #a #b #c #m elim m -m normalize // +qed. + +lemma tri_gt: ∀A,a,b,c,m,n. n < m → tri A m n a b c = c. +#A #a #b #c #m elim m -m +[ #n #H elim (lt_zero_false … H) +| #m #IH #n elim n -n // /3 width=1/ +] +qed. + diff --git a/matita/matita/contribs/lambda_delta/ground_2/xoa.conf.xml b/matita/matita/contribs/lambda_delta/ground_2/xoa.conf.xml new file mode 100644 index 000000000..2a2da2a5c --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/xoa.conf.xml @@ -0,0 +1,37 @@ + + +
+ $(MATITA_RT_BASE_DIR) + +
+
+ contribs/lambda_delta/Ground_2/ + xoa + xoa_notation + basics/pts.ma + 1 2 + 2 1 + 2 2 + 2 3 + 3 1 + 3 2 + 3 3 + 4 2 + 4 3 + 4 4 + 5 2 + 5 3 + 5 4 + 6 4 + 6 6 + 7 6 + 3 + 4 + 3 +
+
diff --git a/matita/matita/contribs/lambda_delta/ground_2/xoa.ma b/matita/matita/contribs/lambda_delta/ground_2/xoa.ma new file mode 100644 index 000000000..5a3a40024 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/xoa.ma @@ -0,0 +1,175 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +(* This file was generated by xoa.native: do not edit *********************) + +include "basics/pts.ma". + +(* multiple existental quantifier (1, 2) *) + +inductive ex1_2 (A0,A1:Type[0]) (P0:A0→A1→Prop) : Prop ≝ + | ex1_2_intro: ∀x0,x1. P0 x0 x1 → ex1_2 ? ? ? +. + +interpretation "multiple existental quantifier (1, 2)" 'Ex P0 = (ex1_2 ? ? P0). + +(* multiple existental quantifier (2, 1) *) + +inductive ex2_1 (A0:Type[0]) (P0,P1:A0→Prop) : Prop ≝ + | ex2_1_intro: ∀x0. P0 x0 → P1 x0 → ex2_1 ? ? ? +. + +interpretation "multiple existental quantifier (2, 1)" 'Ex P0 P1 = (ex2_1 ? P0 P1). + +(* multiple existental quantifier (2, 2) *) + +inductive ex2_2 (A0,A1:Type[0]) (P0,P1:A0→A1→Prop) : Prop ≝ + | ex2_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → ex2_2 ? ? ? ? +. + +interpretation "multiple existental quantifier (2, 2)" 'Ex P0 P1 = (ex2_2 ? ? P0 P1). + +(* multiple existental quantifier (2, 3) *) + +inductive ex2_3 (A0,A1,A2:Type[0]) (P0,P1:A0→A1→A2→Prop) : Prop ≝ + | ex2_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → ex2_3 ? ? ? ? ? +. + +interpretation "multiple existental quantifier (2, 3)" 'Ex P0 P1 = (ex2_3 ? ? ? P0 P1). + +(* multiple existental quantifier (3, 1) *) + +inductive ex3_1 (A0:Type[0]) (P0,P1,P2:A0→Prop) : Prop ≝ + | ex3_1_intro: ∀x0. P0 x0 → P1 x0 → P2 x0 → ex3_1 ? ? ? ? +. + +interpretation "multiple existental quantifier (3, 1)" 'Ex P0 P1 P2 = (ex3_1 ? P0 P1 P2). + +(* multiple existental quantifier (3, 2) *) + +inductive ex3_2 (A0,A1:Type[0]) (P0,P1,P2:A0→A1→Prop) : Prop ≝ + | ex3_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → ex3_2 ? ? ? ? ? +. + +interpretation "multiple existental quantifier (3, 2)" 'Ex P0 P1 P2 = (ex3_2 ? ? P0 P1 P2). + +(* multiple existental quantifier (3, 3) *) + +inductive ex3_3 (A0,A1,A2:Type[0]) (P0,P1,P2:A0→A1→A2→Prop) : Prop ≝ + | ex3_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → ex3_3 ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (3, 3)" 'Ex P0 P1 P2 = (ex3_3 ? ? ? P0 P1 P2). + +(* multiple existental quantifier (4, 2) *) + +inductive ex4_2 (A0,A1:Type[0]) (P0,P1,P2,P3:A0→A1→Prop) : Prop ≝ + | ex4_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → P3 x0 x1 → ex4_2 ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (4, 2)" 'Ex P0 P1 P2 P3 = (ex4_2 ? ? P0 P1 P2 P3). + +(* multiple existental quantifier (4, 3) *) + +inductive ex4_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3:A0→A1→A2→Prop) : Prop ≝ + | ex4_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → ex4_3 ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (4, 3)" 'Ex P0 P1 P2 P3 = (ex4_3 ? ? ? P0 P1 P2 P3). + +(* multiple existental quantifier (4, 4) *) + +inductive ex4_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3:A0→A1→A2→A3→Prop) : Prop ≝ + | ex4_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → ex4_4 ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (4, 4)" 'Ex P0 P1 P2 P3 = (ex4_4 ? ? ? ? P0 P1 P2 P3). + +(* multiple existental quantifier (5, 2) *) + +inductive ex5_2 (A0,A1:Type[0]) (P0,P1,P2,P3,P4:A0→A1→Prop) : Prop ≝ + | ex5_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → P3 x0 x1 → P4 x0 x1 → ex5_2 ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (5, 2)" 'Ex P0 P1 P2 P3 P4 = (ex5_2 ? ? P0 P1 P2 P3 P4). + +(* multiple existental quantifier (5, 3) *) + +inductive ex5_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→Prop) : Prop ≝ + | ex5_3_intro: ∀x0,x1,x2. P0 x0 x1 x2 → P1 x0 x1 x2 → P2 x0 x1 x2 → P3 x0 x1 x2 → P4 x0 x1 x2 → ex5_3 ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (5, 3)" 'Ex P0 P1 P2 P3 P4 = (ex5_3 ? ? ? P0 P1 P2 P3 P4). + +(* multiple existental quantifier (5, 4) *) + +inductive ex5_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4:A0→A1→A2→A3→Prop) : Prop ≝ + | ex5_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → ex5_4 ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (5, 4)" 'Ex P0 P1 P2 P3 P4 = (ex5_4 ? ? ? ? P0 P1 P2 P3 P4). + +(* multiple existental quantifier (6, 4) *) + +inductive ex6_4 (A0,A1,A2,A3:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→Prop) : Prop ≝ + | ex6_4_intro: ∀x0,x1,x2,x3. P0 x0 x1 x2 x3 → P1 x0 x1 x2 x3 → P2 x0 x1 x2 x3 → P3 x0 x1 x2 x3 → P4 x0 x1 x2 x3 → P5 x0 x1 x2 x3 → ex6_4 ? ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (6, 4)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_4 ? ? ? ? P0 P1 P2 P3 P4 P5). + +(* multiple existental quantifier (6, 6) *) + +inductive ex6_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ + | ex6_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → ex6_6 ? ? ? ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (6, 6)" 'Ex P0 P1 P2 P3 P4 P5 = (ex6_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5). + +(* multiple existental quantifier (7, 6) *) + +inductive ex7_6 (A0,A1,A2,A3,A4,A5:Type[0]) (P0,P1,P2,P3,P4,P5,P6:A0→A1→A2→A3→A4→A5→Prop) : Prop ≝ + | ex7_6_intro: ∀x0,x1,x2,x3,x4,x5. P0 x0 x1 x2 x3 x4 x5 → P1 x0 x1 x2 x3 x4 x5 → P2 x0 x1 x2 x3 x4 x5 → P3 x0 x1 x2 x3 x4 x5 → P4 x0 x1 x2 x3 x4 x5 → P5 x0 x1 x2 x3 x4 x5 → P6 x0 x1 x2 x3 x4 x5 → ex7_6 ? ? ? ? ? ? ? ? ? ? ? ? ? +. + +interpretation "multiple existental quantifier (7, 6)" 'Ex P0 P1 P2 P3 P4 P5 P6 = (ex7_6 ? ? ? ? ? ? P0 P1 P2 P3 P4 P5 P6). + +(* multiple disjunction connective (3) *) + +inductive or3 (P0,P1,P2:Prop) : Prop ≝ + | or3_intro0: P0 → or3 ? ? ? + | or3_intro1: P1 → or3 ? ? ? + | or3_intro2: P2 → or3 ? ? ? +. + +interpretation "multiple disjunction connective (3)" 'Or P0 P1 P2 = (or3 P0 P1 P2). + +(* multiple disjunction connective (4) *) + +inductive or4 (P0,P1,P2,P3:Prop) : Prop ≝ + | or4_intro0: P0 → or4 ? ? ? ? + | or4_intro1: P1 → or4 ? ? ? ? + | or4_intro2: P2 → or4 ? ? ? ? + | or4_intro3: P3 → or4 ? ? ? ? +. + +interpretation "multiple disjunction connective (4)" 'Or P0 P1 P2 P3 = (or4 P0 P1 P2 P3). + +(* multiple conjunction connective (3) *) + +inductive and3 (P0,P1,P2:Prop) : Prop ≝ + | and3_intro: P0 → P1 → P2 → and3 ? ? ? +. + +interpretation "multiple conjunction connective (3)" 'And P0 P1 P2 = (and3 P0 P1 P2). + diff --git a/matita/matita/contribs/lambda_delta/ground_2/xoa_notation.ma b/matita/matita/contribs/lambda_delta/ground_2/xoa_notation.ma new file mode 100644 index 000000000..edd014ef0 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/xoa_notation.ma @@ -0,0 +1,194 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +(* This file was generated by xoa.native: do not edit *********************) + +(* multiple existental quantifier (1, 2) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) }. + +(* multiple existental quantifier (2, 1) *) + +notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) }. + +notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) }. + +(* multiple existental quantifier (2, 2) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) }. + +(* multiple existental quantifier (2, 3) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) }. + +(* multiple existental quantifier (3, 1) *) + +notation > "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.$P0) (λ${ident x0}.$P1) (λ${ident x0}.$P2) }. + +notation < "hvbox(∃∃ ident x0 break . term 19 P0 break & term 19 P1 break & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.$P0) (λ${ident x0}:$T0.$P1) (λ${ident x0}:$T0.$P2) }. + +(* multiple existental quantifier (3, 2) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) }. + +(* multiple existental quantifier (3, 3) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) }. + +(* multiple existental quantifier (4, 2) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) (λ${ident x0}.λ${ident x1}.$P3) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P3) }. + +(* multiple existental quantifier (4, 3) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P3) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P3) }. + +(* multiple existental quantifier (4, 4) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) }. + +(* multiple existental quantifier (5, 2) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) (λ${ident x0}.λ${ident x1}.$P3) (λ${ident x0}.λ${ident x1}.$P4) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P4) }. + +(* multiple existental quantifier (5, 3) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.$P4) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P4) }. + +(* multiple existental quantifier (5, 4) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P4) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) }. + +(* multiple existental quantifier (6, 4) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.$P5) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.$P5) }. + +(* multiple existental quantifier (6, 6) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P5) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) }. + +(* multiple existental quantifier (7, 6) *) + +notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P0) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P1) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P2) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P3) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P4) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P5) (λ${ident x0}.λ${ident x1}.λ${ident x2}.λ${ident x3}.λ${ident x4}.λ${ident x5}.$P6) }. + +notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 , ident x3 , ident x4 , ident x5 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3 break & term 19 P4 break & term 19 P5 break & term 19 P6)" + non associative with precedence 20 + for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P3) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P4) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P5) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.λ${ident x3}:$T3.λ${ident x4}:$T4.λ${ident x5}:$T5.$P6) }. + +(* multiple disjunction connective (3) *) + +notation "hvbox(∨∨ term 29 P0 break | term 29 P1 break | term 29 P2)" + non associative with precedence 30 + for @{ 'Or $P0 $P1 $P2 }. + +(* multiple disjunction connective (4) *) + +notation "hvbox(∨∨ term 29 P0 break | term 29 P1 break | term 29 P2 break | term 29 P3)" + non associative with precedence 30 + for @{ 'Or $P0 $P1 $P2 $P3 }. + +(* multiple conjunction connective (3) *) + +notation "hvbox(∧∧ term 34 P0 break & term 34 P1 break & term 34 P2)" + non associative with precedence 35 + for @{ 'And $P0 $P1 $P2 }. + diff --git a/matita/matita/contribs/lambda_delta/ground_2/xoa_props.ma b/matita/matita/contribs/lambda_delta/ground_2/xoa_props.ma new file mode 100644 index 000000000..121eb7e74 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/ground_2/xoa_props.ma @@ -0,0 +1,20 @@ +(**************************************************************************) +(* ___ *) +(* ||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/xoa_notation.ma". +include "Ground_2/xoa.ma". + +lemma ex2_1_comm: ∀A0. ∀P0,P1:A0→Prop. (∃∃x0. P0 x0 & P1 x0) → ∃∃x0. P1 x0 & P0 x0. +#A0 #P0 #P1 * /2 width=3/ +qed.