From: Ferruccio Guidi Date: Wed, 27 Apr 2022 17:24:30 +0000 (+0200) Subject: update un delayed updating X-Git-Tag: make_still_working~76^2~5 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=e6482f7f499fbb88aff8a205d2a2c808158042c5;p=helm.git update un delayed updating + some parked files --- diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/relocation/sbr_pap.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/relocation/sbr_pap.etc new file mode 100644 index 000000000..c35244d88 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/relocation/sbr_pap.etc @@ -0,0 +1,48 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM 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/arith/pnat_plus.ma". +include "ground/relocation/tr_map.ma". +include "delayed_updating/notation/functions/verticalbarat_2.ma". + +(* POSITIVE APPLICATION FOR STACKS OF BASIC RELOCATION MAPS *****************) + +rec definition sbr_pap (i: pnat) on i: tr_map → pnat. +* #p #f cases i -i +[ @(p) +| #i cases p -p + [ lapply (sbr_pap i f) -sbr_pap -i -f + #i @(↑i) + | #p @(i+↑p) + ] +] +defined. + +interpretation + "functional positive application (stack of relocation maps)" + 'VerticalBarAt f i = (sbr_pap i f). + +(* Basic constructions ******************************************************) + +lemma sbr_pap_unfold_unit (f): + ∀p. p = (p⨮f)❘@❨𝟏❩. +// qed. + +lemma sbr_pap_unfold_unit_succ (f): + ∀i. ↑(f❘@❨i❩) = (𝟏⨮f)❘@❨↑i❩. +// qed. + +lemma sbr_pap_unfold_succ_bi (f): + ∀p,i. i+↑p = (↑p⨮f)❘@❨↑i❩. +// qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/relocation/sbr_pap_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/relocation/sbr_pap_eq.etc new file mode 100644 index 000000000..10b0d939e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/relocation/sbr_pap_eq.etc @@ -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 "delayed_updating/relocation/sbr_pap.ma". +include "ground/relocation/tr_eq.ma". + +(* POSITIVE APPLICATION FOR STACKS OF BASIC RELOCATION MAPS *****************) + +(* Main constructions with stream_eq ****************************************) + +theorem sbr_pap_eq_repl_sn (i): + stream_eq_repl … (λf1,f2. f1❘@❨i❩ = f2❘@❨i❩). +#i elim i -i [| #i #IH ] +* [ #p1 | * [| #p1 ]] #f1 +* #p2 #f2 #H +elim (stream_eq_inv_cons_bi … H) -H [1,8,15: |*: // ] #Hp #Hf destruct // +(tr_pap_eq_repl … (▼[p]f) … (unwind_rmap_decompose …)) +unwind_lcons_alt // >unwind_append_rcons_sn // + unwind_lcons_alt >unwind_append_rcons_sn // + unwind_lcons_alt >unwind_append_rcons_sn // + unwind_lcons_alt (unwind_path_eq_repl … (tr_compose_assoc …)) // +| tr_compose_push_bi // +] +qed. + +(* Advanced constructions with proj_rmap and stream_tls *********************) + +lemma unwind_rmap_tls_d_dx (f) (p) (m) (n): + ⇂*[m+n]▼[p]f ≗ ⇂*[m]▼[p◖𝗱n]f. +#f #p #m #n +nrplus_inj_dx +/2 width=1 by tr_tls_compose_uni_dx/ +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_etc.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_etc.etc new file mode 100644 index 000000000..a9e563db6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_etc.etc @@ -0,0 +1,10 @@ +include "delayed_updating/unwind2/unwind_depth.ma". +include "ground/relocation/tr_uni_compose.ma". + +lemma pippo (q) (p): + ▼[p]𝐢 ≗ ⇂*[↑❘q❘]▼[p●𝗟◗q]𝐢. +#q @(list_ind_rcons … q) -q // +#q * [ #n ] #IH #p // +list_cons_shift unwind_append_proper_dx + /4 width=5 by in_comp_unwind_bi, in_ppc_comp_trans, or_introl, ex2_intro/ +| * #q #Hq #H1 #H0 + @(ex2_intro … H1) @or_intror @conj // * + [ nsucc_inj #H0 + IH -IH // +| // +| nsucc_inj // +| // +| // +] +qed. + +lemma unwind_rmap_tls_eq (f) (p): + ⇂*[⧣p]f ≗ ⇂*[↑❘p❘]▼[p]⫯f. +(* +#f #p @(list_ind_rcons … p) -p // +#p * [ #n ] #IH // +unwind_rmap_S_dx >structure_S_dx + @unwind_grafted_sn // +| /2 width=1 by unwind_grafted_S_dx/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm.etc new file mode 100644 index 000000000..1a5d46b77 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm.etc @@ -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 "delayed_updating/unwind2/unwind.ma". +include "delayed_updating/syntax/prototerm.ma". +include "ground/lib/subset_ext.ma". + +(* UNWIND FOR PROTOTERM *****************************************************) + +interpretation + "unwind (prototerm)" + 'BlackDownTriangle f t = (subset_ext_f1 ? ? (unwind_gen ? proj_path f) t). + +(* Basic constructions ******************************************************) + +lemma in_comp_unwind_bi (f) (p) (t): + p ϵ t → ▼[f]p ϵ ▼[f]t. +/2 width=1 by subset_in_ext_f1_dx/ +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm_eq.etc new file mode 100644 index 000000000..15b8cdb2c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm_eq.etc @@ -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/lib/subset_ext_equivalence.ma". +include "delayed_updating/unwind2/unwind_eq.ma". +include "delayed_updating/unwind2/unwind_prototerm.ma". + +(* UNWIND FOR PROTOTERM *****************************************************) + +(* Constructions with subset_equivalence ************************************) + +lemma unwind_term_eq_repl_sn (f1) (f2) (t): + f1 ≗ f2 → ▼[f1]t ⇔ ▼[f2]t. +/3 width=1 by subset_equivalence_ext_f1_exteq, unwind_path_eq_repl/ +qed. + +lemma unwind_term_eq_repl_dx (f) (t1) (t2): + t1 ⇔ t2 → ▼[f]t1 ⇔ ▼[f]t2. +/2 width=1 by subset_equivalence_ext_f1_bi/ +qed. + +lemma unwind_term_after (f1) (f2) (t): + ▼[f2]▼[f1]t ⇔ ▼[f2∘f1]t. +#f1 #f2 #t @subset_eq_trans +[ +| @subset_inclusion_ext_f1_compose +| @subset_equivalence_ext_f1_exteq /2 width=5/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure.etc new file mode 100644 index 000000000..b5d92a577 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure.etc @@ -0,0 +1,269 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM 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/unwind2/unwind_eq.ma". +include "delayed_updating/syntax/path_structure.ma". +include "delayed_updating/syntax/path_inner.ma". +include "delayed_updating/syntax/path_proper.ma". +include "ground/xoa/ex_4_2.ma". + +(* UNWIND FOR PATH **********************************************************) + +(* Basic constructions with structure ***************************************) + +lemma structure_unwind (p) (f): + ⊗p = ⊗▼[f]p. +#p @(path_ind_unwind … p) -p // #p #IH #f +list_cons_shift unwind_lcons_alt // >unwind_append_rcons_sn // + unwind_lcons_alt >unwind_append_rcons_sn // + unwind_lcons_alt >unwind_append_rcons_sn // + unwind_lcons_alt unwind_append_proper_dx + /4 width=5 by in_comp_unwind_bi, in_ppc_comp_trans, or_introl, ex2_intro/ +| * #q #Hq #H1 #H0 + @(ex2_intro … H1) @or_intror @conj // * + [ list_cons_shift list_cons_shift unwind_rmap_S_dx >structure_S_dx + @unwind_grafted_sn // +| /2 width=1 by unwind_grafted_S_dx/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm.etc new file mode 100644 index 000000000..c4a103b61 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm.etc @@ -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 "delayed_updating/unwind3/unwind.ma". +include "delayed_updating/syntax/prototerm.ma". +include "ground/lib/subset_ext.ma". + +(* UNWIND FOR PROTOTERM *****************************************************) + +interpretation + "unwind (prototerm)" + 'BlackDownTriangle f t = (subset_ext_f1 ? ? (unwind_gen ? proj_path f) t). + +(* Basic constructions ******************************************************) + +lemma in_comp_unwind_bi (f) (p) (t): + p ϵ t → ▼[f]p ϵ ▼[f]t. +/2 width=1 by subset_in_ext_f1_dx/ +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm_eq.etc new file mode 100644 index 000000000..01d302f3e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm_eq.etc @@ -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 "ground/lib/subset_ext_equivalence.ma". +include "delayed_updating/unwind3/unwind_eq.ma". +include "delayed_updating/unwind3/unwind_prototerm.ma". + +(* UNWIND FOR PROTOTERM *****************************************************) + +(* Constructions with subset_equivalence ************************************) + +lemma unwind_term_eq_repl_sn (f1) (f2) (t): + f1 ≗ f2 → ▼[f1]t ⇔ ▼[f2]t. +/3 width=1 by subset_equivalence_ext_f1_exteq, unwind_path_eq_repl/ +qed. + +lemma unwind_term_eq_repl_dx (f) (t1) (t2): + t1 ⇔ t2 → ▼[f]t1 ⇔ ▼[f]t2. +/2 width=1 by subset_equivalence_ext_f1_bi/ +qed. + +lemma unwind_term_after_id_sn (f) (t): + ▼[𝐢]▼[f]t ⇔ ▼[f]t. +#f #t @subset_eq_trans +[ +| @subset_inclusion_ext_f1_compose +| @subset_equivalence_ext_f1_exteq + #p @unwind_path_after_id_sn +] +qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure.etc new file mode 100644 index 000000000..c52a43871 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure.etc @@ -0,0 +1,269 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM 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/unwind3/unwind_eq.ma". +include "delayed_updating/syntax/path_structure.ma". +include "delayed_updating/syntax/path_inner.ma". +include "delayed_updating/syntax/path_proper.ma". +include "ground/xoa/ex_4_2.ma". + +(* UNWIND FOR PATH *********************************************************) + +(* Basic constructions with structure **************************************) + +lemma structure_unwind (p) (f): + ⊗p = ⊗▼[f]p. +#p @(path_ind_unwind … p) -p // #p #IH #f +(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/unwind2/unwind.ma b/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind.ma deleted file mode 100644 index 1713625e8..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind.ma +++ /dev/null @@ -1,189 +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/relocation/tr_compose_pap.ma". -include "ground/relocation/tr_uni_pap.ma". -include "delayed_updating/syntax/path.ma". -include "delayed_updating/notation/functions/black_downtriangle_4.ma". -include "delayed_updating/notation/functions/black_downtriangle_2.ma". - -(* UNWIND FOR PATH **********************************************************) - -definition unwind_continuation (A:Type[0]) ≝ -tr_map → path → A. - -rec definition unwind_gen (A:Type[0]) (k:unwind_continuation A) (f) (p) on p ≝ -match p with -[ list_empty ⇒ k f (𝐞) -| list_lcons l q ⇒ - match l with - [ label_d n ⇒ - match q with - [ list_empty ⇒ unwind_gen (A) (λg,p. k g (𝗱(f@❨n❩)◗p)) (f∘𝐮❨n❩) q - | list_lcons _ _ ⇒ unwind_gen (A) k (f∘𝐮❨n❩) q - ] - | label_m ⇒ unwind_gen (A) k f q - | label_L ⇒ unwind_gen (A) (λg,p. k g (𝗟◗p)) (⫯f) q - | label_A ⇒ unwind_gen (A) (λg,p. k g (𝗔◗p)) f q - | label_S ⇒ unwind_gen (A) (λg,p. k g (𝗦◗p)) f q - ] -]. - -interpretation - "unwind (gneric)" - 'BlackDownTriangle A k f p = (unwind_gen A k f p). - -definition proj_path: unwind_continuation … ≝ - λf,p.p. - -definition proj_rmap: unwind_continuation … ≝ - λf,p.f. - -interpretation - "unwind (path)" - 'BlackDownTriangle f p = (unwind_gen ? proj_path f p). - -interpretation - "unwind (relocation map)" - 'BlackDownTriangle p f = (unwind_gen ? proj_rmap f p). - -(* Basic constructions ******************************************************) - -lemma unwind_empty (A) (k) (f): - k f (𝐞) = ▼{A}❨k, f, 𝐞❩. -// qed. - -lemma unwind_d_empty_sn (A) (k) (n) (f): - ▼❨(λg,p. k g (𝗱(f@❨n❩)◗p)), f∘𝐮❨ninj n❩, 𝐞❩ = ▼{A}❨k, f, -𝗱n◗𝐞❩. -// qed. - -lemma unwind_d_lcons_sn (A) (k) (p) (l) (n) (f): - ▼❨k, f∘𝐮❨ninj n❩, l◗p❩ = ▼{A}❨k, f, 𝗱n◗l◗p❩. -// qed. - -lemma unwind_m_sn (A) (k) (p) (f): - ▼❨k, f, p❩ = ▼{A}❨k, f, 𝗺◗p❩. -// qed. - -lemma unwind_L_sn (A) (k) (p) (f): - ▼❨(λg,p. k g (𝗟◗p)), ⫯f, p❩ = ▼{A}❨k, f, 𝗟◗p❩. -// qed. - -lemma unwind_A_sn (A) (k) (p) (f): - ▼❨(λg,p. k g (𝗔◗p)), f, p❩ = ▼{A}❨k, f, 𝗔◗p❩. -// qed. - -lemma unwind_S_sn (A) (k) (p) (f): - ▼❨(λg,p. k g (𝗦◗p)), f, p❩ = ▼{A}❨k, f, 𝗦◗p❩. -// qed. - -(* Basic constructions with proj_path ***************************************) - -lemma unwind_path_empty (f): - (𝐞) = ▼[f]𝐞. -// qed. - -lemma unwind_path_d_empty_sn (f) (n): - 𝗱(f@❨n❩)◗𝐞 = ▼[f](𝗱n◗𝐞). -// qed. - -lemma unwind_path_d_lcons_sn (f) (p) (l) (n): - ▼[f∘𝐮❨ninj n❩](l◗p) = ▼[f](𝗱n◗l◗p). -// qed. - -lemma unwind_path_m_sn (f) (p): - ▼[f]p = ▼[f](𝗺◗p). -// qed. - -(* Basic constructions with proj_rmap ***************************************) - -lemma unwind_rmap_empty (f): - f = ▼[𝐞]f. -// qed. - -lemma unwind_rmap_d_sn (f) (p) (n): - ▼[p](f∘𝐮❨ninj n❩) = ▼[𝗱n◗p]f. -#f * // qed. - -lemma unwind_rmap_m_sn (f) (p): - ▼[p]f = ▼[𝗺◗p]f. -// qed. - -lemma unwind_rmap_L_sn (f) (p): - ▼[p](⫯f) = ▼[𝗟◗p]f. -// qed. - -lemma unwind_rmap_A_sn (f) (p): - ▼[p]f = ▼[𝗔◗p]f. -// qed. - -lemma unwind_rmap_S_sn (f) (p): - ▼[p]f = ▼[𝗦◗p]f. -// qed. - -(* Advanced constructions with proj_rmap and path_append ********************) - -lemma unwind_rmap_append (p2) (p1) (f): - ▼[p2]▼[p1]f = ▼[p1●p2]f. -#p2 #p1 elim p1 -p1 // * [ #n ] #p1 #IH #f // -[ (tr_pap_eq_repl … (▼[p]f) … (unwind_rmap_decompose …)) -unwind_lcons_alt // >unwind_append_rcons_sn // - unwind_lcons_alt >unwind_append_rcons_sn // - unwind_lcons_alt >unwind_append_rcons_sn // - unwind_lcons_alt (unwind_path_eq_repl … (tr_compose_assoc …)) // -| tr_compose_push_bi // -] -qed. - -(* Advanced constructions with proj_rmap and stream_tls *********************) - -lemma unwind_rmap_tls_d_dx (f) (p) (m) (n): - ⇂*[m+n]▼[p]f ≗ ⇂*[m]▼[p◖𝗱n]f. -#f #p #m #n -nrplus_inj_dx -/2 width=1 by tr_tls_compose_uni_dx/ -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_fsubst.ma b/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_fsubst.ma deleted file mode 100644 index fedc665fb..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_fsubst.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 "delayed_updating/unwind2/unwind_prototerm_eq.ma". -include "delayed_updating/unwind2/unwind_structure.ma". -include "delayed_updating/substitution/fsubst.ma". -include "delayed_updating/syntax/preterm.ma". -include "delayed_updating/syntax/prototerm_proper.ma". - -(* UNWIND FOR PROTOTERM *****************************************************) - -(* Constructions with fsubst ************************************************) - -lemma unwind_fsubst_sn (f) (t) (u) (p): u ϵ 𝐏 → - (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u] ⊆ ▼[f](t[⋔p←u]). -#f #t #u #p #Hu #ql * * -[ #rl * #r #Hr #H1 #H2 destruct - >unwind_append_proper_dx - /4 width=5 by in_comp_unwind_bi, in_ppc_comp_trans, or_introl, ex2_intro/ -| * #q #Hq #H1 #H0 - @(ex2_intro … H1) @or_intror @conj // * - [ nsucc_inj #H0 - IH -IH // -| // -| nsucc_inj // -| // -| // -] -qed. - -lemma unwind_rmap_tls_eq (f) (p): - ⇂*[⧣p]f ≗ ⇂*[↑❘p❘]▼[p]⫯f. -(* -#f #p @(list_ind_rcons … p) -p // -#p * [ #n ] #IH // -unwind_rmap_S_dx >structure_S_dx - @unwind_grafted_sn // -| /2 width=1 by unwind_grafted_S_dx/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_prototerm.ma b/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_prototerm.ma deleted file mode 100644 index 1a5d46b77..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_prototerm.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 "delayed_updating/unwind2/unwind.ma". -include "delayed_updating/syntax/prototerm.ma". -include "ground/lib/subset_ext.ma". - -(* UNWIND FOR PROTOTERM *****************************************************) - -interpretation - "unwind (prototerm)" - 'BlackDownTriangle f t = (subset_ext_f1 ? ? (unwind_gen ? proj_path f) t). - -(* Basic constructions ******************************************************) - -lemma in_comp_unwind_bi (f) (p) (t): - p ϵ t → ▼[f]p ϵ ▼[f]t. -/2 width=1 by subset_in_ext_f1_dx/ -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_prototerm_eq.ma b/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_prototerm_eq.ma deleted file mode 100644 index 15b8cdb2c..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_prototerm_eq.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/lib/subset_ext_equivalence.ma". -include "delayed_updating/unwind2/unwind_eq.ma". -include "delayed_updating/unwind2/unwind_prototerm.ma". - -(* UNWIND FOR PROTOTERM *****************************************************) - -(* Constructions with subset_equivalence ************************************) - -lemma unwind_term_eq_repl_sn (f1) (f2) (t): - f1 ≗ f2 → ▼[f1]t ⇔ ▼[f2]t. -/3 width=1 by subset_equivalence_ext_f1_exteq, unwind_path_eq_repl/ -qed. - -lemma unwind_term_eq_repl_dx (f) (t1) (t2): - t1 ⇔ t2 → ▼[f]t1 ⇔ ▼[f]t2. -/2 width=1 by subset_equivalence_ext_f1_bi/ -qed. - -lemma unwind_term_after (f1) (f2) (t): - ▼[f2]▼[f1]t ⇔ ▼[f2∘f1]t. -#f1 #f2 #t @subset_eq_trans -[ -| @subset_inclusion_ext_f1_compose -| @subset_equivalence_ext_f1_exteq /2 width=5/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_structure.ma b/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_structure.ma deleted file mode 100644 index b5d92a577..000000000 --- a/matita/matita/contribs/lambdadelta/delayed_updating/unwind2/unwind_structure.ma +++ /dev/null @@ -1,269 +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/unwind2/unwind_eq.ma". -include "delayed_updating/syntax/path_structure.ma". -include "delayed_updating/syntax/path_inner.ma". -include "delayed_updating/syntax/path_proper.ma". -include "ground/xoa/ex_4_2.ma". - -(* UNWIND FOR PATH **********************************************************) - -(* Basic constructions with structure ***************************************) - -lemma structure_unwind (p) (f): - ⊗p = ⊗▼[f]p. -#p @(path_ind_unwind … p) -p // #p #IH #f -