--- /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/prototerm_constructors.ma".
+include "delayed_updating/syntax/prototerm_equivalence.ma".
+include "delayed_updating/substitution/fsubst.ma".
+include "delayed_updating/substitution/unwind.ma".
+include "delayed_updating/syntax/path_structure.ma".
+include "delayed_updating/syntax/path_balanced.ma".
+include "delayed_updating/syntax/path_depth.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 ϵ 𝐁 & ∀f. ↑❘q❘ = (↑[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
+ t1[⋔r←𝛗(n+❘b❘).(t1⋔(p◖𝗦))] ⇔ t2
+.
+
+interpretation
+ "focused balanced reduction with delayed updating (prototerm)"
+ 'BlackRightArrowDF t1 p q t2 = (dfr p q t1 t2).
--- /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/reduction/dfr.ma".
+include "delayed_updating/reduction/ifr.ma".
+include "delayed_updating/substitution/fsubst_unwind.ma".
+include "delayed_updating/substitution/fsubst_eq.ma".
+include "delayed_updating/substitution/unwind_constructors.ma".
+include "delayed_updating/substitution/unwind_preterm_eq.ma".
+include "delayed_updating/substitution/unwind_structure_depth.ma".
+include "delayed_updating/substitution/unwind_depth.ma".
+include "delayed_updating/syntax/prototerm_proper_constructors.ma".
+include "delayed_updating/syntax/path_structure_depth.ma".
+include "ground/relocation/tr_uni_compose.ma".
+include "ground/relocation/tr_pap_pushs.ma".
+
+(* DELAYED FOCUSED REDUCTION ************************************************)
+
+lemma tr_uni_eq_repl (n1) (n2):
+ n1 = n2 → 𝐮❨n1❩ ≗ 𝐮❨n2❩.
+// qed.
+
+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
+
+@(stream_eq_canc_dx) … (unwind_rmap_decompose …))
+
+lemma dfr_unwind_bi (f) (p) (q) (t1) (t2): t1 ϵ 𝐓 →
+ t1 ➡𝐝𝐟[p,q] t2 → ↑[f]t1 ➡𝐟[⊗p,⊗q] ↑[f]t2.
+#f #p #q #t1 #t2 #H0t1
+* #b #n * #Hb #Hn #Ht1 #Ht2
+@(ex1_2_intro … (⊗b) (↑❘⊗q❘)) @and4_intro
+[ //
+| #g <unwind_rmap_structure <depth_structure
+ >tr_pushs_swap <tr_pap_pushs_le //
+| lapply (in_comp_unwind_bi f … Ht1) -Ht1 -H0t1 -Hb -Ht2
+ <unwind_path_d_empty_dx //
+| lapply (unwind_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 …))
+ @(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 //
+ |
+ ]
+ | //
+ | /2 width=2 by ex_intro/
+ | //
+ ]
+]
--- /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/prototerm.ma".
+include "delayed_updating/notation/functions/pitchforkleftarrow_3.ma".
+
+(* FOCALIZED SUBSTITUTION ***************************************************)
+
+definition fsubst (p) (u): prototerm → prototerm ≝
+ λt,q.
+ ∨∨ ∃∃r. r ϵ u & p●r = q
+ | ∧∧ q ϵ t & (∀r. p●r = q → ⊥)
+.
+
+interpretation
+ "focalized substitution (prototerm)"
+ 'PitchforkLeftArrow t p u = (fsubst p u t).
--- /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/substitution/fsubst.ma".
+include "delayed_updating/syntax/prototerm_equivalence.ma".
+
+(* Constructions with subset_equivalence ************************************)
+
+lemma subset_inclusion_fsubst_bi (t1) (t2) (u1) (u2) (p):
+ t1 ⊆ t2 → u1 ⊆ u2 → t1[⋔p←u1] ⊆ t2[⋔p←u2].
+#t1 #t2 #u1 #u2 #p #Ht #Hu #q * *
+[ #r #Hr #H0 destruct
+ /4 width=3 by ex2_intro, or_introl/
+| /4 width=2 by or_intror, conj/
+]
+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
+/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/substitution/fsubst.ma".
-include "delayed_updating/substitution/lift_prototerm_eq.ma".
-include "delayed_updating/substitution/lift_structure.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/syntax/prototerm_proper.ma".
-
-(* FOCALIZED SUBSTITUTION ***************************************************)
-
-lemma lift_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
- >lift_append_proper_dx
- /4 width=5 by in_comp_lift_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 (lift_path_root f q) #r #_ #Hr /2 width=2 by/
- | #l #r #H2 destruct
- @H0 -H0 [| <lift_append_proper_dx /2 width=3 by ppc_lcons/ ]
- ]
-]
-qed-.
-
-lemma lift_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
- [|| <lift_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_lift_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 (lift_inv_append_proper_dx … Hr) -Hr // #s1 #s2 #Hs1 #_ #H1 destruct
- lapply (H2p … Hs1) -H2p -Hs1 /2 width=2 by ex_intro/
- ]
-]
-qed-.
-
-lemma lift_fsubst (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
- (↑[f]t)[⋔(⊗p)←↑[↑[p]f]u] ⇔ ↑[f](t[⋔p←u]).
-/4 width=3 by lift_fsubst_sn, conj, lift_fsubst_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/substitution/fsubst.ma".
+include "delayed_updating/substitution/unwind_prototerm.ma".
+include "delayed_updating/syntax/prototerm_equivalence.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/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 ϵ 𝐁 & ∀f. ↑❘q❘ = (↑[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
+ t1[⋔r←↑[𝐮❨❘b●𝗟◗q❘❩](t1⋔(p◖𝗦))] ⇔ t2
+.
+
+interpretation
+ "focused balanced reduction with immediate updating (prototerm)"
+ 'BlackRightArrowF t1 p q t2 = (ifr p q t1 t2).
+++ /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/substitution/lift.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".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Constructions with depth ************************************************)
-
-lemma lift_rmap_decompose (p) (f):
- ↑[p]f ≗ (⫯*[❘p❘]f)∘(↑[p]𝐢).
-#p @(list_ind_rcons … p) -p
-[ #f <lift_rmap_empty <lift_rmap_empty <tr_pushs_zero //
-| #p * [ #n ] #IH #f //
- [ <lift_rmap_d_dx <lift_rmap_d_dx <depth_d_dx
- @(stream_eq_trans … (tr_compose_assoc …))
- /2 width=1 by tr_compose_eq_repl/
- | <lift_rmap_L_dx <lift_rmap_L_dx <depth_L_dx
- <tr_pushs_succ <tr_compose_push_bi
- /2 width=1 by tr_push_eq_repl/
- ]
-]
-qed.
-
-lemma lift_rmap_pap_le (f) (p) (n):
- (↑[p]𝐢)@❨n❩ < ↑❘p❘ → (↑[p]𝐢)@❨n❩ = (↑[p]f)@❨n❩.
-#f #p #n #Hn
->(tr_pap_eq_repl … (↑[p]f) … (lift_rmap_decompose …))
-<tr_compose_pap @tr_pap_pushs_le //
-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 "ground/lib/subset_equivalence.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/substitution/lift_structure.ma".
-include "delayed_updating/substitution/lift_prototerm.ma".
-
-lemma pippo (p):
- ⊗p ϵ 𝐈.
-#p @(list_ind_rcons … p) -p
-[ <structure_empty //
-| #p * [ #n ] #IH
- [ <structure_d_dx //
- | <structure_m_dx //
- | <structure_L_dx //
- | <structure_A_dx //
- | <structure_S_dx //
- ]
-]
-qed.
-
-(* LIFT FOR PRETERM *********************************************************)
-
-(* Constructions with subset_equivalence ************************************)
-
-lemma lift_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
-<lift_append_inner_sn //
-qed-.
-
-lemma lift_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 (lift_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_lift_bi/
-qed-.
-
-lemma lift_grafted (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
- ↑[↑[p]f](t⋔p) ⇔ (↑[f]t)⋔(⊗p).
-/3 width=1 by lift_grafted_sn, conj, lift_grafted_dx/ qed.
-
-
-lemma lift_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 (lift_inv_append_proper_dx … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-elim (lift_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_lift_bi/
-qed-.
-
-lemma lift_grafted_S (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
- ↑[↑[p]f](t⋔(p◖𝗦)) ⇔ (↑[f]t)⋔((⊗p)◖𝗦).
-#f #t #p #Hp #Ht
-@conj
-[ >lift_rmap_S_dx >structure_S_dx
- @lift_grafted_sn //
-| /2 width=1 by lift_grafted_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/substitution/lift_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".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Basic constructions with structure **************************************)
-
-lemma structure_lift (p) (f):
- ⊗p = ⊗↑[f]p.
-#p @(path_ind_lift … p) -p // #p #IH #f
-<lift_path_L_sn //
-qed.
-
-lemma lift_structure (p) (f):
- ⊗p = ↑[f]⊗p.
-#p @(path_ind_lift … p) -p //
-qed.
-
-(* Destructions with structure **********************************************)
-
-lemma lift_des_structure (q) (p) (f):
- ⊗q = ↑[f]p → ⊗q = ⊗p.
-// qed-.
-
-(* Constructions with proper condition for path *****************************)
-
-lemma lift_append_proper_dx (p2) (p1) (f): p2 ϵ 𝐏 →
- (⊗p1)●(↑[↑[p1]f]p2) = ↑[f](p1●p2).
-#p2 #p1 @(path_ind_lift … p1) -p1 //
-[ #n | #n #l #p1 |*: #p1 ] #IH #f #Hp2
-[ elim (ppc_inv_lcons … Hp2) -Hp2 #l #q #H destruct //
-| <lift_path_d_lcons_sn <IH //
-| <lift_path_m_sn <IH //
-| <lift_path_L_sn <IH //
-| <lift_path_A_sn <IH //
-| <lift_path_S_sn <IH //
-]
-qed-.
-
-(* Constructions with inner condition for path ******************************)
-
-lemma lift_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 <lift_append_proper_dx //
-| <list_append_rcons_sn <lift_append_proper_dx //
- <structure_L_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <lift_append_proper_dx //
- <structure_A_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <lift_append_proper_dx //
- <structure_S_dx <list_append_rcons_sn //
-]
-qed-.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma lift_path_d_empty_dx (n) (p) (f):
- (⊗p)◖𝗱((↑[p]f)@❨n❩) = ↑[f](p◖𝗱n).
-#n #p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_m_dx (p) (f):
- ⊗p = ↑[f](p◖𝗺).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_L_dx (p) (f):
- (⊗p)◖𝗟 = ↑[f](p◖𝗟).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_A_dx (p) (f):
- (⊗p)◖𝗔 = ↑[f](p◖𝗔).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_S_dx (p) (f):
- (⊗p)◖𝗦 = ↑[f](p◖𝗦).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_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 lift_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_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct -IH
- /2 width=5 by ex4_2_intro/
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
- /2 width=5 by ex4_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
- /2 width=5 by ex4_2_intro/
-| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_m_sn (q) (p) (f):
- (𝗺◗q) = ↑[f]p → ⊥.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H /2 width=2 by/
-| <lift_path_m_sn #H /2 width=2 by/
-| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_L_sn (q) (p) (f):
- (𝗟◗q) = ↑[f]p →
- ∃∃r1,r2. 𝐞 = ⊗r1 & q = ↑[⫯↑[r1]f]r2 & r1●𝗟◗r2 = p.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_L_sn #H destruct -IH
- /2 width=5 by ex3_2_intro/
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_A_sn (q) (p) (f):
- (𝗔◗q) = ↑[f]p →
- ∃∃r1,r2. 𝐞 = ⊗r1 & q = ↑[↑[r1]f]r2 & r1●𝗔◗r2 = p.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct -IH
- /2 width=5 by ex3_2_intro/
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_S_sn (q) (p) (f):
- (𝗦◗q) = ↑[f]p →
- ∃∃r1,r2. 𝐞 = ⊗r1 & q = ↑[↑[r1]f]r2 & r1●𝗦◗r2 = p.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct -IH
- /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Inversions with proper condition for path ********************************)
-
-lemma lift_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 (lift_path_inv_d_sn … H) -H #r1 #m1 #_ #_ #H0 #_ -IH
- elim (eq_inv_list_empty_append … H0) -H0 #_ #H0 destruct
- elim Hq2 -Hq2 //
- | elim (lift_path_inv_m_sn … H)
- | elim (lift_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 (lift_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 (lift_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 lift_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 (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_path_inv_m_sn … (sym_eq … H2))
- | <list_append_rcons_sn #H0
- elim (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_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 (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_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 (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_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-.
+++ /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/substitution/lift.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/arith/nat_pred_succ.ma".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Basic constructions with structure and depth ****************************)
-
-lemma lift_rmap_structure (p) (f):
- (⫯*[❘p❘]f) = ↑[⊗p]f.
-#p elim p -p //
-* [ #n ] #p #IH #f //
-[ <lift_rmap_A_sn //
-| <lift_rmap_S_sn //
-]
-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/substitution/lift_eq.ma".
-include "delayed_updating/syntax/path_update.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/lib/stream_eq_eq.ma".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Constructions with update ***********************************************)
-
-lemma lift_rmap_pap_le (f1) (f2) (p) (m:pnat) (l:nat):
- ninj (m+⧣p+l) = ❘p❘ →
- (↑[p]f1)@❨m❩ = (↑[p]f2)@❨m❩.
-#f1 #f2 #p @(list_ind_rcons … p) -p
-[ #m #l <depth_empty #H0 destruct
-| #p * [ #n ] #IH #m #l
- [ <update_d_dx <depth_d_dx <lift_rmap_pap_d_dx <lift_rmap_pap_d_dx
- <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
- #H0 <(IH … l) -IH //
- | /2 width=2 by/
- | <update_L_dx <depth_L_dx <lift_rmap_L_dx <lift_rmap_L_dx
- cases m -m // #m
- <nrplus_succ_sn <nrplus_succ_sn >nsucc_inj #H0
- <tr_pap_push <tr_pap_push
- <(IH … l) -IH //
- | /2 width=2 by/
- | /2 width=2 by/
- ]
-]
-qed.
-
-lemma lift_rmap_pap_gt (f) (p) (m):
- f@❨m+⧣p❩+❘p❘ = (↑[p]f)@❨m+❘p❘❩.
-#f #p @(list_ind_rcons … p) -p [ // ]
-#p * [ #n ] #IH #m
-[ <update_d_dx <depth_d_dx
- <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
- <lift_rmap_pap_d_dx >IH -IH //
-| //
-| <update_L_dx <depth_L_dx
- <nrplus_succ_dx <nrplus_succ_dx <lift_rmap_L_dx <tr_pap_push //
-| //
-| //
-]
-qed.
-
-lemma lift_rmap_tls_gt (f) (p) (m:pnat):
- ⇂*[ninj (m+⧣p)]f ≗ ⇂*[ninj (m+❘p❘)]↑[p]f.
-#f #p @(list_ind_rcons … p) -p [ // ]
-#p * [ #n ] #IH #m
-[ <update_d_dx <depth_d_dx
- <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
- @(stream_eq_trans … (lift_rmap_tls_d_dx …))
- @(stream_eq_canc_dx … (IH …)) -IH //
-| //
-| <update_L_dx <depth_L_dx
- <nrplus_succ_dx >nsucc_inj //
-| //
-| //
-]
-qed.
-
-lemma lift_rmap_tls_eq (f) (p):
- ⇂*[⧣p]f ≗ ⇂*[↑❘p❘]↑[p]⫯f.
-(*
-#f #p @(list_ind_rcons … p) -p //
-#p * [ #n ] #IH //
-<update_d_dx <depth_d_dx <lift_rmap_d_dx
-@(stream_eq_trans … (tr_tls_compose_uni_dx …))
-*)
lemma pr_pat_after_uni_tls (i2) (i1):
∀f2. @❨i1, f2❩ ≘ i2 →
∀f. 𝐮❨i2❩ ⊚ ⫰*[i2] f2 ≘ f → f2 ⊚ 𝐮❨i1❩ ≘ f.
+
+. replace.sh . "/substitution/" "/unwind/"
+. replace.sh . ↑ ▼
(* *)
(**************************************************************************)
+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/substitution/fsubst.ma".
-include "delayed_updating/substitution/lift.ma".
include "delayed_updating/syntax/path_structure.ma".
include "delayed_updating/syntax/path_balanced.ma".
include "delayed_updating/syntax/path_depth.ma".
definition dfr (p) (q): relation2 prototerm prototerm ≝
λt1,t2. ∃∃b,n.
let r ≝ p●𝗔◗b●𝗟◗q in
- â\88§â\88§ â\8a\97b ϵ ð\9d\90\81 & â\88\80f. â\86\91â\9d\98qâ\9d\98 = (â\86\91[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
+ â\88§â\88§ â\8a\97b ϵ ð\9d\90\81 & â\86\91â\9d\98qâ\9d\98 = (â\96¼[r]ð\9d\90¢)@❨n❩ & r◖𝗱n ϵ t1 &
t1[⋔r←𝛗(n+❘b❘).(t1⋔(p◖𝗦))] ⇔ t2
.
include "delayed_updating/reduction/dfr.ma".
include "delayed_updating/reduction/ifr.ma".
-include "delayed_updating/substitution/fsubst_lift.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/substitution/fsubst_eq.ma".
-include "delayed_updating/substitution/lift_constructors.ma".
-include "delayed_updating/substitution/lift_preterm_eq.ma".
-include "delayed_updating/substitution/lift_structure_depth.ma".
-include "delayed_updating/substitution/lift_depth.ma".
include "delayed_updating/syntax/prototerm_proper_constructors.ma".
include "delayed_updating/syntax/path_structure_depth.ma".
include "ground/relocation/tr_uni_compose.ma".
(* DELAYED FOCUSED REDUCTION ************************************************)
-lemma tr_uni_eq_repl (n1) (n2):
- n1 = n2 → 𝐮❨n1❩ ≗ 𝐮❨n2❩.
-// qed.
-
+(* COMMENT
axiom pippo (b) (q) (n):
↑❘q❘ = (↑[q]𝐢)@❨n❩ →
↑❘q❘+❘b❘= (↑[b●𝗟◗q]𝐢)@❨n+❘b❘❩.
-lemma lift_rmap_tls_eq_id (p) (n):
+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 <lift_rmap_pap_d_dx #H0
- @(stream_eq_trans … (lift_rmap_tls_d_dx …))
+ [ <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 <lift_rmap_L_dx
+ | <depth_L_dx <unwind_rmap_L_dx
cases n -n [| #n ] #H0
[
|
(* (↑❘q❘+❘b❘=↑[b●𝗟◗q]𝐢@❨n+❘b❘❩ *)
(* [↑[p]𝐢@❨n❩]⫯*[❘p❘]f∘⇂*[n]↑[p]𝐢) *)
-lemma lift_rmap_tls_eq (f) (p) (n):
+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 …))
-[| @lift_rmap_decompose | skip ]
+[| @unwind_rmap_decompose | skip ]
<tr_compose_tls <Hp
-@(stream_eq_canc_dx) … (lift_rmap_decompose …))
+@(stream_eq_canc_dx) … (unwind_rmap_decompose …))
-lemma dfr_lift_bi (f) (p) (q) (t1) (t2): t1 ϵ 𝐓 →
- t1 ➡𝐝𝐟[p,q] t2 → ↑[f]t1 ➡𝐟[⊗p,⊗q] ↑[f]t2.
-#f #p #q #t1 #t2 #H0t1
-* #b #n * #Hb #Hn #Ht1 #Ht2
+*)
+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
[ //
-| #g <lift_rmap_structure <depth_structure
- >tr_pushs_swap <tr_pap_pushs_le //
-| lapply (in_comp_lift_bi f … Ht1) -Ht1 -H0t1 -Hb -Ht2
- <lift_path_d_empty_dx //
-| lapply (lift_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
+| //
+| 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
@(subset_eq_trans … Ht2) -t2
- @(subset_eq_trans … (lift_fsubst …))
- [ <lift_rmap_append <lift_rmap_A_sn (* <lift_rmap_append <lift_rmap_L_sn *)
+ @(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 … (lift_iref …))
- @(subset_eq_canc_sn … (lift_term_eq_repl_dx …))
- [ @lift_grafted_S /2 width=2 by ex_intro/ | skip ]
- @(subset_eq_trans … (lift_term_after …))
- @(subset_eq_canc_dx … (lift_term_after …))
- @lift_term_eq_repl_sn -t1
+ @(subset_eq_trans … (unwind_iref …))
+(*
+ @(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
- [ <lift_rmap_pap_le //
+ [ <unwind_rmap_pap_le //
<Hn <nrplus_inj_sn //
|
]
+*)
| //
| /2 width=2 by ex_intro/
| //
]
]
+
+(*
+Hn : ↑❘q❘ = ↑[p●𝗔◗b●𝗟◗q]𝐢@❨n❩
+---------------------------
+↑[𝐮❨↑❘q❘+❘b❘❩] ↑[↑[p]𝐢] t ⇔ ↑[𝐮❨↑[p●𝗔◗b●𝗟◗q]𝐢@❨n+❘b❘❩❩] t
+*)
(* *)
(**************************************************************************)
+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".
definition ifr (p) (q): relation2 prototerm prototerm ≝
λt1,t2. ∃∃b,n.
let r ≝ p●𝗔◗b●𝗟◗q in
- â\88§â\88§ â\8a\97b ϵ ð\9d\90\81 & â\88\80f. â\86\91â\9d\98qâ\9d\98 = (â\86\91[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
+ â\88§â\88§ â\8a\97b ϵ ð\9d\90\81 & â\86\91â\9d\98qâ\9d\98 = (â\96¼[r]ð\9d\90¢)@❨n❩ & r◖𝗱n ϵ t1 &
t1[⋔r←↑[𝐮❨❘b●𝗟◗q❘❩](t1⋔(p◖𝗦))] ⇔ t2
.
+++ /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/substitution/fsubst.ma".
-include "delayed_updating/substitution/lift_prototerm_eq.ma".
-include "delayed_updating/substitution/lift_structure.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/syntax/prototerm_proper.ma".
-
-(* FOCALIZED SUBSTITUTION ***************************************************)
-
-lemma lift_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
- >lift_append_proper_dx
- /4 width=5 by in_comp_lift_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 (lift_path_root f q) #r #_ #Hr /2 width=2 by/
- | #l #r #H2 destruct
- @H0 -H0 [| <lift_append_proper_dx /2 width=3 by ppc_lcons/ ]
- ]
-]
-qed-.
-
-lemma lift_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
- [|| <lift_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_lift_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 (lift_inv_append_proper_dx … Hr) -Hr // #s1 #s2 #Hs1 #_ #H1 destruct
- lapply (H2p … Hs1) -H2p -Hs1 /2 width=2 by ex_intro/
- ]
-]
-qed-.
-
-lemma lift_fsubst (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
- (↑[f]t)[⋔(⊗p)←↑[↑[p]f]u] ⇔ ↑[f](t[⋔p←u]).
-/4 width=3 by lift_fsubst_sn, conj, lift_fsubst_dx/ qed.
(* *)
(**************************************************************************)
-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/uparrow_4.ma".
include "delayed_updating/notation/functions/uparrow_2.ma".
+include "delayed_updating/syntax/path.ma".
+include "ground/relocation/tr_uni.ma".
+include "ground/relocation/tr_pap_tls.ma".
(* LIFT FOR PATH ***********************************************************)
definition lift_continuation (A:Type[0]) ≝
tr_map → path → A.
-(* Note: inner numeric labels are not liftable, so they are removed *)
rec definition lift_gen (A:Type[0]) (k:lift_continuation A) (f) (p) on p ≝
match p with
[ list_empty ⇒ k f (𝐞)
| list_lcons l q ⇒
match l with
- [ label_d n ⇒
- match q with
- [ list_empty ⇒ lift_gen (A) (λg,p. k g (𝗱(f@❨n❩)◗p)) (f∘𝐮❨n❩) q
- | list_lcons _ _ ⇒ lift_gen (A) k (f∘𝐮❨n❩) q
- ]
- | label_m ⇒ lift_gen (A) k f q
+ [ label_d n ⇒ lift_gen (A) (λg,p. k g (𝗱(f@❨n❩)◗p)) (⇂*[n]f) q
+ | label_m ⇒ lift_gen (A) (λg,p. k g (𝗺◗p)) f q
| label_L ⇒ lift_gen (A) (λg,p. k g (𝗟◗p)) (⫯f) q
| label_A ⇒ lift_gen (A) (λg,p. k g (𝗔◗p)) f q
| label_S ⇒ lift_gen (A) (λg,p. k g (𝗦◗p)) f q
k f (𝐞) = ↑{A}❨k, f, 𝐞❩.
// qed.
-lemma lift_d_empty_sn (A) (k) (n) (f):
- ↑❨(λg,p. k g (𝗱(f@❨n❩)◗p)), f∘𝐮❨ninj n❩, 𝐞❩ = ↑{A}❨k, f, 𝗱n◗𝐞❩.
-// qed.
-
-lemma lift_d_lcons_sn (A) (k) (p) (l) (n) (f):
- ↑❨k, f∘𝐮❨ninj n❩, l◗p❩ = ↑{A}❨k, f, 𝗱n◗l◗p❩.
+lemma lift_d_sn (A) (k) (p) (n) (f):
+ ↑❨(λg,p. k g (𝗱(f@❨n❩)◗p)), ⇂*[n]f, p❩ = ↑{A}❨k, f, 𝗱n◗p❩.
// qed.
lemma lift_m_sn (A) (k) (p) (f):
- ↑❨k, f, p❩ = ↑{A}❨k, f, 𝗺◗p❩.
+ ↑❨(λg,p. k g (𝗺◗p)), f, p❩ = ↑{A}❨k, f, 𝗺◗p❩.
// qed.
lemma lift_L_sn (A) (k) (p) (f):
(𝐞) = ↑[f]𝐞.
// qed.
-lemma lift_path_d_empty_sn (f) (n):
- 𝗱(f@❨n❩)◗𝐞 = ↑[f](𝗱n◗𝐞).
-// qed.
-
-lemma lift_path_d_lcons_sn (f) (p) (l) (n):
- ↑[f∘𝐮❨ninj n❩](l◗p) = ↑[f](𝗱n◗l◗p).
-// qed.
-
-lemma lift_path_m_sn (f) (p):
- ↑[f]p = ↑[f](𝗺◗p).
-// qed.
-
(* Basic constructions with proj_rmap ***************************************)
lemma lift_rmap_empty (f):
// qed.
lemma lift_rmap_d_sn (f) (p) (n):
- ↑[p](f∘𝐮❨ninj n❩) = ↑[𝗱n◗p]f.
-#f * // qed.
+ ↑[p](⇂*[ninj n]f) = ↑[𝗱n◗p]f.
+// qed.
lemma lift_rmap_m_sn (f) (p):
↑[p]f = ↑[𝗺◗p]f.
(* Advanced constructions with proj_rmap and path_rcons *********************)
lemma lift_rmap_d_dx (f) (p) (n):
- (↑[p]f)∘𝐮❨ninj n❩ = ↑[p◖𝗱n]f.
+ ⇂*[ninj n](↑[p]f) = ↑[p◖𝗱n]f.
// qed.
lemma lift_rmap_m_dx (f) (p):
// qed.
lemma lift_rmap_pap_d_dx (f) (p) (n) (m):
- ↑[p]f@❨m+n❩ = ↑[p◖𝗱n]f@❨m❩.
-#f #p #n #m
-<lift_rmap_d_dx <tr_compose_pap <tr_uni_pap //
-qed.
-
-(* Advanced eliminations with path ******************************************)
-
-lemma path_ind_lift (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-.
+ ↑[p]f@❨m+n❩ = ↑[p◖𝗱n]f@❨m❩+↑[p]f@❨n❩.
+// 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/substitution/lift_eq.ma".
+include "ground/relocation/tr_compose_pap.ma".
+include "ground/relocation/tr_compose_pn.ma".
+include "ground/relocation/tr_compose_tls.ma".
+
+(* LIFT FOR PATH ***********************************************************)
+
+(* Constructions with tr_after *********************************************)
+
+lemma lift_path_after (p) (f1) (f2):
+ ↑[f2]↑[f1]p = ↑[f2∘f1]p.
+#p elim p -p [| * ] // [ #n ] #p #IH #f1 #f2
+[ <lift_path_d_sn <lift_path_d_sn <lift_path_d_sn //
+| <lift_path_L_sn <lift_path_L_sn <lift_path_L_sn
+ >tr_compose_push_bi //
+]
+qed.
(* LIFT FOR PROTOTERM *******************************************************)
-lemma lift_iref_after_sn (f) (t:prototerm) (n:pnat):
- ↑[f∘𝐮❨n❩]t ⊆ ↑[f](𝛗n.t).
-#f #t #n #p * #q #Hq #H0 destruct
-@(ex2_intro … (𝗱n◗𝗺◗q))
+lemma lift_iref_sn (f) (t:prototerm) (n:pnat):
+ (𝛗f@❨n❩.↑[⇂*[n]f]t) ⊆ ↑[f](𝛗n.t).
+#f #t #n #p * #q * #r #Hr #H1 #H2 destruct
+@(ex2_intro … (𝗱n◗𝗺◗r))
/2 width=1 by in_comp_iref/
qed-.
-lemma lift_iref_after_dx (f) (t) (n:pnat):
- ↑[f](𝛗n.t) ⊆ ↑[f∘𝐮❨n❩]t.
+lemma lift_iref_dx (f) (t) (n:pnat):
+ ↑[f](𝛗n.t) ⊆ 𝛗f@❨n❩.↑[⇂*[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_lift_bi/
+elim (in_comp_inv_iref … Hq) -Hq #p #H0 #Hp destruct
+/3 width=1 by in_comp_iref, in_comp_lift_bi/
qed-.
-lemma lift_iref_after (f) (t) (n:pnat):
- ↑[f∘𝐮❨n❩]t ⇔ ↑[f](𝛗n.t).
-/3 width=1 by conj, lift_iref_after_sn, lift_iref_after_dx/
-qed.
-
lemma lift_iref (f) (t) (n:pnat):
- ↑[f]↑[𝐮❨n❩]t ⇔ ↑[f](𝛗n.t).
-/3 width=3 by lift_term_after, subset_eq_trans/
+ (𝛗f@❨n❩.↑[⇂*[n]f]t) ⇔ ↑[f](𝛗n.t).
+/3 width=1 by conj, lift_iref_sn, lift_iref_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/substitution/lift.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".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Constructions with depth ************************************************)
-
-lemma lift_rmap_decompose (p) (f):
- ↑[p]f ≗ (⫯*[❘p❘]f)∘(↑[p]𝐢).
-#p @(list_ind_rcons … p) -p
-[ #f <lift_rmap_empty <lift_rmap_empty <tr_pushs_zero //
-| #p * [ #n ] #IH #f //
- [ <lift_rmap_d_dx <lift_rmap_d_dx <depth_d_dx
- @(stream_eq_trans … (tr_compose_assoc …))
- /2 width=1 by tr_compose_eq_repl/
- | <lift_rmap_L_dx <lift_rmap_L_dx <depth_L_dx
- <tr_pushs_succ <tr_compose_push_bi
- /2 width=1 by tr_push_eq_repl/
- ]
-]
-qed.
-
-lemma lift_rmap_pap_le (f) (p) (n):
- (↑[p]𝐢)@❨n❩ < ↑❘p❘ → (↑[p]𝐢)@❨n❩ = (↑[p]f)@❨n❩.
-#f #p #n #Hn
->(tr_pap_eq_repl … (↑[p]f) … (lift_rmap_decompose …))
-<tr_compose_pap @tr_pap_pushs_le //
-qed.
(**************************************************************************)
include "delayed_updating/substitution/lift.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_pap_eq.ma".
include "ground/relocation/tr_pn_eq.ma".
+include "ground/lib/stream_tls_eq.ma".
(* LIFT FOR PATH ***********************************************************)
lemma lift_eq_repl (A) (p) (k1) (k2):
k1 ≗{A} k2 → stream_eq_repl … (λf1,f2. ↑❨k1, f1, p❩ = ↑❨k2, f2, p❩).
-#A #p @(path_ind_lift … p) -p [| #n #IH | #n #l0 #q #IH |*: #q #IH ]
-#k1 #k2 #f1 #f2 #Hk #Hf
+#A #p elim p -p [| * [ #n ] #q #IH ]
+#k1 #k2 #Hk #f1 #f2 #Hf
[ <lift_empty <lift_empty /2 width=1 by/
-| <lift_d_empty_sn <lift_d_empty_sn <(tr_pap_eq_repl … Hf)
- /3 width=1 by tr_compose_eq_repl, stream_eq_refl/
-| <lift_d_lcons_sn <lift_d_lcons_sn
- /3 width=1 by tr_compose_eq_repl, stream_eq_refl/
-| /2 width=1 by/
+| <lift_d_sn <lift_d_sn <(tr_pap_eq_repl … Hf)
+ /3 width=3 by stream_tls_eq_repl, compose_repl_fwd_sn/
+| /3 width=1 by/
| /3 width=1 by tr_push_eq_repl/
| /3 width=1 by/
| /3 width=1 by/
lemma lift_path_append_sn (p) (f) (q):
q●↑[f]p = ↑❨(λg,p. proj_path g (q●p)), f, p❩.
-#p @(path_ind_lift … p) -p // [ #n #l #p |*: #p ] #IH #f #q
-[ <lift_d_lcons_sn <lift_d_lcons_sn <IH -IH //
-| <lift_m_sn <lift_m_sn //
-| <lift_L_sn <lift_L_sn >lift_lcons_alt // >lift_append_rcons_sn //
- <IH <IH -IH <list_append_rcons_sn //
-| <lift_A_sn <lift_A_sn >lift_lcons_alt >lift_append_rcons_sn //
- <IH <IH -IH <list_append_rcons_sn //
-| <lift_S_sn <lift_S_sn >lift_lcons_alt >lift_append_rcons_sn //
- <IH <IH -IH <list_append_rcons_sn //
-]
+#p elim p -p // * [ #n ] #p #IH #f #q
+[ <lift_d_sn <lift_d_sn
+| <lift_m_sn <lift_m_sn
+| <lift_L_sn <lift_L_sn
+| <lift_A_sn <lift_A_sn
+| <lift_S_sn <lift_S_sn
+]
+>lift_lcons_alt // >lift_append_rcons_sn //
+<IH <IH -IH <list_append_rcons_sn //
qed.
lemma lift_path_lcons (f) (p) (l):
>lift_lcons_alt <lift_path_append_sn //
qed.
+lemma lift_path_d_sn (f) (p) (n):
+ (𝗱(f@❨n❩)◗↑[⇂*[n]f]p) = ↑[f](𝗱n◗p).
+// qed.
+
+lemma lift_path_m_sn (f) (p):
+ (𝗺◗↑[f]p) = ↑[f](𝗺◗p).
+// qed.
+
lemma lift_path_L_sn (f) (p):
(𝗟◗↑[⫯f]p) = ↑[f](𝗟◗p).
// qed.
lemma lift_path_S_sn (f) (p):
(𝗦◗↑[f]p) = ↑[f](𝗦◗p).
// qed.
-
-lemma lift_path_after (p) (f1) (f2):
- ↑[f2]↑[f1]p = ↑[f2∘f1]p.
-#p @(path_ind_lift … p) -p // [ #n #l #p | #p ] #IH #f1 #f2
-[ <lift_path_d_lcons_sn <lift_path_d_lcons_sn
- >(lift_path_eq_repl … (tr_compose_assoc …)) //
-| <lift_path_L_sn <lift_path_L_sn <lift_path_L_sn
- >tr_compose_push_bi //
-]
-qed.
+(* COMMENT
(* Advanced constructions with proj_rmap and stream_tls *********************)
<lift_rmap_d_dx >nrplus_inj_dx
/2 width=1 by tr_tls_compose_uni_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 "ground/lib/subset_equivalence.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/substitution/lift_structure.ma".
-include "delayed_updating/substitution/lift_prototerm.ma".
-
-lemma pippo (p):
- ⊗p ϵ 𝐈.
-#p @(list_ind_rcons … p) -p
-[ <structure_empty //
-| #p * [ #n ] #IH
- [ <structure_d_dx //
- | <structure_m_dx //
- | <structure_L_dx //
- | <structure_A_dx //
- | <structure_S_dx //
- ]
-]
-qed.
-
-(* LIFT FOR PRETERM *********************************************************)
-
-(* Constructions with subset_equivalence ************************************)
-
-lemma lift_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
-<lift_append_inner_sn //
-qed-.
-
-lemma lift_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 (lift_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_lift_bi/
-qed-.
-
-lemma lift_grafted (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
- ↑[↑[p]f](t⋔p) ⇔ (↑[f]t)⋔(⊗p).
-/3 width=1 by lift_grafted_sn, conj, lift_grafted_dx/ qed.
-
-
-lemma lift_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 (lift_inv_append_proper_dx … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-elim (lift_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_lift_bi/
-qed-.
-
-lemma lift_grafted_S (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
- ↑[↑[p]f](t⋔(p◖𝗦)) ⇔ (↑[f]t)⋔((⊗p)◖𝗦).
-#f #t #p #Hp #Ht
-@conj
-[ >lift_rmap_S_dx >structure_S_dx
- @lift_grafted_sn //
-| /2 width=1 by lift_grafted_S_dx/
-]
-qed.
(**************************************************************************)
include "ground/lib/subset_ext_equivalence.ma".
-include "delayed_updating/substitution/lift_eq.ma".
+include "delayed_updating/substitution/lift_after.ma".
include "delayed_updating/substitution/lift_prototerm.ma".
(* LIFT FOR PROTOTERM *******************************************************)
+++ /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/substitution/lift_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".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Basic constructions with structure **************************************)
-
-lemma structure_lift (p) (f):
- ⊗p = ⊗↑[f]p.
-#p @(path_ind_lift … p) -p // #p #IH #f
-<lift_path_L_sn //
-qed.
-
-lemma lift_structure (p) (f):
- ⊗p = ↑[f]⊗p.
-#p @(path_ind_lift … p) -p //
-qed.
-
-(* Destructions with structure **********************************************)
-
-lemma lift_des_structure (q) (p) (f):
- ⊗q = ↑[f]p → ⊗q = ⊗p.
-// qed-.
-
-(* Constructions with proper condition for path *****************************)
-
-lemma lift_append_proper_dx (p2) (p1) (f): p2 ϵ 𝐏 →
- (⊗p1)●(↑[↑[p1]f]p2) = ↑[f](p1●p2).
-#p2 #p1 @(path_ind_lift … p1) -p1 //
-[ #n | #n #l #p1 |*: #p1 ] #IH #f #Hp2
-[ elim (ppc_inv_lcons … Hp2) -Hp2 #l #q #H destruct //
-| <lift_path_d_lcons_sn <IH //
-| <lift_path_m_sn <IH //
-| <lift_path_L_sn <IH //
-| <lift_path_A_sn <IH //
-| <lift_path_S_sn <IH //
-]
-qed-.
-
-(* Constructions with inner condition for path ******************************)
-
-lemma lift_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 <lift_append_proper_dx //
-| <list_append_rcons_sn <lift_append_proper_dx //
- <structure_L_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <lift_append_proper_dx //
- <structure_A_dx <list_append_rcons_sn //
-| <list_append_rcons_sn <lift_append_proper_dx //
- <structure_S_dx <list_append_rcons_sn //
-]
-qed-.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma lift_path_d_empty_dx (n) (p) (f):
- (⊗p)◖𝗱((↑[p]f)@❨n❩) = ↑[f](p◖𝗱n).
-#n #p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_m_dx (p) (f):
- ⊗p = ↑[f](p◖𝗺).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_L_dx (p) (f):
- (⊗p)◖𝗟 = ↑[f](p◖𝗟).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_A_dx (p) (f):
- (⊗p)◖𝗔 = ↑[f](p◖𝗔).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_path_S_dx (p) (f):
- (⊗p)◖𝗦 = ↑[f](p◖𝗦).
-#p #f <lift_append_proper_dx //
-qed.
-
-lemma lift_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 lift_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_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct -IH
- /2 width=5 by ex4_2_intro/
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
- /2 width=5 by ex4_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r #h #Hr #Hh #Hq #Hp destruct
- /2 width=5 by ex4_2_intro/
-| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_m_sn (q) (p) (f):
- (𝗺◗q) = ↑[f]p → ⊥.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H /2 width=2 by/
-| <lift_path_m_sn #H /2 width=2 by/
-| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_L_sn (q) (p) (f):
- (𝗟◗q) = ↑[f]p →
- ∃∃r1,r2. 𝐞 = ⊗r1 & q = ↑[⫯↑[r1]f]r2 & r1●𝗟◗r2 = p.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_L_sn #H destruct -IH
- /2 width=5 by ex3_2_intro/
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_A_sn (q) (p) (f):
- (𝗔◗q) = ↑[f]p →
- ∃∃r1,r2. 𝐞 = ⊗r1 & q = ↑[↑[r1]f]r2 & r1●𝗔◗r2 = p.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct -IH
- /2 width=5 by ex3_2_intro/
-| <lift_path_S_sn #H destruct
-]
-qed-.
-
-lemma lift_path_inv_S_sn (q) (p) (f):
- (𝗦◗q) = ↑[f]p →
- ∃∃r1,r2. 𝐞 = ⊗r1 & q = ↑[↑[r1]f]r2 & r1●𝗦◗r2 = p.
-#q #p @(path_ind_lift … p) -p
-[| #n | #n #l #p |*: #p ] [|*: #IH ] #f
-[ <lift_path_empty #H destruct
-| <lift_path_d_empty_sn #H destruct
-| <lift_path_d_lcons_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/
-| <lift_path_m_sn #H
- elim (IH … H) -IH -H #r1 #r2 #Hr1 #Hq #Hp destruct
- /2 width=5 by ex3_2_intro/| <lift_path_L_sn #H destruct
-| <lift_path_A_sn #H destruct
-| <lift_path_S_sn #H destruct -IH
- /2 width=5 by ex3_2_intro/
-]
-qed-.
-
-(* Inversions with proper condition for path ********************************)
-
-lemma lift_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 (lift_path_inv_d_sn … H) -H #r1 #m1 #_ #_ #H0 #_ -IH
- elim (eq_inv_list_empty_append … H0) -H0 #_ #H0 destruct
- elim Hq2 -Hq2 //
- | elim (lift_path_inv_m_sn … H)
- | elim (lift_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 (lift_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 (lift_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 lift_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 (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_path_inv_m_sn … (sym_eq … H2))
- | <list_append_rcons_sn #H0
- elim (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_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 (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_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 (lift_inv_append_proper_dx … H0) -H0 // #p1 #p2 #H1 #H2 #H3 destruct
- elim (lift_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-.
+++ /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/substitution/lift.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/arith/nat_pred_succ.ma".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Basic constructions with structure and depth ****************************)
-
-lemma lift_rmap_structure (p) (f):
- (⫯*[❘p❘]f) = ↑[⊗p]f.
-#p elim p -p //
-* [ #n ] #p #IH #f //
-[ <lift_rmap_A_sn //
-| <lift_rmap_S_sn //
-]
-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/substitution/lift_eq.ma".
-include "delayed_updating/syntax/path_update.ma".
-include "delayed_updating/syntax/path_depth.ma".
-include "ground/lib/stream_eq_eq.ma".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Constructions with update ***********************************************)
-
-lemma lift_rmap_pap_le (f1) (f2) (p) (m:pnat) (l:nat):
- ninj (m+⧣p+l) = ❘p❘ →
- (↑[p]f1)@❨m❩ = (↑[p]f2)@❨m❩.
-#f1 #f2 #p @(list_ind_rcons … p) -p
-[ #m #l <depth_empty #H0 destruct
-| #p * [ #n ] #IH #m #l
- [ <update_d_dx <depth_d_dx <lift_rmap_pap_d_dx <lift_rmap_pap_d_dx
- <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
- #H0 <(IH … l) -IH //
- | /2 width=2 by/
- | <update_L_dx <depth_L_dx <lift_rmap_L_dx <lift_rmap_L_dx
- cases m -m // #m
- <nrplus_succ_sn <nrplus_succ_sn >nsucc_inj #H0
- <tr_pap_push <tr_pap_push
- <(IH … l) -IH //
- | /2 width=2 by/
- | /2 width=2 by/
- ]
-]
-qed.
-
-lemma lift_rmap_pap_gt (f) (p) (m):
- f@❨m+⧣p❩+❘p❘ = (↑[p]f)@❨m+❘p❘❩.
-#f #p @(list_ind_rcons … p) -p [ // ]
-#p * [ #n ] #IH #m
-[ <update_d_dx <depth_d_dx
- <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
- <lift_rmap_pap_d_dx >IH -IH //
-| //
-| <update_L_dx <depth_L_dx
- <nrplus_succ_dx <nrplus_succ_dx <lift_rmap_L_dx <tr_pap_push //
-| //
-| //
-]
-qed.
-
-lemma lift_rmap_tls_gt (f) (p) (m:pnat):
- ⇂*[ninj (m+⧣p)]f ≗ ⇂*[ninj (m+❘p❘)]↑[p]f.
-#f #p @(list_ind_rcons … p) -p [ // ]
-#p * [ #n ] #IH #m
-[ <update_d_dx <depth_d_dx
- <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
- @(stream_eq_trans … (lift_rmap_tls_d_dx …))
- @(stream_eq_canc_dx … (IH …)) -IH //
-| //
-| <update_L_dx <depth_L_dx
- <nrplus_succ_dx >nsucc_inj //
-| //
-| //
-]
-qed.
-
-lemma lift_rmap_tls_eq (f) (p):
- ⇂*[⧣p]f ≗ ⇂*[↑❘p❘]↑[p]⫯f.
-(*
-#f #p @(list_ind_rcons … p) -p //
-#p * [ #n ] #IH //
-<update_d_dx <depth_d_dx <lift_rmap_d_dx
-@(stream_eq_trans … (tr_tls_compose_uni_dx …))
-*)
--- /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 "ground/arith/nat_plus.ma".
+include "delayed_updating/syntax/path.ma".
+include "delayed_updating/notation/functions/hash_1.ma".
+
+(* HEIGHT FOR PATH **********************************************************)
+
+rec definition height (p) on p: nat ≝
+match p with
+[ list_empty ⇒ 𝟎
+| list_lcons l q ⇒
+ match l with
+ [ label_d n ⇒ n + height q
+ | label_m ⇒ height q
+ | label_L ⇒ height q
+ | label_A ⇒ height q
+ | label_S ⇒ height q
+ ]
+].
+
+interpretation
+ "height (path)"
+ 'Hash p = (height p).
+
+(* Basic constructions ******************************************************)
+
+lemma height_empty: 𝟎 = ⧣𝐞.
+// qed.
+
+lemma height_d_sn (q) (n): ninj n+⧣q = ⧣(𝗱n◗q).
+// qed.
+
+lemma height_m_sn (q): ⧣q = ⧣(𝗺◗q).
+// qed.
+
+lemma height_L_sn (q): ⧣q = ⧣(𝗟◗q).
+// qed.
+
+lemma height_A_sn (q): ⧣q = ⧣(𝗔◗q).
+// qed.
+
+lemma height_S_sn (q): ⧣q = ⧣(𝗦◗q).
+// qed.
+
+(* Main constructions *******************************************************)
+
+theorem height_append (p1) (p2):
+ (⧣p2+⧣p1) = ⧣(p1●p2).
+#p1 elim p1 -p1 //
+* [ #n ] #p1 #IH #p2 <list_append_lcons_sn
+[ <height_d_sn <height_d_sn //
+| <height_m_sn <height_m_sn //
+| <height_L_sn <height_L_sn //
+| <height_A_sn <height_A_sn //
+| <height_S_sn <height_S_sn //
+]
+qed.
+
+(* Constructions with list_rcons ********************************************)
+
+lemma height_d_dx (p) (n):
+ (⧣p)+(ninj n) = ⧣(p◖𝗱n).
+// qed.
+
+lemma height_m_dx (p):
+ (⧣p) = ⧣(p◖𝗺).
+// qed.
+
+lemma height_L_dx (p):
+ (⧣p) = ⧣(p◖𝗟).
+// qed.
+
+lemma height_A_dx (p):
+ (⧣p) = ⧣(p◖𝗔).
+// qed.
+
+lemma height_S_dx (p):
+ (⧣p) = ⧣(p◖𝗦).
+// 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 "ground/arith/nat_plus.ma".
-include "delayed_updating/syntax/path.ma".
-include "delayed_updating/notation/functions/hash_1.ma".
-
-(* UPDATE COUNT FOR PATH ****************************************************)
-
-rec definition update (p) on p: nat ≝
-match p with
-[ list_empty ⇒ 𝟎
-| list_lcons l q ⇒
- match l with
- [ label_d n ⇒ n + update q
- | label_m ⇒ update q
- | label_L ⇒ update q
- | label_A ⇒ update q
- | label_S ⇒ update q
- ]
-].
-
-interpretation
- "update count (path)"
- 'Hash p = (update p).
-
-(* Basic constructions ******************************************************)
-
-lemma update_empty: 𝟎 = ⧣𝐞.
-// qed.
-
-lemma update_d_sn (q) (n): ninj n+⧣q = ⧣(𝗱n◗q).
-// qed.
-
-lemma update_m_sn (q): ⧣q = ⧣(𝗺◗q).
-// qed.
-
-lemma update_L_sn (q): ⧣q = ⧣(𝗟◗q).
-// qed.
-
-lemma update_A_sn (q): ⧣q = ⧣(𝗔◗q).
-// qed.
-
-lemma update_S_sn (q): ⧣q = ⧣(𝗦◗q).
-// qed.
-
-(* Main constructions *******************************************************)
-
-theorem update_append (p1) (p2):
- (⧣p2+⧣p1) = ⧣(p1●p2).
-#p1 elim p1 -p1 //
-* [ #n ] #p1 #IH #p2 <list_append_lcons_sn
-[ <update_d_sn <update_d_sn //
-| <update_m_sn <update_m_sn //
-| <update_L_sn <update_L_sn //
-| <update_A_sn <update_A_sn //
-| <update_S_sn <update_S_sn //
-]
-qed.
-
-(* Constructions with list_rcons ********************************************)
-
-lemma update_d_dx (p) (n):
- (⧣p)+(ninj n) = ⧣(p◖𝗱n).
-// qed.
-
-lemma update_m_dx (p):
- (⧣p) = ⧣(p◖𝗺).
-// qed.
-
-lemma update_L_dx (p):
- (⧣p) = ⧣(p◖𝗟).
-// qed.
-
-lemma update_A_dx (p):
- (⧣p) = ⧣(p◖𝗔).
-// qed.
-
-lemma update_S_dx (p):
- (⧣p) = ⧣(p◖𝗦).
-// 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 "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@❨n❩❩, 𝐞❩ = ▼{A}❨k, f, 𝗱n◗𝐞❩.
+// qed.
+
+lemma unwind_d_lcons_sn (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_sn (f) (n):
+ 𝗱(f@❨n❩)◗𝐞 = ▼[f](𝗱n◗𝐞).
+// qed.
+
+lemma unwind_path_d_lcons_sn (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-.
--- /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/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.
--- /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/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.
+*)
--- /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/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_sn <unwind_d_empty_sn <(tr_pap_eq_repl … Hf)
+ /2 width=1 by stream_eq_refl/
+| <unwind_d_lcons_sn <unwind_d_lcons_sn
+ /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_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 // [ #n | #n #l #p | #p ] #IH #f
+[ <unwind_path_d_empty_sn //
+| <unwind_path_d_lcons_sn //
+| <unwind_path_L_sn <unwind_path_L_sn //
+]
+qed.
--- /dev/null
+
+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.
+*)
--- /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/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.
--- /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 "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.
--- /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/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.
--- /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 "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.
--- /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/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_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-.
--- /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/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.
--- /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/unwind1/unwind_eq.ma".
+include "delayed_updating/syntax/path_height.ma".
+include "delayed_updating/syntax/path_depth.ma".
+include "ground/lib/stream_eq_eq.ma".
+
+(* UNWIND FOR PATH *********************************************************)
+(* COMMMENT
+(* Constructions with update ***********************************************)
+
+lemma unwind_rmap_pap_le (f1) (f2) (p) (m:pnat) (l:nat):
+ ninj (m+⧣p+l) = ❘p❘ →
+ (▼[p]f1)@❨m❩ = (▼[p]f2)@❨m❩.
+#f1 #f2 #p @(list_ind_rcons … p) -p
+[ #m #l <depth_empty #H0 destruct
+| #p * [ #n ] #IH #m #l
+ [ <update_d_dx <depth_d_dx <unwind_rmap_pap_d_dx <unwind_rmap_pap_d_dx
+ <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
+ #H0 <(IH … l) -IH //
+ | /2 width=2 by/
+ | <update_L_dx <depth_L_dx <unwind_rmap_L_dx <unwind_rmap_L_dx
+ cases m -m // #m
+ <nrplus_succ_sn <nrplus_succ_sn >nsucc_inj #H0
+ <tr_pap_push <tr_pap_push
+ <(IH … l) -IH //
+ | /2 width=2 by/
+ | /2 width=2 by/
+ ]
+]
+qed.
+
+lemma unwind_rmap_pap_gt (f) (p) (m):
+ f@❨m+⧣p❩+❘p❘ = (▼[p]f)@❨m+❘p❘❩.
+#f #p @(list_ind_rcons … p) -p [ // ]
+#p * [ #n ] #IH #m
+[ <update_d_dx <depth_d_dx
+ <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
+ <unwind_rmap_pap_d_dx >IH -IH //
+| //
+| <update_L_dx <depth_L_dx
+ <nrplus_succ_dx <nrplus_succ_dx <unwind_rmap_L_dx <tr_pap_push //
+| //
+| //
+]
+qed.
+
+lemma unwind_rmap_tls_gt (f) (p) (m:pnat):
+ ⇂*[ninj (m+⧣p)]f ≗ ⇂*[ninj (m+❘p❘)]▼[p]f.
+#f #p @(list_ind_rcons … p) -p [ // ]
+#p * [ #n ] #IH #m
+[ <update_d_dx <depth_d_dx
+ <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
+ @(stream_eq_trans … (unwind_rmap_tls_d_dx …))
+ @(stream_eq_canc_dx … (IH …)) -IH //
+| //
+| <update_L_dx <depth_L_dx
+ <nrplus_succ_dx >nsucc_inj //
+| //
+| //
+]
+qed.
+
+lemma unwind_rmap_tls_eq (f) (p):
+ ⇂*[⧣p]f ≗ ⇂*[▼❘p❘]▼[p]⫯f.
+(*
+#f #p @(list_ind_rcons … p) -p //
+#p * [ #n ] #IH //
+<update_d_dx <depth_d_dx <unwind_rmap_d_dx
+@(stream_eq_trans … (tr_tls_compose_uni_dx …))
+*)
+*)
--- /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 "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-.
--- /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/unwind2/unwind_prototerm_eq.ma".
+include "delayed_updating/syntax/prototerm_constructors.ma".
+
+(* UNWIND FOR PROTOTERM *****************************************************)
+
+lemma unwind_iref_after_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_after_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_after (f) (t) (n:pnat):
+ ▼[f∘𝐮❨n❩]t ⇔ ▼[f](𝛗n.t).
+/3 width=1 by conj, unwind_iref_after_sn, unwind_iref_after_dx/
+qed.
+
+lemma unwind_iref (f) (t) (n:pnat):
+ ▼[f]▼[𝐮❨n❩]t ⇔ ▼[f](𝛗n.t).
+/3 width=3 by unwind_term_after, subset_eq_trans/
+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/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.
--- /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/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.
--- /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/unwind2/unwind_prototerm_eq.ma".
+include "delayed_updating/unwind2/unwind_structure.ma".
+include "delayed_updating/substitution/fsubst.ma".
+include "delayed_updating/syntax/preterm.ma".
+include "delayed_updating/syntax/prototerm_proper.ma".
+
+(* UNWIND FOR PROTOTERM *****************************************************)
+
+(* Constructions with fsubst ************************************************)
+
+lemma unwind_fsubst_sn (f) (t) (u) (p): u ϵ 𝐏 →
+ (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u] ⊆ ▼[f](t[⋔p←u]).
+#f #t #u #p #Hu #ql * *
+[ #rl * #r #Hr #H1 #H2 destruct
+ >unwind_append_proper_dx
+ /4 width=5 by in_comp_unwind_bi, in_ppc_comp_trans, or_introl, ex2_intro/
+| * #q #Hq #H1 #H0
+ @(ex2_intro … H1) @or_intror @conj // *
+ [ <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.
--- /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/unwind2/unwind_eq.ma".
+include "delayed_updating/syntax/path_height.ma".
+include "delayed_updating/syntax/path_depth.ma".
+include "ground/lib/stream_eq_eq.ma".
+
+(* UNWIND FOR PATH **********************************************************)
+
+(* Constructions with height ************************************************)
+
+lemma unwind_rmap_pap_le (f1) (f2) (p) (m:pnat) (l:nat):
+ ninj (m+⧣p+l) = ❘p❘ →
+ (▼[p]f1)@❨m❩ = (▼[p]f2)@❨m❩.
+#f1 #f2 #p @(list_ind_rcons … p) -p
+[ #m #l <depth_empty #H0 destruct
+| #p * [ #n ] #IH #m #l
+ [ <height_d_dx <depth_d_dx <unwind_rmap_pap_d_dx <unwind_rmap_pap_d_dx
+ <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
+ #H0 <(IH … l) -IH //
+ | /2 width=2 by/
+ | <height_L_dx <depth_L_dx <unwind_rmap_L_dx <unwind_rmap_L_dx
+ cases m -m // #m
+ <nrplus_succ_sn <nrplus_succ_sn >nsucc_inj #H0
+ <tr_pap_push <tr_pap_push
+ <(IH … l) -IH //
+ | /2 width=2 by/
+ | /2 width=2 by/
+ ]
+]
+qed.
+
+lemma unwind_rmap_pap_gt (f) (p) (m):
+ f@❨m+⧣p❩+❘p❘ = (▼[p]f)@❨m+❘p❘❩.
+#f #p @(list_ind_rcons … p) -p [ // ]
+#p * [ #n ] #IH #m
+[ <height_d_dx <depth_d_dx
+ <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
+ <unwind_rmap_pap_d_dx >IH -IH //
+| //
+| <height_L_dx <depth_L_dx
+ <nrplus_succ_dx <nrplus_succ_dx <unwind_rmap_L_dx <tr_pap_push //
+| //
+| //
+]
+qed.
+
+lemma unwind_rmap_tls_gt (f) (p) (m:pnat):
+ ⇂*[ninj (m+⧣p)]f ≗ ⇂*[ninj (m+❘p❘)]▼[p]f.
+#f #p @(list_ind_rcons … p) -p [ // ]
+#p * [ #n ] #IH #m
+[ <height_d_dx <depth_d_dx
+ <nplus_comm <nrplus_inj_sn <nrplus_inj_dx <nrplus_pplus_assoc
+ @(stream_eq_trans … (unwind_rmap_tls_d_dx …))
+ @(stream_eq_canc_dx … (IH …)) -IH //
+| //
+| <height_L_dx <depth_L_dx
+ <nrplus_succ_dx >nsucc_inj //
+| //
+| //
+]
+qed.
+
+lemma unwind_rmap_tls_eq (f) (p):
+ ⇂*[⧣p]f ≗ ⇂*[↑❘p❘]▼[p]⫯f.
+(*
+#f #p @(list_ind_rcons … p) -p //
+#p * [ #n ] #IH //
+<height_d_dx <depth_d_dx <unwind_rmap_d_dx
+@(stream_eq_trans … (tr_tls_compose_uni_dx …))
+*)
--- /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 "ground/lib/subset_equivalence.ma".
+include "delayed_updating/syntax/path_structure_inner.ma".
+include "delayed_updating/syntax/preterm.ma".
+include "delayed_updating/unwind2/unwind_structure.ma".
+include "delayed_updating/unwind2/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.
--- /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/unwind2/unwind.ma".
+include "delayed_updating/syntax/prototerm.ma".
+include "ground/lib/subset_ext.ma".
+
+(* UNWIND FOR PROTOTERM *****************************************************)
+
+interpretation
+ "unwind (prototerm)"
+ 'BlackDownTriangle f t = (subset_ext_f1 ? ? (unwind_gen ? proj_path f) t).
+
+(* Basic constructions ******************************************************)
+
+lemma in_comp_unwind_bi (f) (p) (t):
+ p ϵ t → ▼[f]p ϵ ▼[f]t.
+/2 width=1 by subset_in_ext_f1_dx/
+qed.
--- /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 "ground/lib/subset_ext_equivalence.ma".
+include "delayed_updating/unwind2/unwind_eq.ma".
+include "delayed_updating/unwind2/unwind_prototerm.ma".
+
+(* UNWIND FOR PROTOTERM *****************************************************)
+
+(* Constructions with subset_equivalence ************************************)
+
+lemma unwind_term_eq_repl_sn (f1) (f2) (t):
+ f1 ≗ f2 → ▼[f1]t ⇔ ▼[f2]t.
+/3 width=1 by subset_equivalence_ext_f1_exteq, unwind_path_eq_repl/
+qed.
+
+lemma unwind_term_eq_repl_dx (f) (t1) (t2):
+ t1 ⇔ t2 → ▼[f]t1 ⇔ ▼[f]t2.
+/2 width=1 by subset_equivalence_ext_f1_bi/
+qed.
+
+lemma unwind_term_after (f1) (f2) (t):
+ ▼[f2]▼[f1]t ⇔ ▼[f2∘f1]t.
+#f1 #f2 #t @subset_eq_trans
+[
+| @subset_inclusion_ext_f1_compose
+| @subset_equivalence_ext_f1_exteq /2 width=5/
+]
+qed.
--- /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/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-.
--- /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/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.