(* NOTATION FOR DELAYED UPDATING ********************************************)
notation "hvbox( p ᴿ )"
- non associative with precedence 70
+ non associative with precedence 71
for @{ 'NEcR $p }.
(* *)
(**************************************************************************)
-include "delayed_updating/unwind1/unwind.ma".
include "delayed_updating/substitution/fsubst.ma".
include "delayed_updating/syntax/prototerm_constructors.ma".
include "delayed_updating/syntax/prototerm_equivalence.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_balanced.ma".
+include "delayed_updating/syntax/path_head.ma".
include "delayed_updating/syntax/path_depth.ma".
+include "delayed_updating/syntax/path_reverse.ma".
include "delayed_updating/notation/relations/black_rightarrow_df_4.ma".
-include "ground/xoa/ex_1_2.ma".
-include "ground/xoa/and_4.ma".
(* DELAYED FOCUSED REDUCTION ************************************************)
definition dfr (p) (q): relation2 prototerm prototerm ≝
- λt1,t2. ∃∃b,n.
- let r ≝ p●𝗔◗b●𝗟◗q in
- ∧∧ (⊗b ϵ 𝐁 ∧ 𝟎 = ♭b) & ↑♭q = (▼[r]𝐢)@⧣❨n❩ & r◖𝗱n ϵ t1 &
- t1[⋔r←𝛗(n+♭b).(t1⋔(p◖𝗦))] ⇔ t2
+ λt1,t2. ∃n:pnat.
+ let r ≝ p●𝗔◗𝗟◗q in
+ ∧∧ (𝗟◗q)ᴿ = ↳[n](rᴿ) & r◖𝗱n ϵ t1 &
+ t1[⋔r←𝛗n.(t1⋔(p◖𝗦))] ⇔ t2
.
interpretation
include "delayed_updating/reduction/dfr.ma".
include "delayed_updating/reduction/ifr.ma".
-include "delayed_updating/unwind1/unwind_fsubst.ma".
-include "delayed_updating/unwind1/unwind_constructors.ma".
-include "delayed_updating/unwind1/unwind_preterm_eq.ma".
-include "delayed_updating/unwind1/unwind_structure_depth.ma".
-include "delayed_updating/unwind1/unwind_depth.ma".
+include "delayed_updating/unwind/unwind2_constructors.ma".
+include "delayed_updating/unwind/unwind2_preterm_fsubst.ma".
+include "delayed_updating/unwind/unwind2_preterm_eq.ma".
+include "delayed_updating/unwind/unwind2_prototerm_lift.ma".
+include "delayed_updating/unwind/unwind2_rmap_head.ma".
include "delayed_updating/substitution/fsubst_eq.ma".
include "delayed_updating/substitution/lift_prototerm_eq.ma".
+
include "delayed_updating/syntax/prototerm_proper_constructors.ma".
+include "delayed_updating/syntax/path_head_structure.ma".
include "delayed_updating/syntax/path_structure_depth.ma".
-include "ground/relocation/tr_uni_compose.ma".
-include "ground/relocation/tr_pap_pushs.ma".
+include "delayed_updating/syntax/path_structure_reverse.ma".
+include "delayed_updating/syntax/path_depth_reverse.ma".
(* DELAYED FOCUSED REDUCTION ************************************************)
-(* COMMENT
-axiom pippo (b) (q) (n):
- ↑❘q❘ = (↑[q]𝐢)@⧣❨n❩ →
- ↑❘q❘+❘b❘= (↑[b●𝗟◗q]𝐢)@⧣❨n+❘b❘❩.
-
-lemma unwind_rmap_tls_eq_id (p) (n):
- ❘p❘ = ↑[p]𝐢@⧣❨n❩ →
- (𝐢) ≗ ⇂*[n]↑[p]𝐢.
-#p @(list_ind_rcons … p) -p
-[ #n <depth_empty #H destruct
-| #p * [ #m ] #IH #n
- [ <depth_d_dx <unwind_rmap_pap_d_dx #H0
- @(stream_eq_trans … (unwind_rmap_tls_d_dx …))
- @(stream_eq_trans … (IH …)) -IH //
- | /2 width=1 by/
- | <depth_L_dx <unwind_rmap_L_dx
- cases n -n [| #n ] #H0
- [
- |
- ]
- | /2 width=1 by/
- | /2 width=1 by/
- ]
-]
-
-
-(* (↑❘q❘+❘b❘=↑[b●𝗟◗q]𝐢@⧣❨n+❘b❘❩ *)
-(* [↑[p]𝐢@⧣❨n❩]⫯*[❘p❘]f∘⇂*[n]↑[p]𝐢) *)
-lemma unwind_rmap_tls_eq (f) (p) (n):
- ❘p❘ = ↑[p]𝐢@⧣❨n❩ →
- f ≗ ⇂*[n]↑[p]f.
-#f #p #n #Hp
-@(stream_eq_canc_dx … (stream_tls_eq_repl …))
-[| @unwind_rmap_decompose | skip ]
-<tr_compose_tls <Hp
+(* Main destructions with ifr ***********************************************)
-@(stream_eq_canc_dx) … (unwind_rmap_decompose …))
-
-*)
-lemma dfr_unwind_id_bi (p) (q) (t1) (t2): t1 ϵ 𝐓 →
- t1 ➡𝐝𝐟[p,q] t2 → ▼[𝐢]t1 ➡𝐟[⊗p,⊗q] ▼[𝐢]t2.
-#p #q #t1 #t2 #H0t1
-* #b #n * #Hb #Hn #Ht1 #Ht2
-@(ex1_2_intro … (⊗b) (↑♭⊗q)) @and4_intro
-[ //
-| (*//*)
-| lapply (in_comp_unwind_bi (𝐢) … Ht1) -Ht1 -H0t1 -Hb -Ht2
- <unwind_path_d_empty_dx <depth_structure //
-| lapply (unwind_term_eq_repl_dx (𝐢) … Ht2) -Ht2 #Ht2
+theorem dfr_des_ifr (f) (p) (q) (t1) (t2): t1 ϵ 𝐓 →
+ t1 ➡𝐝𝐟[p,q] t2 → ▼[f]t1 ➡𝐟[⊗p,⊗q] ▼[f]t2.
+#f #p #q #t1 #t2 #H0t1
+* #n * #H1n #Ht1 #Ht2
+@(ex_intro … (↑♭⊗q)) @and3_intro
+[ -H0t1 -H1n -Ht1 -Ht2
+ >list_append_rcons_sn <reverse_append
+ >nsucc_unfold >depth_reverse >depth_L_dx >reverse_lcons
+ >structure_L_sn >structure_reverse
+ <path_head_structure //
+| lapply (in_comp_unwind2_path_term f … Ht1) -Ht2 -Ht1 -H0t1
+ <unwind2_path_d_dx <depth_structure
+ >list_append_rcons_sn in H1n; <reverse_append #H1n
+ lapply (unwind2_rmap_append_pap_closed f … H1n)
+ <reverse_lcons <depth_L_dx #H2n
+ lapply (eq_inv_ninj_bi … H2n) -H2n #H2n <H2n -H2n -H1n #Ht1 //
+| lapply (unwind2_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
@(subset_eq_trans … Ht2) -t2
- @(subset_eq_trans … (unwind_fsubst …))
- [ (*<unwind_rmap_append <unwind_rmap_A_sn <unwind_rmap_append <unwind_rmap_L_sn *)
- <structure_append <structure_A_sn <structure_append <structure_L_sn
- <depth_append <depth_L_sn <depth_structure <depth_structure
- @fsubst_eq_repl [ // ]
- @(subset_eq_trans … (unwind_iref …))
-
- elim Hb -Hb #Hb #H0 <H0 -H0 <nrplus_zero_dx <nplus_zero_dx <nsucc_unfold
- >Hn
+ @(subset_eq_trans … (unwind2_term_fsubst …))
+ [ @fsubst_eq_repl [ // | // ]
+ @(subset_eq_trans … (unwind2_term_iref …))
@(subset_eq_canc_sn … (lift_term_eq_repl_dx …))
- [ @unwind_grafted_S /2 width=2 by ex_intro/ | skip ]
- <Hn <Hn
-(*
- @(subset_eq_trans … (lift_term_eq_repl_dx …))
- [ @(unwind_term_eq_repl_sn … (tls_succ_unwind q …)) | skip ]
-*)
-(*
-
- @subset_eq_trans
- [2: @unwind_term_eq_repl_dx
- @(subset_eq_canc_sn … (unwind_term_eq_repl_dx …))
-
- @(subset_eq_canc_sn … (unwind_term_eq_repl_dx …))
- [ @unwind_grafted_S /2 width=2 by ex_intro/ | skip ]
-
- @(subset_eq_trans … (unwind_term_after …))
- @(subset_eq_canc_dx … (unwind_term_after …))
- @unwind_term_eq_repl_sn -t1
- @(stream_eq_trans … (tr_compose_uni_dx …))
- lapply (Hn (𝐢)) -Hn >tr_id_unfold #Hn
- lapply (pippo … b … Hn) -Hn #Hn
- @tr_compose_eq_repl
- [ <unwind_rmap_pap_le //
- <Hn <nrplus_inj_sn //
- |
- ]
-*)
+ [ @unwind2_term_grafted_S /2 width=2 by ex_intro/ | skip ] -Ht1
+ @(subset_eq_trans … (unwind2_lift_term_after …))
+ @unwind2_term_eq_repl_sn
+(* Note: crux of the proof begins *)
+ @nstream_eq_inv_ext #m
+ <tr_compose_pap <tr_compose_pap
+ <tr_uni_pap <tr_uni_pap <tr_pap_plus
+ >list_append_rcons_sn in H1n; <reverse_append #H1n
+ lapply (unwind2_rmap_append_pap_closed f … H1n) #H2n
+ >nrplus_inj_dx in ⊢ (???%); <H2n -H2n
+ lapply (tls_unwind2_rmap_append_closed f … H1n) #H2n
+ <(tr_pap_eq_repl … H2n) -H2n -H1n //
+(* Note: crux of the proof ends *)
| //
| /2 width=2 by ex_intro/
| //
]
]
-
-(*
-Hn : ↑❘q❘ = ↑[p●𝗔◗b●𝗟◗q]𝐢@⧣❨n❩
----------------------------
-↑[𝐮❨↑❘q❘+❘b❘❩] ↑[↑[p]𝐢] t ⇔ ↑[𝐮❨↑[p●𝗔◗b●𝗟◗q]𝐢@⧣❨n+❘b❘❩❩] t
-*)
-(*
-(↑[𝐮❨↑❘q❘❩]▼[⇂*[↑❘q❘]▼[p●𝗟◗q]𝐢](t1⋔(p◖𝗦))⇔▼[𝐮❨↑❘q❘❩∘▼[p●𝗔◗b●𝗟◗q]𝐢](t1⋔(p◖𝗦))
-*)
+qed.
(* *)
(**************************************************************************)
-include "delayed_updating/unwind1/unwind_prototerm.ma".
include "delayed_updating/substitution/fsubst.ma".
include "delayed_updating/substitution/lift_prototerm.ma".
include "delayed_updating/syntax/prototerm_equivalence.ma".
+include "delayed_updating/syntax/path_head.ma".
include "delayed_updating/syntax/path_depth.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_balanced.ma".
+include "delayed_updating/syntax/path_reverse.ma".
include "delayed_updating/notation/relations/black_rightarrow_f_4.ma".
-include "ground/xoa/ex_1_2.ma".
-include "ground/xoa/and_4.ma".
(* IMMEDIATE FOCUSED REDUCTION ************************************************)
definition ifr (p) (q): relation2 prototerm prototerm ≝
- λt1,t2. ∃∃b,n.
- let r ≝ p●𝗔◗b●𝗟◗q in
- ∧∧ (⊗b ϵ 𝐁 ∧ 𝟎 = ♭b) & ↑♭q = (▼[r]𝐢)@⧣❨n❩ & r◖𝗱n ϵ t1 &
- t1[⋔r←↑[𝐮❨♭(b●𝗟◗q)❩](t1⋔(p◖𝗦))] ⇔ t2
+ λt1,t2. ∃n:pnat.
+ let r ≝ p●𝗔◗𝗟◗q in
+ ∧∧ (𝗟◗q)ᴿ = ↳[n](rᴿ) & r◖𝗱n ϵ t1 &
+ t1[⋔r←↑[𝐮❨♭(𝗟◗q)❩](t1⋔(p◖𝗦))] ⇔ t2
.
interpretation
]
qed.
-lemma fsubst_eq_repl (t1) (t2) (u1) (u2) (p):
- t1 ⇔ t2 → u1 ⇔ u2 → t1[⋔p←u1] ⇔ t2[⋔p←u2].
-#t1 #t2 #u1 #u2 #p * #H1t #H2t * #H1u #H2u
+lemma fsubst_eq_repl (t1) (t2) (u1) (u2) (p1) (p2):
+ t1 ⇔ t2 → p1 = p2 → u1 ⇔ u2 → t1[⋔p1←u1] ⇔ t2[⋔p2←u2].
+#t1 #t2 #u1 #u2 #p1 #p2 * #H1t #H2t #Hp * #H1u #H2u
/3 width=5 by conj, subset_inclusion_fsubst_bi/
qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM 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/syntax/path_depth.ma".
+include "delayed_updating/syntax/path_reverse.ma".
+
+(* DEPTH FOR PATH ***********************************************************)
+
+(* Constructions with reverse ***********************************************)
+
+lemma depth_reverse (p):
+ ♭p = ♭pᴿ.
+#p elim p -p //
+* [ #n ] #p #IH <reverse_lcons
+[ <depth_d_sn <depth_d_dx //
+| <depth_m_sn <depth_m_dx //
+| <depth_L_sn <depth_L_dx //
+| <depth_A_sn <depth_A_dx //
+| <depth_S_sn <depth_S_dx //
+]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM 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/syntax/path_head.ma".
+include "delayed_updating/syntax/path_structure.ma".
+include "delayed_updating/syntax/path_depth.ma".
+
+(* HEAD FOR PATH ************************************************************)
+
+(* Constructions with structure *********************************************)
+
+axiom path_head_structure (p) (q):
+ ⊗p = ↳[♭⊗p]((⊗p)●q).
+(*
+#p elim p -p
+[ #q <path_head_zero //
+| * [ #n ] #p #IH #q
+ [ <structure_d_sn //
+ | <structure_m_sn //
+ | <structure_L_sn
+ | <structure_A_sn <list_append_lcons_sn
+ <path_head_A_sn //
+ | <structure_S_sn //
+ ]
+]
+*)
\ No newline at end of file
#l #n
<labels_unfold <labels_unfold <niter_succ //
qed.
+
+(* Basic inversions *********************************************************)
+
+lemma eq_inv_empty_labels (l) (n):
+ (𝐞) = l∗∗n → 𝟎 = n.
+#l #n @(nat_ind_succ … n) -n //
+#n #_ <labels_succ #H0 destruct
+qed-.
#p #l
<reverse_append //
qed.
+
+(* Main constructions *******************************************************)
+
+theorem reverse_revrse (p):
+ p = pᴿᴿ.
+#p elim p -p //
+#l #p #IH
+<reverse_lcons <reverse_rcons //
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM 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/syntax/path_structure.ma".
+include "delayed_updating/syntax/path_reverse.ma".
+
+(* STRUCTURE FOR PATH *******************************************************)
+
+(* Constructions with reverse ***********************************************)
+
+lemma structure_reverse (p):
+ (⊗p)ᴿ = ⊗(pᴿ).
+#p elim p -p //
+* [ #n ] #p #IH <reverse_lcons //
+[ <structure_d_sn <structure_d_dx //
+| <structure_m_sn <structure_m_dx //
+]
+qed.
include "delayed_updating/unwind/unwind2_rmap_eq.ma".
include "delayed_updating/unwind/unwind_gen_eq.ma".
include "ground/relocation/tr_compose_compose.ma".
-include "ground/relocation/tr_compose_pn.ma".
(* UNWIND FOR PATH **********************************************************)
include "delayed_updating/unwind/unwind2_path_eq.ma".
include "delayed_updating/substitution/lift_eq.ma".
-include "ground/relocation/tr_uni_compose.ma".
(* UNWIND FOR PATH **********************************************************)
include "delayed_updating/unwind/unwind2_path.ma".
include "delayed_updating/unwind/unwind_gen_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 **********************************************************)
(**************************************************************************)
include "delayed_updating/unwind/unwind2_rmap.ma".
+include "ground/relocation/tr_uni_compose.ma".
include "ground/relocation/tr_compose_eq.ma".
include "ground/relocation/tr_pn_eq.ma".
+include "ground/arith/nat_rplus_pplus.ma".
(* UNWIND MAP FOR PATH ******************************************************)
| /2 width=1 by/
]
qed-.
+
+lemma tls_unwind2_rmap_d_sn (f) (p) (m) (n):
+ ⇂*[m+n]▶[f]p ≗ ⇂*[m]▶[f](𝗱n◗p).
+#f #p #m #n
+<unwind2_rmap_d_sn >nrplus_inj_dx
+/2 width=1 by tr_tls_compose_uni_dx/
+qed.
(**************************************************************************)
include "delayed_updating/unwind/unwind2_rmap_labels.ma".
+include "delayed_updating/unwind/unwind2_rmap_eq.ma".
include "delayed_updating/unwind/xap.ma".
include "delayed_updating/syntax/path_head_depth.ma".
+include "ground/lib/stream_eq_eq.ma".
include "ground/arith/nat_le_plus.ma".
(* UNWIND MAP FOR PATH ******************************************************)
(* Constructions with path_head *********************************************)
-lemma unwind2_rmap_head_xap_closed (f) (p) (n) (k):
- (∃q. p = (↳[n]p)●q) → k ≤ n →
+lemma unwind2_rmap_head_xap_le_closed (f) (p) (q) (n) (k):
+ p = (↳[n]p)●q → k ≤ n →
(▶[f]p)@❨k❩ = (▶[f]↳[n]p)@❨k❩.
#f #p elim p -p
-[ #n #k * #q #Hq #Hk
+[ #q #n #k #Hq #Hk
elim (eq_inv_list_empty_append … Hq) -Hq * #_ //
-| #l #p #IH #n @(nat_ind_succ … n) -n
+| #l #p #IH #q #n @(nat_ind_succ … n) -n
[ #k #_ #Hk <(nle_inv_zero_dx … Hk) -k -IH
<path_head_zero <unwind2_rmap_empty //
- | #n #_ #k cases l [ #m ] * #q
+ | #n #_ #k cases l [ #m ]
[ <path_head_d_sn <list_append_lcons_sn #Hq #Hk
elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
<unwind2_rmap_d_sn <unwind2_rmap_d_sn
<tr_compose_xap <tr_compose_xap
- @IH [ /2 width=2 by ex_intro/ ] (**) (* auto too slow *)
+ @(IH … Hq) -IH -Hq (**) (* auto too slow *)
@nle_trans [| @tr_uni_xap ]
/2 width=1 by nle_plus_bi_dx/
| <path_head_m_sn <list_append_lcons_sn #Hq #Hk
elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
<unwind2_rmap_m_sn <unwind2_rmap_m_sn
- /3 width=2 by ex_intro/
+ /2 width=2 by/
| <path_head_L_sn <list_append_lcons_sn #Hq
@(nat_ind_succ … k) -k // #k #_ #Hk
lapply (nle_inv_succ_bi … Hk) -Hk #Hk
elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
<unwind2_rmap_L_sn <unwind2_rmap_L_sn
<tr_xap_push <tr_xap_push
- /4 width=2 by ex_intro, eq_f/
+ /3 width=2 by eq_f/
| <path_head_A_sn <list_append_lcons_sn #Hq #Hk
elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
<unwind2_rmap_A_sn <unwind2_rmap_A_sn
- /3 width=2 by ex_intro/
+ /2 width=2 by/
| <path_head_S_sn <list_append_lcons_sn #Hq #Hk
elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
<unwind2_rmap_S_sn <unwind2_rmap_S_sn
- /3 width=2 by ex_intro/
+ /2 width=2 by/
]
]
]
qed-.
+lemma unwind2_rmap_head_append_xap_closed (f) (p) (q) (n):
+ p = ↳[n](p●q) →
+ ▶[f](p●q)@❨n❩ = ▶[f]↳[n](p●q)@❨n❩.
+/2 width=2 by unwind2_rmap_head_xap_le_closed/
+qed-.
+
lemma unwind2_rmap_head_xap (f) (p) (n):
n + ♯(↳[n]p) = (▶[f]↳[n]p)@❨n❩.
#f #p elim p -p
]
]
qed.
+
+lemma unwind2_rmap_append_pap_closed (f) (p) (q) (n:pnat):
+ p = ↳[n](p●q) →
+ ♭p = ninj (▶[f](p●q)@⧣❨n❩).
+#f #p #q #n #Hn
+>tr_xap_ninj >Hn in ⊢ (??%?);
+>(unwind2_rmap_head_append_xap_closed … Hn) -Hn //
+qed.
+
+lemma tls_unwind2_rmap_plus_closed (f) (p) (q) (n) (k):
+ p = (↳[n]p)●q →
+ ⇂*[k]▶[f]q ≗ ⇂*[n+k]▶[f]p.
+#f #p elim p -p
+[ #q #n #k #Hq
+ elim (eq_inv_list_empty_append … Hq) -Hq #Hn #H0 destruct
+ <path_head_empty in Hn; #Hn
+ <(eq_inv_empty_labels … Hn) -n //
+| #l #p #IH #q #n @(nat_ind_succ … n) -n //
+ #n #_ #k cases l [ #m ]
+ [ <path_head_d_sn <list_append_lcons_sn #Hq
+ elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq <nrplus_inj_sn
+ @(stream_eq_trans … (tls_unwind2_rmap_d_sn …))
+ >nrplus_inj_dx >nrplus_inj_sn >nrplus_inj_sn <nplus_plus_comm_23
+ /2 width=1 by/
+ | <path_head_m_sn <list_append_lcons_sn #Hq
+ elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
+ <unwind2_rmap_m_sn /2 width=1 by/
+ | <path_head_L_sn <list_append_lcons_sn #Hq
+ elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
+ <unwind2_rmap_L_sn <nplus_succ_sn /2 width=1 by/
+ | <path_head_A_sn <list_append_lcons_sn #Hq
+ elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
+ <unwind2_rmap_A_sn /2 width=2 by/
+ | <path_head_S_sn <list_append_lcons_sn #Hq
+ elim (eq_inv_list_lcons_bi ????? Hq) -Hq #_ #Hq
+ <unwind2_rmap_S_sn /2 width=2 by/
+ ]
+]
+qed-.
+
+lemma tls_unwind2_rmap_append_closed (f) (p) (q) (n):
+ p = ↳[n](p●q) →
+ ▶[f]q ≗ ⇂*[n]▶[f](p●q).
+/2 width=1 by tls_unwind2_rmap_plus_closed/
+qed.