]> matita.cs.unibo.it Git - helm.git/commitdiff
update in delayed_updating
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Sun, 7 Aug 2022 17:36:15 +0000 (19:36 +0200)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Sun, 7 Aug 2022 17:36:15 +0000 (19:36 +0200)
+ three old lemmas restored
+ some parked lemmas removed
+ minor renaming

29 files changed:
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_constructors.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_depth.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_eq_etc.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_fsubst.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_preterm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_prototerm.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_prototerm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_structure.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_structure_depth.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_depth.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_etc.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure_depth.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_constructors.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_fsubst.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_lift.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_preterm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_prototerm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure_depth.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/unwind/unwind2_rmap_depth.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/delayed_updating/unwind/unwind2_rmap_head.ma

diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind.etc
deleted file mode 100644 (file)
index 3e76dee..0000000
+++ /dev/null
@@ -1,187 +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_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 (A) (k) (n) (f):
-      ▼❨(λg,p. k g (𝗱(f@⧣❨n❩)◗p)), 𝐮❨f@⧣❨n❩❩, 𝐞❩ = ▼{A}❨k, f, 𝗱n◗𝐞❩.
-// qed.
-
-lemma unwind_d_lcons (A) (k) (p) (l) (n) (f):
-      ▼❨k, 𝐮❨f@⧣❨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 (f) (n):
-      𝗱(f@⧣❨n❩)◗𝐞 = ▼[f](𝗱n◗𝐞).
-// qed.
-
-lemma unwind_path_d_lcons (f) (p) (l) (n):
-      ▼[𝐮❨f@⧣❨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@⧣❨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 //
-[ <unwind_rmap_m_sn <unwind_rmap_m_sn //
-| <unwind_rmap_A_sn <unwind_rmap_A_sn //
-| <unwind_rmap_S_sn <unwind_rmap_S_sn //
-]
-qed.
-
-(* Advanced constructions with proj_rmap and path_rcons *********************)
-
-lemma unwind_rmap_d_dx (f) (p) (n):
-      (𝐮❨(▼[p]f)@⧣❨n❩❩) = ▼[p◖𝗱n]f.
-// qed.
-
-lemma unwind_rmap_m_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗺]f.
-// qed.
-
-lemma unwind_rmap_L_dx (f) (p):
-      (⫯▼[p]f) = ▼[p◖𝗟]f.
-// qed.
-
-lemma unwind_rmap_A_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗔]f.
-// qed.
-
-lemma unwind_rmap_S_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗦]f.
-// qed.
-
-lemma unwind_rmap_pap_d_dx (f) (p) (n) (m):
-      ▼[p]f@⧣❨n❩+m = ▼[p◖𝗱n]f@⧣❨m❩.
-#f #p #n #m
-<unwind_rmap_d_dx <tr_uni_pap <pplus_comm //
-qed.
-
-(* Advanced eliminations with path ******************************************)
-
-lemma path_ind_unwind (Q:predicate …):
-      Q (𝐞) →
-      (∀n. Q (𝐞) → Q (𝗱n◗𝐞)) →
-      (∀n,l,p. Q (l◗p) → Q (𝗱n◗l◗p)) →
-      (∀p. Q p → Q (𝗺◗p)) →
-      (∀p. Q p → Q (𝗟◗p)) →
-      (∀p. Q p → Q (𝗔◗p)) →
-      (∀p. Q p → Q (𝗦◗p)) →
-      ∀p. Q p.
-#Q #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #p
-elim p -p [| * [ #n * ] ]
-/2 width=1 by/
-qed-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_constructors.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_constructors.etc
deleted file mode 100644 (file)
index a412797..0000000
+++ /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 "delayed_updating/unwind1/unwind_prototerm_eq.ma".
-include "delayed_updating/syntax/prototerm_constructors.ma".
-
-(* UNWIND FOR PROTOTERM *****************************************************)
-
-lemma unwind_iref_sn (f) (t:prototerm) (n:pnat):
-      ▼[𝐮❨f@⧣❨n❩❩]t ⊆ ▼[f](𝛗n.t).
-#f #t #n #p * #q #Hq #H0 destruct
-@(ex2_intro … (𝗱n◗𝗺◗q))
-/2 width=1 by in_comp_iref/
-qed-.
-
-lemma unwind_iref_dx (f) (t) (n:pnat):
-      ▼[f](𝛗n.t) ⊆ ▼[𝐮❨f@⧣❨n❩❩]t.
-#f #t #n #p * #q #Hq #H0 destruct
-elim (in_comp_inv_iref … Hq) -Hq #p #Hp #Ht destruct
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_iref (f) (t) (n:pnat):
-      ▼[𝐮❨f@⧣❨n❩❩]t ⇔ ▼[f](𝛗n.t).
-/3 width=1 by conj, unwind_iref_sn, unwind_iref_dx/
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_depth.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_depth.etc
deleted file mode 100644 (file)
index 05e2d52..0000000
+++ /dev/null
@@ -1,49 +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/unwind1/unwind.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/relocation/tr_id_compose.ma".
-include "ground/relocation/tr_compose_compose.ma".
-include "ground/relocation/tr_compose_pn.ma".
-include "ground/relocation/tr_compose_eq.ma".
-include "ground/relocation/tr_pn_eq.ma".
-include "ground/lib/stream_eq_eq.ma".
-
-(* UNWIND FOR PATH *********************************************************)
-(* COMMENT
-(* Constructions with depth ************************************************)
-
-lemma unwind_rmap_decompose (p) (f):
-      ▼[p]f ≗ (▼[p]𝐢)∘(⫯*[❘p❘]f).
-#p @(list_ind_rcons … p) -p
-[ #f <unwind_rmap_empty <unwind_rmap_empty <tr_pushs_zero //
-| #p * [ #n ] #IH #f //
-  [ <unwind_rmap_d_dx <unwind_rmap_d_dx <depth_d_dx
-    @(stream_eq_canc_dx … (tr_compose_assoc …))
-    /2 width=1 by tr_compose_eq_repl/
-  | <unwind_rmap_L_dx <unwind_rmap_L_dx <depth_L_dx
-    <tr_pushs_succ <tr_compose_push_bi
-    /2 width=1 by tr_push_eq_repl/
-  ]
-]
-qed.
-
-lemma unwind_rmap_pap_le (f) (p) (n):
-      n < ▼❘p❘ → (▼[p]𝐢)@⧣❨n❩ = (▼[p]f)@⧣❨n❩.
-#f #p #n #Hn
->(tr_pap_eq_repl … (▼[p]f) … (unwind_rmap_decompose …))
-<tr_compose_pap <tr_pap_pushs_le //
-qed.
-*)
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_eq.etc
deleted file mode 100644 (file)
index 8943b73..0000000
+++ /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 "delayed_updating/unwind1/unwind.ma".
-include "ground/relocation/tr_uni_eq.ma".
-include "ground/relocation/tr_pap_eq.ma".
-include "ground/relocation/tr_pn_eq.ma".
-
-(* UNWIND FOR PATH **********************************************************)
-
-definition unwind_exteq (A): relation2 (unwind_continuation A) (unwind_continuation A) ≝
-           λk1,k2. ∀f1,f2,p. f1 ≗ f2 → k1 f1 p = k2 f2 p.
-
-interpretation
-  "extensional equivalence (unwind continuation)"
-  'RingEq A k1 k2 = (unwind_exteq A k1 k2).
-
-(* Constructions with unwind_exteq ******************************************)
-
-lemma unwind_eq_repl (A) (p) (k1) (k2):
-      k1 ≗{A} k2 → stream_eq_repl … (λf1,f2. ▼❨k1, f1, p❩ = ▼❨k2, f2, p❩).
-#A #p @(path_ind_unwind … p) -p [| #n #IH | #n #l0 #q #IH |*: #q #IH ]
-#k1 #k2 #Hk #f1 #f2 #Hf
-[ <unwind_empty <unwind_empty /2 width=1 by/
-| <unwind_d_empty <unwind_d_empty <(tr_pap_eq_repl … Hf)
-  /2 width=1 by stream_eq_refl/
-| <unwind_d_lcons <unwind_d_lcons
-  /5 width=1 by tr_uni_eq_repl, tr_pap_eq_repl, eq_f/ 
-| /2 width=1 by/
-| /3 width=1 by tr_push_eq_repl/
-| /3 width=1 by/
-| /3 width=1 by/
-]
-qed-.
-
-(* Advanced constructions ***************************************************)
-
-lemma unwind_lcons_alt (A) (k) (f) (p) (l): k ≗ k →
-      ▼❨λg,p2. k g (l◗p2), f, p❩ = ▼{A}❨λg,p2. k g ((l◗𝐞)●p2), f, p❩.
-#A #k #f #p #l #Hk
-@unwind_eq_repl // #g1 #g2 #p2 #Hg @Hk -Hk // (**) (* auto fail *)
-qed.
-
-lemma unwind_append_rcons_sn (A) (k) (f) (p1) (p) (l): k ≗ k →
-      ▼❨λg,p2. k g (p1●l◗p2), f, p❩ = ▼{A}❨λg,p2. k g (p1◖l●p2), f, p❩.
-#A #k #f #p1 #p #l #Hk
-@unwind_eq_repl // #g1 #g2 #p2 #Hg
-<list_append_rcons_sn @Hk -Hk // (**) (* auto fail *)
-qed.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma proj_path_proper:
-      proj_path ≗ proj_path.
-// qed.
-
-lemma unwind_path_eq_repl (p):
-      stream_eq_repl … (λf1,f2. ▼[f1]p = ▼[f2]p).
-/2 width=1 by unwind_eq_repl/ qed.
-
-lemma unwind_path_append_sn (p) (f) (q):
-      q●▼[f]p = ▼❨(λg,p. proj_path g (q●p)), f, p❩.
-#p @(path_ind_unwind … p) -p // [ #n #l #p |*: #p ] #IH #f #q
-[ <unwind_d_lcons <unwind_d_lcons <IH -IH //
-| <unwind_m_sn <unwind_m_sn //
-| <unwind_L_sn <unwind_L_sn >unwind_lcons_alt // >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-| <unwind_A_sn <unwind_A_sn >unwind_lcons_alt >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-| <unwind_S_sn <unwind_S_sn >unwind_lcons_alt >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-]
-qed.
-
-lemma unwind_path_lcons (f) (p) (l):
-      l◗▼[f]p = ▼❨(λg,p. proj_path g (l◗p)), f, p❩.
-#f #p #l
->unwind_lcons_alt <unwind_path_append_sn //
-qed.
-
-lemma unwind_path_L_sn (f) (p):
-      (𝗟◗▼[⫯f]p) = ▼[f](𝗟◗p).
-// qed.
-
-lemma unwind_path_A_sn (f) (p):
-      (𝗔◗▼[f]p) = ▼[f](𝗔◗p).
-// qed.
-
-lemma unwind_path_S_sn (f) (p):
-      (𝗦◗▼[f]p) = ▼[f](𝗦◗p).
-// qed.
-
-lemma unwind_path_after_id_sn (p) (f):
-      ▼[𝐢]▼[f]p = ▼[f]p.
-#p @(path_ind_unwind … p) -p // [ #n | #n #l #p | #p ] #IH #f
-[ <unwind_path_d_empty //
-| <unwind_path_d_lcons //
-| <unwind_path_L_sn <unwind_path_L_sn //
-]
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_eq_etc.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_eq_etc.etc
deleted file mode 100644 (file)
index e9b82e0..0000000
+++ /dev/null
@@ -1,15 +0,0 @@
-
-include "delayed_updating/unwind1/unwind.ma".
-(*
-include "ground/relocation/tr_compose_compose.ma".
-include "ground/relocation/tr_compose_eq.ma".
-*)
-(* Advanced constructions with proj_rmap and stream_tls *********************)
-(* COMMENT
-lemma unwind_rmap_tls_d_dx (f) (p) (m) (n):
-      ⇂*[m+n]▼[p]f ≗ ⇂*[m]▼[p◖𝗱n]f.
-#f #p #m #n
-<unwind_rmap_d_dx >nrplus_inj_dx
-/2 width=1 by tr_tls_compose_uni_dx/
-qed.
-*)
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_fsubst.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_fsubst.etc
deleted file mode 100644 (file)
index 4053fb8..0000000
+++ /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/unwind1/unwind_prototerm_eq.ma".
-include "delayed_updating/unwind1/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 // *
-  [ <list_append_empty_dx #H2 destruct
-    elim (unwind_path_root f q) #r #_ #Hr /2 width=2 by/
-  | #l #r #H2 destruct
-    @H0 -H0 [| <unwind_append_proper_dx /2 width=3 by ppc_lcons/ ]
-  ]
-]
-qed-.
-
-lemma unwind_fsubst_dx (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
-      ▼[f](t[⋔p←u]) ⊆ (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u].
-#f #t #u #p #Hu #H1p #H2p #ql * #q * *
-[ #r #Hu #H1 #H2 destruct
-  @or_introl @ex2_intro
-  [|| <unwind_append_proper_dx /2 width=5 by in_ppc_comp_trans/ ]
-  /2 width=3 by ex2_intro/
-| #Hq #H0 #H1 destruct
-  @or_intror @conj [ /2 width=1 by in_comp_unwind_bi/ ] *
-  [ <list_append_empty_dx #Hr @(H0 … (𝐞)) -H0
-    <list_append_empty_dx @H2p -H2p /2 width=1 by prototerm_in_comp_root/
-  | #l #r #Hr
-    elim (unwind_inv_append_proper_dx … Hr) -Hr // #s1 #s2 #Hs1 #_ #H1 destruct
-    lapply (H2p … Hs1) -H2p -Hs1 /2 width=2 by ex_intro/
-  ]
-]
-qed-.
-
-lemma unwind_fsubst (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u] ⇔ ▼[f](t[⋔p←u]).
-/4 width=3 by unwind_fsubst_sn, conj, unwind_fsubst_dx/ qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_preterm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_preterm_eq.etc
deleted file mode 100644 (file)
index 825a85d..0000000
+++ /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 "ground/lib/subset_equivalence.ma".
-include "delayed_updating/syntax/path_structure_inner.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/unwind1/unwind_structure.ma".
-include "delayed_updating/unwind1/unwind_prototerm.ma".
-
-(* UNWIND FOR PRETERM *******************************************************)
-
-(* Constructions with subset_equivalence ************************************)
-
-lemma unwind_grafted_sn (f) (t) (p): p ϵ 𝐈 →
-      ▼[▼[p]f](t⋔p) ⊆ (▼[f]t)⋔(⊗p).
-#f #t #p #Hp #q * #r #Hr #H0 destruct
-@(ex2_intro … Hr) -Hr
-<unwind_append_inner_sn //
-qed-.
-
-lemma unwind_grafted_dx (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)⋔(⊗p) ⊆ ▼[▼[p]f](t⋔p).
-#f #t #p #H1p #H2p #Ht #q * #r #Hr #H0
-elim (unwind_inv_append_inner_sn … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_grafted (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
-      ▼[▼[p]f](t⋔p) ⇔ (▼[f]t)⋔(⊗p).
-/3 width=1 by unwind_grafted_sn, conj, unwind_grafted_dx/ qed.
-
-
-lemma unwind_grafted_S_dx (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)⋔((⊗p)◖𝗦) ⊆ ▼[▼[p]f](t⋔(p◖𝗦)).
-#f #t #p #Hp #Ht #q * #r #Hr
-<list_append_rcons_sn #H0
-elim (unwind_inv_append_proper_dx … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-elim (unwind_path_inv_S_sn … (sym_eq … Hq0)) -Hq0
-#r1 #r2 #Hr1 #Hr2 #H0 destruct
-lapply (preterm_in_root_append_inv_structure_empty_dx … p0 … Ht Hr1)
-[ /2 width=2 by ex_intro/ ] -Hr1 #Hr1 destruct
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_grafted_S (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
-      ▼[▼[p]f](t⋔(p◖𝗦)) ⇔ (▼[f]t)⋔((⊗p)◖𝗦).
-#f #t #p #Hp #Ht
-@conj
-[ >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/unwind1/unwind_prototerm.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_prototerm.etc
deleted file mode 100644 (file)
index bc55294..0000000
+++ /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/unwind1/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/unwind1/unwind_prototerm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_prototerm_eq.etc
deleted file mode 100644 (file)
index e93a78e..0000000
+++ /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 "ground/lib/subset_ext_equivalence.ma".
-include "delayed_updating/unwind1/unwind_eq.ma".
-include "delayed_updating/unwind1/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/unwind1/unwind_structure.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_structure.etc
deleted file mode 100644 (file)
index 0f52a18..0000000
+++ /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/unwind1/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
-<unwind_path_L_sn //
-qed.
-
-lemma unwind_structure (p) (f):
-      ⊗p = ▼[f]⊗p.
-#p @(path_ind_unwind … p) -p //
-qed.
-
-(* Destructions with structure **********************************************)
-
-lemma unwind_des_structure (q) (p) (f):
-      ⊗q = ▼[f]p → ⊗q = ⊗p.
-// qed-.
-
-(* Constructions with proper condition for path *****************************)
-
-lemma unwind_append_proper_dx (p2) (p1) (f): p2 ϵ 𝐏 →
-      (⊗p1)●(▼[▼[p1]f]p2) = ▼[f](p1●p2).
-#p2 #p1 @(path_ind_unwind … p1) -p1 //
-[ #n | #n #l #p1 |*: #p1 ] #IH #f #Hp2
-[ elim (ppc_inv_lcons … Hp2) -Hp2 #l #q #H destruct //
-| <unwind_path_d_lcons <IH //
-| <unwind_path_m_sn <IH //
-| <unwind_path_L_sn <IH //
-| <unwind_path_A_sn <IH //
-| <unwind_path_S_sn <IH //
-]
-qed-.
-
-(* Constructions with inner condition for path ******************************)
-
-lemma unwind_append_inner_sn (p1) (p2) (f): p1 ϵ 𝐈 →
-      (⊗p1)●(▼[▼[p1]f]p2) = ▼[f](p1●p2).
-#p1 @(list_ind_rcons … p1) -p1 // #p1 *
-[ #n ] #_ #p2 #f #Hp1
-[ elim (pic_inv_d_dx … Hp1)
-| <list_append_rcons_sn <unwind_append_proper_dx //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_L_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_A_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_S_dx <list_append_rcons_sn //
-]
-qed-.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma unwind_path_d_empty_dx (n) (p) (f):
-      (⊗p)◖𝗱((▼[p]f)@⧣❨n❩) = ▼[f](p◖𝗱n).
-#n #p #f <unwind_append_proper_dx // 
-qed.
-
-lemma unwind_path_m_dx (p) (f):
-      ⊗p = ▼[f](p◖𝗺).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_L_dx (p) (f):
-      (⊗p)◖𝗟 = ▼[f](p◖𝗟).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_A_dx (p) (f):
-      (⊗p)◖𝗔 = ▼[f](p◖𝗔).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_S_dx (p) (f):
-      (⊗p)◖𝗦 = ▼[f](p◖𝗦).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_root (f) (p):
-      ∃∃r. 𝐞 = ⊗r & ⊗p●r = ▼[f]p.
-#f #p @(list_ind_rcons … p) -p
-[ /2 width=3 by ex2_intro/
-| #p * [ #n ] /2 width=3 by ex2_intro/
-]
-qed-.
-
-(* Advanced inversions with proj_path ***************************************)
-
-lemma unwind_path_inv_d_sn (k) (q) (p) (f):
-      (𝗱k◗q) = ▼[f]p →
-      ∃∃r,h. 𝐞 = ⊗r & (▼[r]f)@⧣❨h❩ = k & 𝐞 = q & r◖𝗱h = p.
-#k #q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty #H destruct -IH
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_d_lcons #H
-  elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_m_sn (q) (p) (f):
-      (𝗺◗q) = ▼[f]p → ⊥.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty #H destruct
-| <unwind_path_d_lcons #H /2 width=2 by/
-| <unwind_path_m_sn #H /2 width=2 by/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_L_sn (q) (p) (f):
-      (𝗟◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[⫯▼[r1]f]r2 & r1●𝗟◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty #H destruct
-| <unwind_path_d_lcons #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_L_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_A_sn (q) (p) (f):
-      (𝗔◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[▼[r1]f]r2 & r1●𝗔◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty #H destruct
-| <unwind_path_d_lcons #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_S_sn (q) (p) (f):
-      (𝗦◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[▼[r1]f]r2 & r1●𝗦◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty #H destruct
-| <unwind_path_d_lcons #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Inversions with proper condition for path ********************************)
-
-lemma unwind_inv_append_proper_dx (q2) (q1) (p) (f):
-      q2 ϵ 𝐏 → q1●q2 = ▼[f]p →
-      ∃∃p1,p2. ⊗p1 = q1 & ▼[▼[p1]f]p2 = q2 & p1●p2 = p.
-#q2 #q1 elim q1 -q1
-[ #p #f #Hq2 <list_append_empty_sn #H destruct
-  /2 width=5 by ex3_2_intro/
-| * [ #n1 ] #q1 #IH #p #f #Hq2 <list_append_lcons_sn #H
-  [ elim (unwind_path_inv_d_sn … H) -H #r1 #m1 #_ #_ #H0 #_ -IH
-    elim (eq_inv_list_empty_append … H0) -H0 #_ #H0 destruct
-    elim Hq2 -Hq2 //
-  | elim (unwind_path_inv_m_sn … H)
-  | elim (unwind_path_inv_L_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗟◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  | elim (unwind_path_inv_A_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗔◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  | elim (unwind_path_inv_S_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗦◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  ]
-]
-qed-.
-
-(* Inversions with inner condition for path *********************************)
-
-lemma unwind_inv_append_inner_sn (q1) (q2) (p) (f):
-      q1 ϵ 𝐈 → q1●q2 = ▼[f]p →
-      ∃∃p1,p2. ⊗p1 = q1 & ▼[▼[p1]f]p2 = q2 & p1●p2 = p.
-#q1 @(list_ind_rcons … q1) -q1
-[ #q2 #p #f #Hq1 <list_append_empty_sn #H destruct
-  /2 width=5 by ex3_2_intro/
-| #q1 * [ #n1 ] #_ #q2 #p #f #Hq2
-  [ elim (pic_inv_d_dx … Hq2)
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_m_sn … (sym_eq … H2))
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_L_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗟)) [1,3: // ]
-    [ <structure_append <structure_L_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_A_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗔)) [1,3: // ]
-    [ <structure_append <structure_A_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_S_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗦)) [1,3: // ]
-    [ <structure_append <structure_S_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  ]
-]
-qed-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_structure_depth.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind1/unwind_structure_depth.etc
deleted file mode 100644 (file)
index df3423e..0000000
+++ /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/unwind1/unwind.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/arith/nat_pred_succ.ma".
-
-(* UNWIND FOR PATH *********************************************************)
-
-(* Basic constructions with structure and depth ****************************)
-
-lemma unwind_rmap_structure (p) (f):
-      (⫯*[♭p]f) = ▼[⊗p]f.
-#p elim p -p //
-* [ #n ] #p #IH #f //
-[ <unwind_rmap_A_sn //
-| <unwind_rmap_S_sn //
-]
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind.etc
deleted file mode 100644 (file)
index 1713625..0000000
+++ /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 //
-[ <unwind_rmap_m_sn <unwind_rmap_m_sn //
-| <unwind_rmap_A_sn <unwind_rmap_A_sn //
-| <unwind_rmap_S_sn <unwind_rmap_S_sn //
-]
-qed.
-
-(* Advanced constructions with proj_rmap and path_rcons *********************)
-
-lemma unwind_rmap_d_dx (f) (p) (n):
-      (▼[p]f)∘𝐮❨ninj n❩ = ▼[p◖𝗱n]f.
-// qed.
-
-lemma unwind_rmap_m_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗺]f.
-// qed.
-
-lemma unwind_rmap_L_dx (f) (p):
-      (⫯▼[p]f) = ▼[p◖𝗟]f.
-// qed.
-
-lemma unwind_rmap_A_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗔]f.
-// qed.
-
-lemma unwind_rmap_S_dx (f) (p):
-▼[p]f = ▼[p◖𝗦]f.
-// qed.
-
-lemma unwind_rmap_pap_d_dx (f) (p) (n) (m):
-      ▼[p]f@❨m+n❩ = ▼[p◖𝗱n]f@❨m❩.
-#f #p #n #m
-<unwind_rmap_d_dx <tr_compose_pap <tr_uni_pap //
-qed.
-
-(* Advanced eliminations with path ******************************************)
-
-lemma path_ind_unwind (Q:predicate …):
-      Q (𝐞) →
-      (∀n. Q (𝐞) → Q (𝗱n◗𝐞)) →
-      (∀n,l,p. Q (l◗p) → Q (𝗱n◗l◗p)) →
-      (∀p. Q p → Q (𝗺◗p)) →
-      (∀p. Q p → Q (𝗟◗p)) →
-      (∀p. Q p → Q (𝗔◗p)) →
-      (∀p. Q p → Q (𝗦◗p)) →
-      ∀p. Q p.
-#Q #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #p
-elim p -p [| * [ #n * ] ]
-/2 width=1 by/
-qed-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_depth.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_depth.etc
deleted file mode 100644 (file)
index 8715fb5..0000000
+++ /dev/null
@@ -1,48 +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/path_depth.ma".
-include "ground/relocation/tr_id_compose.ma".
-include "ground/relocation/tr_compose_compose.ma".
-include "ground/relocation/tr_compose_pn.ma".
-include "ground/relocation/tr_compose_eq.ma".
-include "ground/relocation/tr_pn_eq.ma".
-include "ground/lib/stream_eq_eq.ma".
-
-(* UNWIND FOR PATH **********************************************************)
-
-(* Constructions with depth *************************************************)
-
-lemma unwind_rmap_decompose (p) (f):
-      ▼[p]f ≗ (⫯*[❘p❘]f)∘(▼[p]𝐢).
-#p @(list_ind_rcons … p) -p
-[ #f <unwind_rmap_empty <unwind_rmap_empty <tr_pushs_zero //
-| #p * [ #n ] #IH #f //
-  [ <unwind_rmap_d_dx <unwind_rmap_d_dx <depth_d_dx
-    @(stream_eq_trans … (tr_compose_assoc …))
-    /2 width=1 by tr_compose_eq_repl/
-  | <unwind_rmap_L_dx <unwind_rmap_L_dx <depth_L_dx
-    <tr_pushs_succ <tr_compose_push_bi
-    /2 width=1 by tr_push_eq_repl/
-  ]
-]
-qed.
-
-lemma unwind_rmap_pap_le (f) (p) (n):
-      (▼[p]𝐢)@❨n❩ < ↑❘p❘ → (▼[p]𝐢)@❨n❩ = (▼[p]f)@❨n❩.
-#f #p #n #Hn
->(tr_pap_eq_repl … (▼[p]f) … (unwind_rmap_decompose …))
-<tr_compose_pap @tr_pap_pushs_le //
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_eq.etc
deleted file mode 100644 (file)
index b3c0b40..0000000
+++ /dev/null
@@ -1,122 +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 "ground/relocation/tr_uni_compose.ma".
-include "ground/relocation/tr_compose_compose.ma".
-include "ground/relocation/tr_compose_eq.ma".
-include "ground/relocation/tr_pn_eq.ma".
-
-(* UNWIND FOR PATH **********************************************************)
-
-definition unwind_exteq (A): relation2 (unwind_continuation A) (unwind_continuation A) ≝
-           λk1,k2. ∀f1,f2,p. f1 ≗ f2 → k1 f1 p = k2 f2 p.
-
-interpretation
-  "extensional equivalence (unwind continuation)"
-  'RingEq A k1 k2 = (unwind_exteq A k1 k2).
-
-(* Constructions with unwind_exteq ******************************************)
-
-lemma unwind_eq_repl (A) (p) (k1) (k2):
-      k1 ≗{A} k2 → stream_eq_repl … (λf1,f2. ▼❨k1, f1, p❩ = ▼❨k2, f2, p❩).
-#A #p @(path_ind_unwind … p) -p [| #n #IH | #n #l0 #q #IH |*: #q #IH ]
-#k1 #k2 #Hk #f1 #f2 #Hf
-[ <unwind_empty <unwind_empty /2 width=1 by/
-| <unwind_d_empty_sn <unwind_d_empty_sn <(tr_pap_eq_repl … Hf)
-  /3 width=1 by tr_compose_eq_repl, stream_eq_refl/
-| <unwind_d_lcons_sn <unwind_d_lcons_sn
-  /3 width=1 by tr_compose_eq_repl, stream_eq_refl/
-| /2 width=1 by/
-| /3 width=1 by tr_push_eq_repl/
-| /3 width=1 by/
-| /3 width=1 by/
-]
-qed-.
-
-(* Advanced constructions ***************************************************)
-
-lemma unwind_lcons_alt (A) (k) (f) (p) (l): k ≗ k →
-      ▼❨λg,p2. k g (l◗p2), f, p❩ = ▼{A}❨λg,p2. k g ((l◗𝐞)●p2), f, p❩.
-#A #k #f #p #l #Hk
-@unwind_eq_repl // #g1 #g2 #p2 #Hg @Hk -Hk // (**) (* auto fail *)
-qed.
-
-lemma unwind_append_rcons_sn (A) (k) (f) (p1) (p) (l): k ≗ k →
-      ▼❨λg,p2. k g (p1●l◗p2), f, p❩ = ▼{A}❨λg,p2. k g (p1◖l●p2), f, p❩.
-#A #k #f #p1 #p #l #Hk
-@unwind_eq_repl // #g1 #g2 #p2 #Hg
-<list_append_rcons_sn @Hk -Hk // (**) (* auto fail *)
-qed.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma proj_path_proper:
-      proj_path ≗ proj_path.
-// qed.
-
-lemma unwind_path_eq_repl (p):
-      stream_eq_repl … (λf1,f2. ▼[f1]p = ▼[f2]p).
-/2 width=1 by unwind_eq_repl/ qed.
-
-lemma unwind_path_append_sn (p) (f) (q):
-      q●▼[f]p = ▼❨(λg,p. proj_path g (q●p)), f, p❩.
-#p @(path_ind_unwind … p) -p // [ #n #l #p |*: #p ] #IH #f #q
-[ <unwind_d_lcons_sn <unwind_d_lcons_sn <IH -IH //
-| <unwind_m_sn <unwind_m_sn //
-| <unwind_L_sn <unwind_L_sn >unwind_lcons_alt // >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-| <unwind_A_sn <unwind_A_sn >unwind_lcons_alt >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-| <unwind_S_sn <unwind_S_sn >unwind_lcons_alt >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-]
-qed.
-
-lemma unwind_path_lcons (f) (p) (l):
-      l◗▼[f]p = ▼❨(λg,p. proj_path g (l◗p)), f, p❩.
-#f #p #l
->unwind_lcons_alt <unwind_path_append_sn //
-qed.
-
-lemma unwind_path_L_sn (f) (p):
-      (𝗟◗▼[⫯f]p) = ▼[f](𝗟◗p).
-// qed.
-
-lemma unwind_path_A_sn (f) (p):
-      (𝗔◗▼[f]p) = ▼[f](𝗔◗p).
-// qed.
-
-lemma unwind_path_S_sn (f) (p):
-      (𝗦◗▼[f]p) = ▼[f](𝗦◗p).
-// qed.
-
-lemma unwind_path_after (p) (f1) (f2):
-      ▼[f2]▼[f1]p = ▼[f2∘f1]p.
-#p @(path_ind_unwind … p) -p // [ #n #l #p | #p ] #IH #f1 #f2
-[ <unwind_path_d_lcons_sn <unwind_path_d_lcons_sn
-  >(unwind_path_eq_repl … (tr_compose_assoc …)) //
-| <unwind_path_L_sn <unwind_path_L_sn <unwind_path_L_sn
-  >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
-<unwind_rmap_d_dx >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
deleted file mode 100644 (file)
index a9e563d..0000000
+++ /dev/null
@@ -1,10 +0,0 @@
-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 //
-<depth_d_dx >list_cons_shift <list_append_assoc
-<unwind_rmap_d_dx
-check tr_tls_compose_uni_sn
\ No newline at end of file
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
deleted file mode 100644 (file)
index b5d92a5..0000000
+++ /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
-<unwind_path_L_sn //
-qed.
-
-lemma unwind_structure (p) (f):
-      ⊗p = ▼[f]⊗p.
-#p @(path_ind_unwind … p) -p //
-qed.
-
-(* Destructions with structure **********************************************)
-
-lemma unwind_des_structure (q) (p) (f):
-      ⊗q = ▼[f]p → ⊗q = ⊗p.
-// qed-.
-
-(* Constructions with proper condition for path *****************************)
-
-lemma unwind_append_proper_dx (p2) (p1) (f): p2 ϵ 𝐏 →
-      (⊗p1)●(▼[▼[p1]f]p2) = ▼[f](p1●p2).
-#p2 #p1 @(path_ind_unwind … p1) -p1 //
-[ #n | #n #l #p1 |*: #p1 ] #IH #f #Hp2
-[ elim (ppc_inv_lcons … Hp2) -Hp2 #l #q #H destruct //
-| <unwind_path_d_lcons_sn <IH //
-| <unwind_path_m_sn <IH //
-| <unwind_path_L_sn <IH //
-| <unwind_path_A_sn <IH //
-| <unwind_path_S_sn <IH //
-]
-qed-.
-
-(* Constructions with inner condition for path ******************************)
-
-lemma unwind_append_inner_sn (p1) (p2) (f): p1 ϵ 𝐈 →
-      (⊗p1)●(▼[▼[p1]f]p2) = ▼[f](p1●p2).
-#p1 @(list_ind_rcons … p1) -p1 // #p1 *
-[ #n ] #_ #p2 #f #Hp1
-[ elim (pic_inv_d_dx … Hp1)
-| <list_append_rcons_sn <unwind_append_proper_dx //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_L_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_A_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_S_dx <list_append_rcons_sn //
-]
-qed-.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma unwind_path_d_empty_dx (n) (p) (f):
-      (⊗p)◖𝗱((▼[p]f)@❨n❩) = ▼[f](p◖𝗱n).
-#n #p #f <unwind_append_proper_dx // 
-qed.
-
-lemma unwind_path_m_dx (p) (f):
-      ⊗p = ▼[f](p◖𝗺).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_L_dx (p) (f):
-      (⊗p)◖𝗟 = ▼[f](p◖𝗟).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_A_dx (p) (f):
-      (⊗p)◖𝗔 = ▼[f](p◖𝗔).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_S_dx (p) (f):
-      (⊗p)◖𝗦 = ▼[f](p◖𝗦).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_root (f) (p):
-      ∃∃r. 𝐞 = ⊗r & ⊗p●r = ▼[f]p.
-#f #p @(list_ind_rcons … p) -p
-[ /2 width=3 by ex2_intro/
-| #p * [ #n ] /2 width=3 by ex2_intro/
-]
-qed-.
-
-(* Advanced inversions with proj_path ***************************************)
-
-lemma unwind_path_inv_d_sn (k) (q) (p) (f):
-      (𝗱k◗q) = ▼[f]p →
-      ∃∃r,h. 𝐞 = ⊗r & (▼[r]f)@❨h❩ = k & 𝐞 = q & r◖𝗱h = p.
-#k #q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct -IH
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_m_sn (q) (p) (f):
-      (𝗺◗q) = ▼[f]p → ⊥.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H /2 width=2 by/
-| <unwind_path_m_sn #H /2 width=2 by/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_L_sn (q) (p) (f):
-      (𝗟◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[⫯▼[r1]f]r2 & r1●𝗟◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_L_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_A_sn (q) (p) (f):
-      (𝗔◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[▼[r1]f]r2 & r1●𝗔◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_S_sn (q) (p) (f):
-      (𝗦◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[▼[r1]f]r2 & r1●𝗦◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Inversions with proper condition for path ********************************)
-
-lemma unwind_inv_append_proper_dx (q2) (q1) (p) (f):
-      q2 ϵ 𝐏 → q1●q2 = ▼[f]p →
-      ∃∃p1,p2. ⊗p1 = q1 & ▼[▼[p1]f]p2 = q2 & p1●p2 = p.
-#q2 #q1 elim q1 -q1
-[ #p #f #Hq2 <list_append_empty_sn #H destruct
-  /2 width=5 by ex3_2_intro/
-| * [ #n1 ] #q1 #IH #p #f #Hq2 <list_append_lcons_sn #H
-  [ elim (unwind_path_inv_d_sn … H) -H #r1 #m1 #_ #_ #H0 #_ -IH
-    elim (eq_inv_list_empty_append … H0) -H0 #_ #H0 destruct
-    elim Hq2 -Hq2 //
-  | elim (unwind_path_inv_m_sn … H)
-  | elim (unwind_path_inv_L_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗟◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  | elim (unwind_path_inv_A_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗔◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  | elim (unwind_path_inv_S_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗦◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  ]
-]
-qed-.
-
-(* Inversions with inner condition for path *********************************)
-
-lemma unwind_inv_append_inner_sn (q1) (q2) (p) (f):
-      q1 ϵ 𝐈 → q1●q2 = ▼[f]p →
-      ∃∃p1,p2. ⊗p1 = q1 & ▼[▼[p1]f]p2 = q2 & p1●p2 = p.
-#q1 @(list_ind_rcons … q1) -q1
-[ #q2 #p #f #Hq1 <list_append_empty_sn #H destruct
-  /2 width=5 by ex3_2_intro/
-| #q1 * [ #n1 ] #_ #q2 #p #f #Hq2
-  [ elim (pic_inv_d_dx … Hq2)
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_m_sn … (sym_eq … H2))
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_L_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗟)) [1,3: // ]
-    [ <structure_append <structure_L_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_A_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗔)) [1,3: // ]
-    [ <structure_append <structure_A_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_S_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗦)) [1,3: // ]
-    [ <structure_append <structure_S_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  ]
-]
-qed-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure_depth.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_structure_depth.etc
deleted file mode 100644 (file)
index 577a95b..0000000
+++ /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/unwind2/unwind.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/arith/nat_pred_succ.ma".
-
-(* UNWIND FOR PATH **********************************************************)
-
-(* Basic constructions with structure and depth *****************************)
-
-lemma unwind_rmap_structure (p) (f):
-      (⫯*[❘p❘]f) = ▼[⊗p]f.
-#p elim p -p //
-* [ #n ] #p #IH #f //
-[ <unwind_rmap_A_sn //
-| <unwind_rmap_S_sn //
-]
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind.etc
deleted file mode 100644 (file)
index 87790b3..0000000
+++ /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 "delayed_updating/relocation/sbr_pap_push.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".
-include "ground/relocation/tr_pn.ma".
-
-(* UNWIND FOR PATH **********************************************************)
-
-definition unwind_continuation (A:Type[0]) ≝
-           tr_map → path → A.
-
-(* Note: f is a stack of update functions *)
-rec definition unwind_gen (A:Type[0]) (k:unwind_continuation A) (f:tr_map) (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❩⫯❘f) q
-    | list_lcons _ _ ⇒ unwind_gen (A) k (f❘@❨n❩⫯❘f) 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❘@❨n❩⫯❘f, 𝐞❩ = ▼{A}❨k, f, 𝗱n◗𝐞❩.
-// qed.
-
-lemma unwind_d_lcons_sn (A) (k) (p) (l) (n) (f):
-      ▼❨k, f❘@❨n❩⫯❘f, 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❘@❨n❩⫯❘f](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❘@❨n❩⫯❘f) = ▼[𝗱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 //
-[ <unwind_rmap_m_sn <unwind_rmap_m_sn //
-| <unwind_rmap_A_sn <unwind_rmap_A_sn //
-| <unwind_rmap_S_sn <unwind_rmap_S_sn //
-]
-qed.
-
-(* Advanced constructions with proj_rmap and path_rcons *********************)
-
-lemma unwind_rmap_d_dx (f) (p) (n):
-      (▼[p]f)❘@❨n❩⫯❘▼[p]f = ▼[p◖𝗱n]f.
-// qed.
-
-lemma unwind_rmap_m_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗺]f.
-// qed.
-
-lemma unwind_rmap_L_dx (f) (p):
-      (⫯▼[p]f) = ▼[p◖𝗟]f.
-// qed.
-
-lemma unwind_rmap_A_dx (f) (p):
-      ▼[p]f = ▼[p◖𝗔]f.
-// qed.
-
-lemma unwind_rmap_S_dx (f) (p):
-▼[p]f = ▼[p◖𝗦]f.
-// qed.
-
-lemma unwind_rmap_pap_d_dx (f) (p) (n) (m):
-      m+▼[p]f❘@❨n❩ = ▼[p◖𝗱n]f❘@❨m❩.
-#f #p #n #m
-<unwind_rmap_d_dx <sbr_push_pap //
-qed.
-
-(* Advanced eliminations with path ******************************************)
-
-lemma path_ind_unwind (Q:predicate …):
-      Q (𝐞) →
-      (∀n. Q (𝐞) → Q (𝗱n◗𝐞)) →
-      (∀n,l,p. Q (l◗p) → Q (𝗱n◗l◗p)) →
-      (∀p. Q p → Q (𝗺◗p)) →
-      (∀p. Q p → Q (𝗟◗p)) →
-      (∀p. Q p → Q (𝗔◗p)) →
-      (∀p. Q p → Q (𝗦◗p)) →
-      ∀p. Q p.
-#Q #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #p
-elim p -p [| * [ #n * ] ]
-/2 width=1 by/
-qed-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_constructors.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_constructors.etc
deleted file mode 100644 (file)
index 4849ecf..0000000
+++ /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 "delayed_updating/unwind3/unwind_prototerm_eq.ma".
-include "delayed_updating/syntax/prototerm_constructors.ma".
-include "delayed_updating/relocation/sbr_push_uni.ma".
-
-(* UNWIND FOR PROTOTERM *****************************************************)
-
-(* Constructions with constructors for prototerm ****************************)
-
-lemma unwind_iref_sn (f) (t:prototerm) (n:pnat):
-      ▼[f❘@❨n❩⫯❘f]t ⊆ ▼[f](𝛗n.t).
-#f #t #n #p * #q #Hq #H0 destruct
-@(ex2_intro … (𝗱n◗𝗺◗q))
-/2 width=1 by in_comp_iref/
-qed-.
-
-lemma unwind_iref_dx (f) (t) (n:pnat):
-      ▼[f](𝛗n.t) ⊆ ▼[f❘@❨n❩⫯❘f]t.
-#f #t #n #p * #q #Hq #H0 destruct
-elim (in_comp_inv_iref … Hq) -Hq #p #Hp #Ht destruct
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_iref (f) (t) (n:pnat):
-      ▼[f❘@❨n❩⫯❘f]t ⇔ ▼[f](𝛗n.t).
-/3 width=1 by conj, unwind_iref_sn, unwind_iref_dx/
-qed.
-
-lemma unwind_iref_id (t) (n:pnat):
-      ▼[𝐮❨n❩]t ⇔ ▼[𝐢](𝛗n.t).
-// qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_eq.etc
deleted file mode 100644 (file)
index 78d755b..0000000
+++ /dev/null
@@ -1,124 +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/unwind3/unwind.ma".
-include "delayed_updating/relocation/sbr_pap_id.ma".
-include "delayed_updating/relocation/sbr_pap_eq.ma".
-include "delayed_updating/relocation/sbr_push_eq.ma".
-include "ground/relocation/tr_pn_eq.ma".
-(*
-include "ground/lib/stream_tls.ma".
-*)
-
-(* UNWIND FOR PATH **********************************************************)
-
-definition unwind_exteq (A): relation2 (unwind_continuation A) (unwind_continuation A) ≝
-           λk1,k2. ∀f1,f2,p. f1 ≗ f2 → k1 f1 p = k2 f2 p.
-
-interpretation
-  "extensional equivalence (unwind continuation)"
-  'RingEq A k1 k2 = (unwind_exteq A k1 k2).
-
-(* Constructions with unwind_exteq ******************************************)
-
-lemma unwind_eq_repl (A) (p) (k1) (k2):
-      k1 ≗{A} k2 → stream_eq_repl … (λf1,f2. ▼❨k1, f1, p❩ = ▼❨k2, f2, p❩).
-#A #p @(path_ind_unwind … p) -p [| #n #IH | #n #l0 #q #IH |*: #q #IH ]
-#k1 #k2 #Hk #f1 #f2 #Hf
-[ <unwind_empty <unwind_empty /2 width=1 by/
-| <unwind_d_empty_sn <unwind_d_empty_sn <(sbr_pap_eq_repl_sn … Hf) 
-  /3 width=1 by sbr_push_eq_repl_dx/
-| <unwind_d_lcons_sn <unwind_d_lcons_sn <(sbr_pap_eq_repl_sn … Hf)
-  /3 width=1 by sbr_push_eq_repl_dx/
-| /2 width=1 by/
-| /3 width=1 by tr_push_eq_repl/
-| /3 width=1 by/
-| /3 width=1 by/
-]
-qed-.
-
-(* Advanced constructions ***************************************************)
-
-lemma unwind_lcons_alt (A) (k) (f) (p) (l): k ≗ k →
-      ▼❨λg,p2. k g (l◗p2), f, p❩ = ▼{A}❨λg,p2. k g ((l◗𝐞)●p2), f, p❩.
-#A #k #f #p #l #Hk
-@unwind_eq_repl // #g1 #g2 #p2 #Hg @Hk -Hk // (**) (* auto fail *)
-qed.
-
-lemma unwind_append_rcons_sn (A) (k) (f) (p1) (p) (l): k ≗ k →
-      ▼❨λg,p2. k g (p1●l◗p2), f, p❩ = ▼{A}❨λg,p2. k g (p1◖l●p2), f, p❩.
-#A #k #f #p1 #p #l #Hk
-@unwind_eq_repl // #g1 #g2 #p2 #Hg
-<list_append_rcons_sn @Hk -Hk // (**) (* auto fail *)
-qed.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma proj_path_proper:
-      proj_path ≗ proj_path.
-// qed.
-
-lemma unwind_path_eq_repl (p):
-      stream_eq_repl … (λf1,f2. ▼[f1]p = ▼[f2]p).
-/2 width=1 by unwind_eq_repl/ qed.
-
-lemma unwind_path_append_sn (p) (f) (q):
-      q●▼[f]p = ▼❨(λg,p. proj_path g (q●p)), f, p❩.
-#p @(path_ind_unwind … p) -p // [ #n #l #p |*: #p ] #IH #f #q
-[ <unwind_d_lcons_sn <unwind_d_lcons_sn <IH -IH //
-| <unwind_m_sn <unwind_m_sn //
-| <unwind_L_sn <unwind_L_sn >unwind_lcons_alt // >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-| <unwind_A_sn <unwind_A_sn >unwind_lcons_alt >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-| <unwind_S_sn <unwind_S_sn >unwind_lcons_alt >unwind_append_rcons_sn //
-  <IH <IH -IH <list_append_rcons_sn //
-]
-qed.
-
-lemma unwind_path_lcons (f) (p) (l):
-      l◗▼[f]p = ▼❨(λg,p. proj_path g (l◗p)), f, p❩.
-#f #p #l
->unwind_lcons_alt <unwind_path_append_sn //
-qed.
-
-lemma unwind_path_L_sn (f) (p):
-      (𝗟◗▼[⫯f]p) = ▼[f](𝗟◗p).
-// qed.
-
-lemma unwind_path_A_sn (f) (p):
-      (𝗔◗▼[f]p) = ▼[f](𝗔◗p).
-// qed.
-
-lemma unwind_path_S_sn (f) (p):
-      (𝗦◗▼[f]p) = ▼[f](𝗦◗p).
-// qed.
-
-lemma unwind_path_after_id_sn (p) (f):
-      ▼[𝐢]▼[f]p = ▼[f]p.
-#p @(path_ind_unwind … p) -p // #p #IH #f
-[ <unwind_path_d_empty_sn <unwind_path_d_empty_sn //
-| <unwind_path_L_sn <unwind_path_L_sn //
-]
-qed.
-
-(* Advanced constructions with proj_rmap and stream_tls *********************)
-(* COMMENT
-lemma unwind_rmap_tls_d_dx (f) (p) (m:pnat) (n):
-      ⇂*[m]▼[p]f ≗ ⇂*[m]▼[p◖𝗱n]f.
-#f #p #m #n
-<unwind_rmap_d_dx
-/2 width=1 by tr_tls_compose_uni_sn/
-qed.
-*)
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_fsubst.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_fsubst.etc
deleted file mode 100644 (file)
index 411c562..0000000
+++ /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/unwind3/unwind_prototerm_eq.ma".
-include "delayed_updating/unwind3/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 // *
-  [ <list_append_empty_dx #H2 destruct
-    elim (unwind_path_root f q) #r #_ #Hr /2 width=2 by/
-  | #l #r #H2 destruct
-    @H0 -H0 [| <unwind_append_proper_dx /2 width=3 by ppc_lcons/ ]
-  ]
-]
-qed-.
-
-lemma unwind_fsubst_dx (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
-      ▼[f](t[⋔p←u]) ⊆ (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u].
-#f #t #u #p #Hu #H1p #H2p #ql * #q * *
-[ #r #Hu #H1 #H2 destruct
-  @or_introl @ex2_intro
-  [|| <unwind_append_proper_dx /2 width=5 by in_ppc_comp_trans/ ]
-  /2 width=3 by ex2_intro/
-| #Hq #H0 #H1 destruct
-  @or_intror @conj [ /2 width=1 by in_comp_unwind_bi/ ] *
-  [ <list_append_empty_dx #Hr @(H0 … (𝐞)) -H0
-    <list_append_empty_dx @H2p -H2p /2 width=1 by prototerm_in_comp_root/
-  | #l #r #Hr
-    elim (unwind_inv_append_proper_dx … Hr) -Hr // #s1 #s2 #Hs1 #_ #H1 destruct
-    lapply (H2p … Hs1) -H2p -Hs1 /2 width=2 by ex_intro/
-  ]
-]
-qed-.
-
-lemma unwind_fsubst (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u] ⇔ ▼[f](t[⋔p←u]).
-/4 width=3 by unwind_fsubst_sn, conj, unwind_fsubst_dx/ qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_lift.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_lift.etc
deleted file mode 100644 (file)
index 9530004..0000000
+++ /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 "delayed_updating/unwind3/unwind_structure.ma".
-include "delayed_updating/substitution/lift_structure.ma".
-
-(* UNWIND FOR PATH **********************************************************)
-
-(* Constructions with lift **************************************************)
-
-lemma pippo (p) (n:pnat) (f):
-      ↑[𝐮❨n❩]▼[⇂*[n]f]p = ▼[𝐮❨n❩∘f]p.
-#p @(list_ind_rcons … p) -p //
-#p * [ #m ] #IH #n #f //
-[
-| <unwind_path_m_dx
-
- //
-<depth_d_dx >list_cons_shift <list_append_assoc <unwind_rmap_d_dx
-/3 width=3 by tr_tls_compose_uni_sn, stream_eq_trans/
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_preterm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_preterm_eq.etc
deleted file mode 100644 (file)
index d245cb1..0000000
+++ /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 "ground/lib/subset_equivalence.ma".
-include "delayed_updating/syntax/path_structure_inner.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/unwind3/unwind_structure.ma".
-include "delayed_updating/unwind3/unwind_prototerm.ma".
-
-(* UNWIND FOR PRETERM *******************************************************)
-
-(* Constructions with subset_equivalence ************************************)
-
-lemma unwind_grafted_sn (f) (t) (p): p ϵ 𝐈 →
-      ▼[▼[p]f](t⋔p) ⊆ (▼[f]t)⋔(⊗p).
-#f #t #p #Hp #q * #r #Hr #H0 destruct
-@(ex2_intro … Hr) -Hr
-<unwind_append_inner_sn //
-qed-.
-
-lemma unwind_grafted_dx (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)⋔(⊗p) ⊆ ▼[▼[p]f](t⋔p).
-#f #t #p #H1p #H2p #Ht #q * #r #Hr #H0
-elim (unwind_inv_append_inner_sn … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_grafted (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
-      ▼[▼[p]f](t⋔p) ⇔ (▼[f]t)⋔(⊗p).
-/3 width=1 by unwind_grafted_sn, conj, unwind_grafted_dx/ qed.
-
-
-lemma unwind_grafted_S_dx (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)⋔((⊗p)◖𝗦) ⊆ ▼[▼[p]f](t⋔(p◖𝗦)).
-#f #t #p #Hp #Ht #q * #r #Hr
-<list_append_rcons_sn #H0
-elim (unwind_inv_append_proper_dx … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-elim (unwind_path_inv_S_sn … (sym_eq … Hq0)) -Hq0
-#r1 #r2 #Hr1 #Hr2 #H0 destruct
-lapply (preterm_in_root_append_inv_structure_empty_dx … p0 … Ht Hr1)
-[ /2 width=2 by ex_intro/ ] -Hr1 #Hr1 destruct
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_grafted_S (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
-      ▼[▼[p]f](t⋔(p◖𝗦)) ⇔ (▼[f]t)⋔((⊗p)◖𝗦).
-#f #t #p #Hp #Ht
-@conj
-[ >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
deleted file mode 100644 (file)
index c4a103b..0000000
+++ /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/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
deleted file mode 100644 (file)
index 01d302f..0000000
+++ /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 "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
deleted file mode 100644 (file)
index c52a438..0000000
+++ /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/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
-<unwind_path_L_sn //
-qed.
-
-lemma unwind_structure (p) (f):
-      ⊗p = ▼[f]⊗p.
-#p @(path_ind_unwind … p) -p //
-qed.
-
-(* Destructions with structure **********************************************)
-
-lemma unwind_des_structure (q) (p) (f):
-      ⊗q = ▼[f]p → ⊗q = ⊗p.
-// qed-.
-
-(* Constructions with proper condition for path *****************************)
-
-lemma unwind_append_proper_dx (p2) (p1) (f): p2 ϵ 𝐏 →
-      (⊗p1)●(▼[▼[p1]f]p2) = ▼[f](p1●p2).
-#p2 #p1 @(path_ind_unwind … p1) -p1 //
-[ #n | #n #l #p1 |*: #p1 ] #IH #f #Hp2
-[ elim (ppc_inv_lcons … Hp2) -Hp2 #l #q #H destruct //
-| <unwind_path_d_lcons_sn <IH //
-| <unwind_path_m_sn <IH //
-| <unwind_path_L_sn <IH //
-| <unwind_path_A_sn <IH //
-| <unwind_path_S_sn <IH //
-]
-qed-.
-
-(* Constructions with inner condition for path ******************************)
-
-lemma unwind_append_inner_sn (p1) (p2) (f): p1 ϵ 𝐈 →
-      (⊗p1)●(▼[▼[p1]f]p2) = ▼[f](p1●p2).
-#p1 @(list_ind_rcons … p1) -p1 // #p1 *
-[ #n ] #_ #p2 #f #Hp1
-[ elim (pic_inv_d_dx … Hp1)
-| <list_append_rcons_sn <unwind_append_proper_dx //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_L_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_A_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <unwind_append_proper_dx //
-  <structure_S_dx <list_append_rcons_sn //
-]
-qed-.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma unwind_path_d_empty_dx (n) (p) (f):
-      (⊗p)◖𝗱((▼[p]f)❘@❨n❩) = ▼[f](p◖𝗱n).
-#n #p #f <unwind_append_proper_dx // 
-qed.
-
-lemma unwind_path_m_dx (p) (f):
-      ⊗p = ▼[f](p◖𝗺).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_L_dx (p) (f):
-      (⊗p)◖𝗟 = ▼[f](p◖𝗟).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_A_dx (p) (f):
-      (⊗p)◖𝗔 = ▼[f](p◖𝗔).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_S_dx (p) (f):
-      (⊗p)◖𝗦 = ▼[f](p◖𝗦).
-#p #f <unwind_append_proper_dx //
-qed.
-
-lemma unwind_path_root (f) (p):
-      ∃∃r. 𝐞 = ⊗r & ⊗p●r = ▼[f]p.
-#f #p @(list_ind_rcons … p) -p
-[ /2 width=3 by ex2_intro/
-| #p * [ #n ] /2 width=3 by ex2_intro/
-]
-qed-.
-
-(* Advanced inversions with proj_path ***************************************)
-
-lemma unwind_path_inv_d_sn (k) (q) (p) (f):
-      (𝗱k◗q) = ▼[f]p →
-      ∃∃r,h. 𝐞 = ⊗r & (▼[r]f)❘@❨h❩ = k & 𝐞 = q & r◖𝗱h = p.
-#k #q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct -IH
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
-  /2 width=5 by ex4_2_intro/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_m_sn (q) (p) (f):
-      (𝗺◗q) = ▼[f]p → ⊥.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H /2 width=2 by/
-| <unwind_path_m_sn #H /2 width=2 by/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_L_sn (q) (p) (f):
-      (𝗟◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[⫯▼[r1]f]r2 & r1●𝗟◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_L_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_A_sn (q) (p) (f):
-      (𝗔◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[▼[r1]f]r2 & r1●𝗔◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_S_sn #H destruct
-]
-qed-.
-
-lemma unwind_path_inv_S_sn (q) (p) (f):
-      (𝗦◗q) = ▼[f]p →
-      ∃∃r1,r2. 𝐞 = ⊗r1 & q = ▼[▼[r1]f]r2 & r1●𝗦◗r2 = p.
-#q #p @(path_ind_unwind … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <unwind_path_empty #H destruct
-| <unwind_path_d_empty_sn #H destruct
-| <unwind_path_d_lcons_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/
-| <unwind_path_m_sn #H
-  elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
-  /2 width=5 by ex3_2_intro/| <unwind_path_L_sn #H destruct
-| <unwind_path_A_sn #H destruct
-| <unwind_path_S_sn #H destruct -IH
-  /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Inversions with proper condition for path ********************************)
-
-lemma unwind_inv_append_proper_dx (q2) (q1) (p) (f):
-      q2 ϵ 𝐏 → q1●q2 = ▼[f]p →
-      ∃∃p1,p2. ⊗p1 = q1 & ▼[▼[p1]f]p2 = q2 & p1●p2 = p.
-#q2 #q1 elim q1 -q1
-[ #p #f #Hq2 <list_append_empty_sn #H destruct
-  /2 width=5 by ex3_2_intro/
-| * [ #n1 ] #q1 #IH #p #f #Hq2 <list_append_lcons_sn #H
-  [ elim (unwind_path_inv_d_sn … H) -H #r1 #m1 #_ #_ #H0 #_ -IH
-    elim (eq_inv_list_empty_append … H0) -H0 #_ #H0 destruct
-    elim Hq2 -Hq2 //
-  | elim (unwind_path_inv_m_sn … H)
-  | elim (unwind_path_inv_L_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗟◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  | elim (unwind_path_inv_A_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗔◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  | elim (unwind_path_inv_S_sn … H) -H #r1 #s1 #Hr1 #Hs1 #H0 destruct
-    elim (IH … Hs1) -IH -Hs1 // -Hq2 #p1 #p2 #H1 #H2 #H3 destruct
-    @(ex3_2_intro … (r1●𝗦◗p1)) //
-    <structure_append <Hr1 -Hr1 //
-  ]
-]
-qed-.
-
-(* Inversions with inner condition for path *********************************)
-
-lemma unwind_inv_append_inner_sn (q1) (q2) (p) (f):
-      q1 ϵ 𝐈 → q1●q2 = ▼[f]p →
-      ∃∃p1,p2. ⊗p1 = q1 & ▼[▼[p1]f]p2 = q2 & p1●p2 = p.
-#q1 @(list_ind_rcons … q1) -q1
-[ #q2 #p #f #Hq1 <list_append_empty_sn #H destruct
-  /2 width=5 by ex3_2_intro/
-| #q1 * [ #n1 ] #_ #q2 #p #f #Hq2
-  [ elim (pic_inv_d_dx … Hq2)
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_m_sn … (sym_eq … H2))
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_L_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗟)) [1,3: // ]
-    [ <structure_append <structure_L_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_A_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗔)) [1,3: // ]
-    [ <structure_append <structure_A_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  | <list_append_rcons_sn #H0
-    elim (unwind_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
-    elim (unwind_path_inv_S_sn … (sym_eq … H2)) -H2 #r2 #s2 #Hr2 #Hs2 #H0 destruct
-    @(ex3_2_intro … (p1●r2◖𝗦)) [1,3: // ]
-    [ <structure_append <structure_S_dx <Hr2 -Hr2 //
-    | <list_append_assoc <list_append_rcons_sn //
-    ]
-  ]
-]
-qed-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure_depth.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind3/unwind_structure_depth.etc
deleted file mode 100644 (file)
index fdb7d78..0000000
+++ /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/unwind3/unwind.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/arith/nat_pred_succ.ma".
-
-(* UNWIND FOR PATH *********************************************************)
-
-(* Basic constructions with structure and depth ****************************)
-
-lemma unwind_rmap_structure (p) (f):
-      (⫯*[❘p❘]f) = ▼[⊗p]f.
-#p elim p -p //
-* [ #n ] #p #IH #f //
-[ <unwind_rmap_A_sn //
-| <unwind_rmap_S_sn //
-]
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/unwind/unwind2_rmap_depth.ma b/matita/matita/contribs/lambdadelta/delayed_updating/unwind/unwind2_rmap_depth.ma
new file mode 100644 (file)
index 0000000..cb80667
--- /dev/null
@@ -0,0 +1,58 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM 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/unwind/unwind2_path_append.ma".
+include "delayed_updating/syntax/path_depth.ma".
+include "ground/relocation/tr_id_compose.ma".
+include "ground/relocation/tr_compose_compose.ma".
+include "ground/relocation/tr_compose_eq.ma".
+include "ground/relocation/xap.ma".
+include "ground/lib/stream_eq_eq.ma".
+
+(* TAILED UNWIND FOR RELOCATION MAP *****************************************)
+
+(* Constructions with depth *************************************************)
+
+lemma unwind2_rmap_decompose (p) (f):
+      ▶[f]p ≗ (⫯*[♭p]f)∘(▶[𝐢]p).
+#p elim p -p
+[ #f <unwind2_rmap_empty <unwind2_rmap_empty <tr_pushs_zero //
+| * [ #k ] #p #IH #f //
+  [ <unwind2_rmap_d_dx <unwind2_rmap_d_dx <depth_d_dx
+    @(stream_eq_trans … (tr_compose_assoc …))
+    /2 width=1 by tr_compose_eq_repl/
+  | <unwind2_rmap_L_dx <unwind2_rmap_L_dx <depth_L_dx
+    <tr_pushs_succ <tr_compose_push_bi
+    /2 width=1 by tr_push_eq_repl/
+  ]
+]
+qed.
+
+lemma unwind2_rmap_pap_le (f) (p) (h):
+      ▶[𝐢]p@⧣❨h❩ < ↑♭p → ▶[𝐢]p@⧣❨h❩ = ▶[f]p@⧣❨h❩.
+#f #p #h #Hh
+>(tr_pap_eq_repl … (▶[f]p) … (unwind2_rmap_decompose …))
+<tr_compose_pap <tr_pap_pushs_le //
+qed.
+
+lemma unwind2_rmap_xap_le (f) (p) (n):
+      ▶[𝐢]p@❨n❩ ≤ ♭p → ▶[𝐢]p@❨n❩ = ▶[f]p@❨n❩.
+(*
+#f #p * // #h <tr_xap_ninj #Hh
+>unwind2_rmap_pap_le
+*)
+#f #p #n #Hn
+>(tr_xap_eq_repl … (▶[f]p) … (unwind2_rmap_decompose …))
+<tr_compose_xap <tr_xap_pushs_le //
+qed-.
index 853f07e9cea4a1ef87954d7cd6e82221d6d26c00..7b568dc8681403e78308694f011135ba0dcf6175 100644 (file)
@@ -89,12 +89,12 @@ lemma unwind2_rmap_head_xap (f) (p) (n):
 ]
 qed.
 
-lemma unwind2_rmap_append_pap_closed (f) (p) (q) (n:pnat):
-      q = ↳[n]q →
-      ♭q = ninj (▶[f](p●q)@⧣❨n❩).
-#f #p #q #n #Hn
->tr_xap_ninj >(path_head_refl_append p … Hn) in ⊢ (??%?);
->(unwind2_rmap_head_xap_closed … Hn) -Hn
+lemma unwind2_rmap_append_pap_closed (f) (p) (q) (h:pnat):
+      q = ↳[h]q →
+      ♭q = ninj (▶[f](p●q)@⧣❨h❩).
+#f #p #q #h #Hh
+>tr_xap_ninj >(path_head_refl_append p … Hh) in ⊢ (??%?);
+>(unwind2_rmap_head_xap_closed … Hh) -Hh
 <path_head_depth //
 qed.