From: Ferruccio Guidi Date: Mon, 13 Jun 2022 22:30:57 +0000 (+0200) Subject: update in delayed_updating X-Git-Tag: make_still_working~60 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;ds=sidebyside;h=4ac2becfaa45abb18acb2bdf3db5d2587cadb6d4;p=helm.git update in delayed_updating + final definition of lift + WIP on dfr_lift_bi + updates and corrections + old definition of lift parked --- diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift.etc new file mode 100644 index 000000000..6933e9dbe --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift.etc @@ -0,0 +1,158 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The 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 "delayed_updating/notation/functions/uparrow_4.ma". +include "delayed_updating/notation/functions/uparrow_2.ma". +include "delayed_updating/syntax/path.ma". +include "ground/relocation/tr_id_pap.ma". + +(* LIFT FOR PATH ************************************************************) + +definition lift_continuation (A:Type[0]) ≝ + tr_map → path → A. + +rec definition lift_gen (A:Type[0]) (k:lift_continuation A) (f) (p) on p ≝ +match p with +[ list_empty ⇒ k f (𝐞) +| list_lcons l q ⇒ + match l with + [ label_d n ⇒ lift_gen (A) (λg,p. k g (𝗱(f@⧣❨n❩)◗p)) (𝐢) q + | label_m ⇒ lift_gen (A) (λg,p. k g (𝗺◗p)) f q + | label_L ⇒ lift_gen (A) (λg,p. k g (𝗟◗p)) (⫯f) q + | label_A ⇒ lift_gen (A) (λg,p. k g (𝗔◗p)) f q + | label_S ⇒ lift_gen (A) (λg,p. k g (𝗦◗p)) f q + ] +]. + +interpretation + "lift (gneric)" + 'UpArrow A k f p = (lift_gen A k f p). + +definition proj_path: lift_continuation … ≝ + λf,p.p. + +definition proj_rmap: lift_continuation … ≝ + λf,p.f. + +interpretation + "lift (path)" + 'UpArrow f p = (lift_gen ? proj_path f p). + +interpretation + "lift (relocation map)" + 'UpArrow p f = (lift_gen ? proj_rmap f p). + +(* Basic constructions ******************************************************) + +lemma lift_empty (A) (k) (f): + k f (𝐞) = ↑{A}❨k, f, 𝐞❩. +// qed. + +lemma lift_d_sn (A) (k) (p) (n) (f): + ↑❨(λg,p. k g (𝗱(f@⧣❨n❩)◗p)), 𝐢, p❩ = ↑{A}❨k, f, 𝗱n◗p❩. +// qed. + +lemma lift_m_sn (A) (k) (p) (f): + ↑❨(λg,p. k g (𝗺◗p)), f, p❩ = ↑{A}❨k, f, 𝗺◗p❩. +// qed. + +lemma lift_L_sn (A) (k) (p) (f): + ↑❨(λg,p. k g (𝗟◗p)), ⫯f, p❩ = ↑{A}❨k, f, 𝗟◗p❩. +// qed. + +lemma lift_A_sn (A) (k) (p) (f): + ↑❨(λg,p. k g (𝗔◗p)), f, p❩ = ↑{A}❨k, f, 𝗔◗p❩. +// qed. + +lemma lift_S_sn (A) (k) (p) (f): + ↑❨(λg,p. k g (𝗦◗p)), f, p❩ = ↑{A}❨k, f, 𝗦◗p❩. +// qed. + +(* Basic constructions with proj_path ***************************************) + +lemma lift_path_empty (f): + (𝐞) = ↑[f]𝐞. +// qed. + +(* Basic constructions with proj_rmap ***************************************) + +lemma lift_rmap_empty (f): + f = ↑[𝐞]f. +// qed. + +lemma lift_rmap_d_sn (f) (p) (n): + ↑[p]𝐢 = ↑[𝗱n◗p]f. +// qed. + +lemma lift_rmap_m_sn (f) (p): + ↑[p]f = ↑[𝗺◗p]f. +// qed. + +lemma lift_rmap_L_sn (f) (p): + ↑[p](⫯f) = ↑[𝗟◗p]f. +// qed. + +lemma lift_rmap_A_sn (f) (p): + ↑[p]f = ↑[𝗔◗p]f. +// qed. + +lemma lift_rmap_S_sn (f) (p): + ↑[p]f = ↑[𝗦◗p]f. +// qed. + +(* Advanced cinstructionswith proj_rmap and tr_id ***************************) + +lemma lift_rmap_id (p): + (𝐢) = ↑[p]𝐢. +#p elim p -p // +* [ #n ] #p #IH // +qed. + +(* Advanced constructions with proj_rmap and path_append ********************) + +lemma lift_rmap_append (p2) (p1) (f): + ↑[p2]↑[p1]f = ↑[p1●p2]f. +#p2 #p1 elim p1 -p1 // * [ #n ] #p1 #IH #f // +[ lift_lcons_alt // >lift_append_rcons_sn // +lift_lcons_alt nrplus_inj_dx +/2 width=1 by tr_tls_compose_uni_dx/ +qed. +*) diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_after.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_after.etc new file mode 100644 index 000000000..938d21534 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_after.etc @@ -0,0 +1,28 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The 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 "delayed_updating/substitution/lift_eq.ma". +include "ground/relocation/tr_compose_pap.ma". +include "ground/relocation/tr_compose_pn.ma". + +(* LIFT FOR PATH ************************************************************) + +(* Constructions with tr_after **********************************************) + +lemma lift_path_after (p) (f1) (f2): + ↑[f2]↑[f1]p = ↑[f2∘f1]p. +#p elim p -p [| * ] // #p #IH #f1 #f2 +tr_compose_push_bi // +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_proper.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_proper.etc new file mode 100644 index 000000000..0eef97896 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_proper.etc @@ -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 "delayed_updating/substitution/lift_eq.ma". +include "delayed_updating/syntax/path_proper.ma". + +(* LIFT FOR PATH ************************************************************) + +(* Constructions with proper condition for path *****************************) + +lemma lift_path_proper (f) (p): + p ϵ 𝐏 → ↑[f]p ϵ 𝐏. +#f * +[ #H0 elim (ppc_inv_empty … H0) +| * [ #n ] #p #_ + [ (lift_path_id p) +/2 width=1 by in_comp_lift_bi/ +qed-. + +lemma lift_term_id_dx (t): + ↑[𝐢]t ⊆ t. +#t #p * #q #Hq #H destruct // +qed-. + +lemma lift_term_id (t): + t ⇔ ↑[𝐢]t. +/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/ +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_proper.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_proper.etc new file mode 100644 index 000000000..a9b16219f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_proper.etc @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "delayed_updating/substitution/lift_prototerm.ma". +include "delayed_updating/substitution/lift_path_proper.ma". +include "delayed_updating/syntax/prototerm_proper.ma". + +(* LIFT FOR PROTOTERM *******************************************************) + +(* Constructions with proper condition for path *****************************) + +lemma lift_term_proper (f) (t): + t ϵ 𝐏 → ↑[f]t ϵ 𝐏. +#f #t #Ht #p * #q #Hq #H0 destruct +@lift_path_proper @Ht -Ht // (**) (* auto fails *) +qed. + +(* Inversions with proper condition for path ********************************) + +lemma lift_term_inv_proper (f) (t): + ↑[f]t ϵ 𝐏 → t ϵ 𝐏. +#f #t #Ht #p #Hp +@(lift_path_inv_proper f) +@Ht -Ht @in_comp_lift_bi // (**) (* auto fails *) +qed-. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift.etc deleted file mode 100644 index b873ed303..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift.etc +++ /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 "delayed_updating/notation/functions/uparrow_4.ma". -include "delayed_updating/notation/functions/uparrow_2.ma". -include "delayed_updating/syntax/path.ma". -include "ground/relocation/tr_uni.ma". -include "ground/relocation/tr_pap_tls.ma". - -(* LIFT FOR PATH ************************************************************) - -definition lift_continuation (A:Type[0]) ≝ - tr_map → path → A. - -rec definition lift_gen (A:Type[0]) (k:lift_continuation A) (f) (p) on p ≝ -match p with -[ list_empty ⇒ k f (𝐞) -| list_lcons l q ⇒ - match l with - [ label_d n ⇒ lift_gen (A) (λg,p. k g (𝗱(f@⧣❨n❩)◗p)) (⇂*[n]f) q - | label_m ⇒ lift_gen (A) (λg,p. k g (𝗺◗p)) f q - | label_L ⇒ lift_gen (A) (λg,p. k g (𝗟◗p)) (⫯f) q - | label_A ⇒ lift_gen (A) (λg,p. k g (𝗔◗p)) f q - | label_S ⇒ lift_gen (A) (λg,p. k g (𝗦◗p)) f q - ] -]. - -interpretation - "lift (gneric)" - 'UpArrow A k f p = (lift_gen A k f p). - -definition proj_path: lift_continuation … ≝ - λf,p.p. - -definition proj_rmap: lift_continuation … ≝ - λf,p.f. - -interpretation - "lift (path)" - 'UpArrow f p = (lift_gen ? proj_path f p). - -interpretation - "lift (relocation map)" - 'UpArrow p f = (lift_gen ? proj_rmap f p). - -(* Basic constructions ******************************************************) - -lemma lift_empty (A) (k) (f): - k f (𝐞) = ↑{A}❨k, f, 𝐞❩. -// qed. - -lemma lift_d_sn (A) (k) (p) (n) (f): - ↑❨(λg,p. k g (𝗱(f@⧣❨n❩)◗p)), ⇂*[n]f, p❩ = ↑{A}❨k, f, 𝗱n◗p❩. -// qed. - -lemma lift_m_sn (A) (k) (p) (f): - ↑❨(λg,p. k g (𝗺◗p)), f, p❩ = ↑{A}❨k, f, 𝗺◗p❩. -// qed. - -lemma lift_L_sn (A) (k) (p) (f): - ↑❨(λg,p. k g (𝗟◗p)), ⫯f, p❩ = ↑{A}❨k, f, 𝗟◗p❩. -// qed. - -lemma lift_A_sn (A) (k) (p) (f): - ↑❨(λg,p. k g (𝗔◗p)), f, p❩ = ↑{A}❨k, f, 𝗔◗p❩. -// qed. - -lemma lift_S_sn (A) (k) (p) (f): - ↑❨(λg,p. k g (𝗦◗p)), f, p❩ = ↑{A}❨k, f, 𝗦◗p❩. -// qed. - -(* Basic constructions with proj_path ***************************************) - -lemma lift_path_empty (f): - (𝐞) = ↑[f]𝐞. -// qed. - -(* Basic constructions with proj_rmap ***************************************) - -lemma lift_rmap_empty (f): - f = ↑[𝐞]f. -// qed. - -lemma lift_rmap_d_sn (f) (p) (n): - ↑[p](⇂*[ninj n]f) = ↑[𝗱n◗p]f. -// qed. - -lemma lift_rmap_m_sn (f) (p): - ↑[p]f = ↑[𝗺◗p]f. -// qed. - -lemma lift_rmap_L_sn (f) (p): - ↑[p](⫯f) = ↑[𝗟◗p]f. -// qed. - -lemma lift_rmap_A_sn (f) (p): - ↑[p]f = ↑[𝗔◗p]f. -// qed. - -lemma lift_rmap_S_sn (f) (p): - ↑[p]f = ↑[𝗦◗p]f. -// qed. - -(* Advanced constructions with proj_rmap and path_append ********************) - -lemma lift_rmap_append (p2) (p1) (f): - ↑[p2]↑[p1]f = ↑[p1●p2]f. -#p2 #p1 elim p1 -p1 // * [ #n ] #p1 #IH #f // -[ nsucc_pnpred lift_lcons_alt // >lift_append_rcons_sn // -lift_lcons_alt nrplus_inj_dx -/2 width=1 by tr_tls_compose_uni_dx/ -qed. - -*) diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_path_after.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_path_after.etc deleted file mode 100644 index 869d680fe..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_path_after.etc +++ /dev/null @@ -1,31 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "delayed_updating/substitution/lift_eq.ma". -include "ground/relocation/tr_compose_pap.ma". -include "ground/relocation/tr_compose_pn.ma". -include "ground/relocation/tr_compose_tls.ma". - -(* LIFT FOR PATH ************************************************************) - -(* Constructions with tr_after **********************************************) - -lemma lift_path_after (p) (f1) (f2): - ↑[f2]↑[f1]p = ↑[f2∘f1]p. -#p elim p -p [| * ] // [ #n ] #p #IH #f1 #f2 -[ tr_compose_push_bi // -] -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_path_id.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_path_id.etc deleted file mode 100644 index 1966c25f2..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_path_id.etc +++ /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 "delayed_updating/substitution/lift_eq.ma". -include "ground/relocation/tr_id_pap.ma". -include "ground/relocation/tr_id_tls.ma". - -(* LIFT FOR PATH ************************************************************) - -(* Constructions with tr_id *************************************************) - -lemma lift_path_id (p): - p = ↑[𝐢]p. -#p elim p -p // -* [ #n ] #p #IH // -[ nsucc_pnpred -(lift_path_id p) -/2 width=1 by in_comp_lift_bi/ -qed-. - -lemma lift_term_id_dx (t): - ↑[𝐢]t ⊆ t. -#t #p * #q #Hq #H destruct // -qed-. - -lemma lift_term_id (t): - t ⇔ ↑[𝐢]t. -/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/ -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_prototerm_proper.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_prototerm_proper.etc deleted file mode 100644 index a9b16219f..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_k/lift_prototerm_proper.etc +++ /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 "delayed_updating/substitution/lift_prototerm.ma". -include "delayed_updating/substitution/lift_path_proper.ma". -include "delayed_updating/syntax/prototerm_proper.ma". - -(* LIFT FOR PROTOTERM *******************************************************) - -(* Constructions with proper condition for path *****************************) - -lemma lift_term_proper (f) (t): - t ϵ 𝐏 → ↑[f]t ϵ 𝐏. -#f #t #Ht #p * #q #Hq #H0 destruct -@lift_path_proper @Ht -Ht // (**) (* auto fails *) -qed. - -(* Inversions with proper condition for path ********************************) - -lemma lift_term_inv_proper (f) (t): - ↑[f]t ϵ 𝐏 → t ϵ 𝐏. -#f #t #Ht #p #Hp -@(lift_path_inv_proper f) -@Ht -Ht @in_comp_lift_bi // (**) (* auto fails *) -qed-. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/notation/relations/black_rightarrow_f_4.ma b/matita/matita/contribs/lambdadelta/delayed_updating/notation/relations/black_rightarrow_f_4.ma deleted file mode 100644 index c69d6c79b..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/notation/relations/black_rightarrow_f_4.ma +++ /dev/null @@ -1,19 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR DELAYED UPDATING ********************************************) - -notation "hvbox( t1 ➡𝐟[ break term 46 p, break term 46 q ] break term 46 t2 )" - non associative with precedence 45 - for @{ 'BlackRightArrowF $t1 $p $q $t2 }. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/notation/relations/black_rightarrow_if_4.ma b/matita/matita/contribs/lambdadelta/delayed_updating/notation/relations/black_rightarrow_if_4.ma new file mode 100644 index 000000000..2b35b91ea --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/notation/relations/black_rightarrow_if_4.ma @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR DELAYED UPDATING ********************************************) + +notation "hvbox( t1 ➡𝐢𝐟[ break term 46 p, break term 46 q ] break term 46 t2 )" + non associative with precedence 45 + for @{ 'BlackRightArrowIF $t1 $p $q $t2 }. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_ifr.ma b/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_ifr.ma index be7f3b3d3..fd5f76b54 100644 --- a/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_ifr.ma +++ b/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_ifr.ma @@ -33,7 +33,7 @@ include "delayed_updating/syntax/path_depth_reverse.ma". (* Main destructions with ifr ***********************************************) theorem dfr_des_ifr (f) (p) (q) (t1) (t2): t1 ϵ 𝐓 → - t1 ➡𝐝𝐟[p,q] t2 → ▼[f]t1 ➡𝐟[⊗p,⊗q] ▼[f]t2. + t1 ➡𝐝𝐟[p,q] t2 → ▼[f]t1 ➡𝐢𝐟[⊗p,⊗q] ▼[f]t2. #f #p #q #t1 #t2 #H0t1 * #n * #H1n #Ht1 #Ht2 @(ex_intro … (↑♭q)) @and3_intro diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_lift.ma b/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_lift.ma index d25ddc502..e33b454be 100644 --- a/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_lift.ma +++ b/matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_lift.ma @@ -23,12 +23,23 @@ include "delayed_updating/unwind/unwind2_rmap_head.ma". *) include "delayed_updating/substitution/fsubst_eq.ma". include "delayed_updating/substitution/lift_prototerm_eq.ma". +include "delayed_updating/substitution/lift_path_head.ma". include "delayed_updating/syntax/prototerm_proper_constructors.ma". + (* DELAYED FOCUSED REDUCTION ************************************************) (* Constructions with lift **************************************************) +(* +lemma pippo (f) (r): + ↑[↑[r]f](rᴿ) = (↑[f]r)ᴿ. +#f #r @(list_ind_rcons … r) -r // +#p * [ #n ] #IH +[ H1n >path_head_structure_depth lift_path_L_sn + >list_append_rcons_sn in H1n; list_append_rcons_sn in H1n; nsucc_pnpred nrplus_inj_dx /2 width=1 by tr_tls_compose_uni_dx/ qed. + *) diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_after.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_after.ma index 938d21534..869d680fe 100644 --- a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_after.ma +++ b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_after.ma @@ -15,6 +15,7 @@ include "delayed_updating/substitution/lift_eq.ma". include "ground/relocation/tr_compose_pap.ma". include "ground/relocation/tr_compose_pn.ma". +include "ground/relocation/tr_compose_tls.ma". (* LIFT FOR PATH ************************************************************) @@ -22,7 +23,9 @@ include "ground/relocation/tr_compose_pn.ma". lemma lift_path_after (p) (f1) (f2): ↑[f2]↑[f1]p = ↑[f2∘f1]p. -#p elim p -p [| * ] // #p #IH #f1 #f2 -tr_compose_push_bi // +#p elim p -p [| * ] // [ #n ] #p #IH #f1 #f2 +[ tr_compose_push_bi // +] qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_head.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_head.ma new file mode 100644 index 000000000..126b691c6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_head.ma @@ -0,0 +1,31 @@ +include "delayed_updating/substitution/lift_eq.ma". +include "delayed_updating/syntax/path_head.ma". +include "delayed_updating/syntax/path_reverse.ma". +include "ground/relocation/xap.ma". + +axiom tr_xap_succ_pos (f) (n): + ↑↓(f@❨↑n❩) = f@❨↑n❩. + +axiom tr_xap_plus (n1) (n2) (f): + (⇂*[n2]f)@❨n1❩+f@❨n2❩ = f@❨n1+n2❩. + +axiom eq_inv_path_empty_head (p) (n): + (𝐞) = ↳[n]p → 𝟎 = n. + +lemma lift_path_head (f) (p) (q) (n): + pᴿ = ↳[n](pᴿ●qᴿ) → + ↳[↑[q●p]f@❨n❩](↑[f](q●p))ᴿ = (↑[↑[q]f]p)ᴿ. +#f #p @(list_ind_rcons … p) -p +[ #q #n #H0 + <(eq_inv_path_empty_head … H0) -H0 + tr_xap_succ_pos + lift_rmap_append nsucc_pnpred +(lift_path_id p) -/2 width=1 by in_comp_lift_bi/ -qed-. - -lemma lift_term_id_dx (t): - ↑[𝐢]t ⊆ t. -#t #p * #q #Hq #H destruct // -qed-. - -lemma lift_term_id (t): - t ⇔ ↑[𝐢]t. -/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/ -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_prototerm_id.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_prototerm_id.ma new file mode 100644 index 000000000..b66b63f46 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_prototerm_id.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 "delayed_updating/substitution/lift_prototerm_eq.ma". +include "delayed_updating/substitution/lift_path_id.ma". + +(* LIFT FOR PROTOTERM *******************************************************) + +(* Constructions with tr_id *************************************************) + +lemma lift_term_id_sn (t): + t ⊆ ↑[𝐢]t. +#t #p #Hp +>(lift_path_id p) +/2 width=1 by in_comp_lift_bi/ +qed-. + +lemma lift_term_id_dx (t): + ↑[𝐢]t ⊆ t. +#t #p * #q #Hq #H destruct // +qed-. + +lemma lift_term_id (t): + t ⇔ ↑[𝐢]t. +/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/ +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_reverse.ma b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_reverse.ma index 046404f03..7fa03a41e 100644 --- a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_reverse.ma +++ b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_reverse.ma @@ -54,7 +54,7 @@ qed. (* Main constructions *******************************************************) -theorem reverse_revrse (p): +theorem reverse_reverse (p): p = pᴿᴿ. #p elim p -p // #l #p #IH