+++ /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_constructors.ma".
-include "delayed_updating/syntax/prototerm_eq.ma".
-include "delayed_updating/syntax/path_closed.ma".
-include "delayed_updating/syntax/path_balanced.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/notation/relations/black_rightarrow_dbf_3.ma".
-include "ground/arith/nat_rplus.ma".
-include "ground/xoa/ex_6_5.ma".
-
-(* DELAYED BALANCED FOCUSED REDUCTION ***************************************)
-
-definition dbfr (r): relation2 prototerm prototerm ≝
- λt1,t2.
- ∃∃p,b,q,m,n. p●𝗔◗b●𝗟◗q = r &
- ⊗b ϵ 𝐁 & b ϵ 𝐂❨m❩ & q ϵ 𝐂❨n❩ & r◖𝗱↑n ϵ t1 &
- t1[⋔r←𝛕↑(m+n).(t1⋔(p◖𝗦))] ⇔ t2
-.
-
-interpretation
- "balanced focused reduction with delayed updating (prototerm)"
- 'BlackRightArrowDBF t1 r t2 = (dbfr r 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/dbfr.ma".
-include "delayed_updating/reduction/ibfr.ma".
-
-include "delayed_updating/unwind/unwind2_constructors.ma".
-include "delayed_updating/unwind/unwind2_preterm_fsubst.ma".
-include "delayed_updating/unwind/unwind2_preterm_eq.ma".
-include "delayed_updating/unwind/unwind2_prototerm_lift.ma".
-include "delayed_updating/unwind/unwind2_rmap_closed.ma".
-
-include "delayed_updating/substitution/fsubst_eq.ma".
-include "delayed_updating/substitution/lift_prototerm_eq.ma".
-
-include "delayed_updating/syntax/prototerm_proper_constructors.ma".
-include "delayed_updating/syntax/path_closed_structure.ma".
-include "delayed_updating/syntax/path_structure_depth.ma".
-
-(* DELAYED BALANCED FOCUSED REDUCTION ***************************************)
-
-(* Main destructions with ibfr **********************************************)
-
-theorem dbfr_des_ibfr (f) (t1) (t2) (r): t1 ϵ 𝐓 →
- t1 ➡𝐝𝐛𝐟[r] t2 → ▼[f]t1 ➡𝐢𝐛𝐟[⊗r] ▼[f]t2.
-#f #t1 #t2 #r #H0t1
-* #p #b #q #m #n #Hr #Hb #Hm #Hn #Ht1 #Ht2 destruct
-@(ex6_5_intro … (⊗p) (⊗b) (⊗q) (♭b) (♭q))
-[ -H0t1 -Hb -Hm -Hn -Ht1 -Ht2 //
-| -H0t1 -Hm -Hn -Ht1 -Ht2 //
-| -H0t1 -Hb -Hn -Ht1 -Ht2
- /2 width=2 by path_closed_structure_depth/
-| -H0t1 -Hb -Hm -Ht1 -Ht2
- /2 width=2 by path_closed_structure_depth/
-| lapply (in_comp_unwind2_path_term f … Ht1) -H0t1 -Hb -Hm -Ht2 -Ht1
- <unwind2_path_d_dx <tr_pap_succ_nap >list_append_rcons_dx >list_append_assoc
- <unwind2_rmap_append_closed_nap //
-| lapply (unwind2_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
-(*
- lapply (path_head_refl_append_bi … Hh H1k) -Hh -H1k <nrplus_inj_sn #H1k
-*)
- @(subset_eq_trans … Ht2) -t2
- @(subset_eq_trans … (unwind2_term_fsubst_ppc …))
- [ @fsubst_eq_repl [ // | // ]
- @(subset_eq_trans … (unwind2_term_iref …))
- @(subset_eq_canc_sn … (lift_term_eq_repl_dx …))
- [ @unwind2_term_grafted_S /2 width=2 by ex_intro/ | skip ] -Ht1
- @(subset_eq_trans … (lift_unwind2_term_after …))
- @unwind2_term_eq_repl_sn
-(* Note: crux of the proof begins *)
- <list_append_rcons_sn
- @(stream_eq_trans … (tr_compose_uni_dx_pap …)) <tr_pap_succ_nap
- @tr_compose_eq_repl
-(* TODO
- [ <unwind2_rmap_append_closed_nap //
- <depth_append <depth_L_sn
- <nplus_comm <nrplus_npsucc_sn <nplus_succ_sn //
- | >unwind2_rmap_A_dx
- /2 width=1 by tls_unwind2_rmap_closed/
- ]
-*)
-(* Note: crux of the proof ends *)
- | //
- | /2 width=2 by ex_intro/
- | //
- ]
-]
-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/reduction/dbfr.ma".
-
-include "delayed_updating/substitution/fsubst_lift.ma".
-include "delayed_updating/substitution/fsubst_eq.ma".
-include "delayed_updating/substitution/lift_constructors.ma".
-include "delayed_updating/substitution/lift_path_head.ma".
-include "delayed_updating/substitution/lift_rmap_head.ma".
-
-(* DELAYED BALANCED FOCUSED REDUCTION ***************************************)
-
-(* Constructions with lift **************************************************)
-
-theorem dfr_lift_bi (f) (t1) (t2) (r):
- t1 ➡𝐝𝐛𝐟[r] t2 → ↑[f]t1 ➡𝐝𝐛𝐟[↑[f]r] ↑[f]t2.
-#f #t1 #t2 #r
-* #p #q #k #Hr #H1k #Ht1 #Ht2 destruct
-@(ex4_3_intro … (↑[f]p) (↑[↑[p◖𝗔◖𝗟]f]q) ((↑[p●𝗔◗𝗟◗q]f)@⧣❨k❩))
-[ -H1k -Ht1 -Ht2 //
-| -Ht1 -Ht2
- <lift_rmap_L_dx >lift_path_L_sn
- <(lift_path_head_closed … H1k) in ⊢ (??%?); -H1k //
-| lapply (in_comp_lift_path_term f … Ht1) -Ht2 -Ht1 -H1k
- <lift_path_d_dx #Ht1 //
-| lapply (lift_term_eq_repl_dx f … Ht2) -Ht2 #Ht2 -Ht1
- @(subset_eq_trans … Ht2) -t2
- @(subset_eq_trans … (lift_term_fsubst …))
- @fsubst_eq_repl [ // | // ]
- @(subset_eq_trans … (lift_term_iref …))
- @iref_eq_repl
- @(subset_eq_canc_sn … (lift_term_grafted_S …))
- @lift_term_eq_repl_sn
-(* Note: crux of the proof begins *)
- >list_append_rcons_sn in H1k; #H1k >lift_rmap_A_dx
- /2 width=1 by tls_lift_rmap_closed/
-(* Note: crux of the proof ends *)
-]
-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.ma".
-include "delayed_updating/syntax/prototerm_eq.ma".
-include "delayed_updating/syntax/path_closed.ma".
-include "delayed_updating/syntax/path_balanced.ma".
-include "delayed_updating/syntax/path_structure.ma".
-include "delayed_updating/notation/relations/black_rightarrow_ibf_3.ma".
-include "ground/relocation/tr_uni.ma".
-include "ground/arith/nat_rplus.ma".
-include "ground/xoa/ex_6_5.ma".
-
-(* IMMEDIATE BALANCED FOCUSED REDUCTION *************************************)
-
-definition ibfr (r): relation2 prototerm prototerm ≝
- λt1,t2.
- ∃∃p,b,q,m,n. p●𝗔◗b●𝗟◗q = r &
- ⊗b ϵ 𝐁 & b ϵ 𝐂❨m❩ & q ϵ 𝐂❨n❩ & r◖𝗱↑n ϵ t1 &
- t1[⋔r←↑[𝐮❨↑(m+n)❩](t1⋔(p◖𝗦))] ⇔ t2
-.
-
-interpretation
- "balanced focused reduction with immediate updating (prototerm)"
- 'BlackRightArrowIBF t1 r t2 = (ibfr r 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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR DELAYED UPDATING ********************************************)
+
+notation "hvbox( 🠢[ break term 46 t1 ] break term 70 t2 )"
+ non associative with precedence 70
+ for @{ 'RightTriangleArrow $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 *)
-(* *)
-(**************************************************************************)
-
-(* NOTATION FOR DELAYED UPDATING ********************************************)
-
-notation "hvbox( ↑[ break term 46 t1 ] break term 70 t2 )"
- non associative with precedence 70
- for @{ 'UpArrow $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 *)
+(* *)
+(**************************************************************************)
+
+(* NOTATION FOR DELAYED UPDATING ********************************************)
+
+notation "hvbox( 🠡[ break term 46 t1 ] break term 70 t2 )"
+ non associative with precedence 70
+ for @{ 'UpTriangleArrow $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/fsubst.ma".
+include "delayed_updating/syntax/prototerm_constructors.ma".
+include "delayed_updating/syntax/prototerm_eq.ma".
+include "delayed_updating/syntax/path_closed.ma".
+include "delayed_updating/syntax/path_balanced.ma".
+include "delayed_updating/syntax/path_structure.ma".
+include "delayed_updating/notation/relations/black_rightarrow_dbf_3.ma".
+include "ground/arith/nat_rplus.ma".
+include "ground/xoa/ex_6_5.ma".
+
+(* DELAYED BALANCED FOCUSED REDUCTION ***************************************)
+
+definition dbfr (r): relation2 prototerm prototerm ≝
+ λt1,t2.
+ ∃∃p,b,q,m,n. p●𝗔◗b●𝗟◗q = r &
+ ⊗b ϵ 𝐁 & b ϵ 𝐂❨Ⓣ,m❩ & q ϵ 𝐂❨Ⓕ,n❩ & r◖𝗱↑n ϵ t1 &
+ t1[⋔r←𝛕↑(m+n).(t1⋔(p◖𝗦))] ⇔ t2
+.
+
+interpretation
+ "balanced focused reduction with delayed updating (prototerm)"
+ 'BlackRightArrowDBF t1 r t2 = (dbfr r 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/dbfr.ma".
+include "delayed_updating/reduction/ibfr.ma".
+
+include "delayed_updating/unwind/unwind2_constructors.ma".
+include "delayed_updating/unwind/unwind2_preterm_fsubst.ma".
+include "delayed_updating/unwind/unwind2_preterm_eq.ma".
+include "delayed_updating/unwind/unwind2_prototerm_lift.ma".
+include "delayed_updating/unwind/unwind2_rmap_closed.ma".
+
+include "delayed_updating/substitution/fsubst_eq.ma".
+include "delayed_updating/substitution/lift_prototerm_eq.ma".
+
+include "delayed_updating/syntax/prototerm_proper_constructors.ma".
+include "delayed_updating/syntax/path_closed_structure.ma".
+include "delayed_updating/syntax/path_structure_depth.ma".
+
+(* DELAYED BALANCED FOCUSED REDUCTION ***************************************)
+
+(* Main destructions with ibfr **********************************************)
+
+theorem dbfr_des_ibfr (f) (t1) (t2) (r): t1 ϵ 𝐓 →
+ t1 ➡𝐝𝐛𝐟[r] t2 → ▼[f]t1 ➡𝐢𝐛𝐟[⊗r] ▼[f]t2.
+#f #t1 #t2 #r #H0t1
+* #p #b #q #m #n #Hr #Hb #Hm #Hn #Ht1 #Ht2 destruct
+@(ex6_5_intro … (⊗p) (⊗b) (⊗q) (♭b) (♭q))
+[ -H0t1 -Hb -Hm -Hn -Ht1 -Ht2 //
+| -H0t1 -Hm -Hn -Ht1 -Ht2 //
+| -H0t1 -Hb -Hn -Ht1 -Ht2
+ /2 width=2 by path_closed_structure_depth/
+| -H0t1 -Hb -Hm -Ht1 -Ht2
+ /2 width=2 by path_closed_structure_depth/
+| lapply (in_comp_unwind2_path_term f … Ht1) -H0t1 -Hb -Hm -Ht2 -Ht1
+ <unwind2_path_d_dx <tr_pap_succ_nap >list_append_rcons_dx >list_append_assoc
+ <unwind2_rmap_append_closed_Lq_dx_nap_depth //
+| lapply (unwind2_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
+ @(subset_eq_trans … Ht2) -t2
+ @(subset_eq_trans … (unwind2_term_fsubst_ppc …))
+ [ @fsubst_eq_repl [ // | // ]
+ @(subset_eq_trans … (unwind2_term_iref …))
+ @(subset_eq_canc_sn … (lift_term_eq_repl_dx …))
+ [ @unwind2_term_grafted_S /2 width=2 by ex_intro/ | skip ] -Ht1
+ @(subset_eq_trans … (lift_unwind2_term_after …))
+ @unwind2_term_eq_repl_sn
+(* Note: crux of the proof begins *)
+ <list_append_rcons_sn
+ @(stream_eq_trans … (tr_compose_uni_dx_pap …)) <tr_pap_succ_nap
+ @tr_compose_eq_repl
+ [ <unwind2_rmap_append_closed_bLq_dx_nap_plus //
+ | >unwind2_rmap_A_dx
+ /2 width=2 by tls_succ_plus_unwind2_rmap_append_closed_bLq_dx/
+ ]
+(* Note: crux of the proof ends *)
+ | //
+ | /2 width=2 by ex_intro/
+ | //
+ ]
+]
+qed.
(* *)
(**************************************************************************)
+include "delayed_updating/reduction/dfr.ma".
+
include "delayed_updating/substitution/fsubst_lift.ma".
include "delayed_updating/substitution/fsubst_eq.ma".
include "delayed_updating/substitution/lift_constructors.ma".
include "delayed_updating/substitution/lift_path_closed.ma".
include "delayed_updating/substitution/lift_rmap_closed.ma".
-(**) (* reverse include *)
-include "delayed_updating/reduction/dfr.ma".
-
(* DELAYED FOCUSED REDUCTION ************************************************)
(* Constructions with lift **************************************************)
theorem dfr_lift_bi (f) (t1) (t2) (r):
- t1 ➡𝐝𝐟[r] t2 → ↑[f]t1 ➡𝐝𝐟[↑[f]r] ↑[f]t2.
+ t1 ➡𝐝𝐟[r] t2 → 🠡[f]t1 ➡𝐝𝐟[🠡[f]r] 🠡[f]t2.
#f #t1 #t2 #r
* #p #q #n #Hr #Hn #Ht1 #Ht2 destruct
-@(ex4_3_intro … (↑[f]p) (↑[↑[p◖𝗔◖𝗟]f]q) ((↑[p●𝗔◗𝗟◗q]f)@§❨n❩))
+@(ex4_3_intro … (🠡[f]p) (🠡[🠢[f](p◖𝗔◖𝗟)]q) (🠢[f](p●𝗔◗𝗟◗q)@§❨n❩))
[ -Hn -Ht1 -Ht2 //
| -Ht1 -Ht2
/2 width=1 by lift_path_rmap_closed_L/
--- /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.ma".
+include "delayed_updating/syntax/prototerm_eq.ma".
+include "delayed_updating/syntax/path_closed.ma".
+include "delayed_updating/syntax/path_balanced.ma".
+include "delayed_updating/syntax/path_structure.ma".
+include "delayed_updating/notation/relations/black_rightarrow_ibf_3.ma".
+include "ground/relocation/tr_uni.ma".
+include "ground/arith/nat_rplus.ma".
+include "ground/xoa/ex_6_5.ma".
+
+(* IMMEDIATE BALANCED FOCUSED REDUCTION *************************************)
+
+definition ibfr (r): relation2 prototerm prototerm ≝
+ λt1,t2.
+ ∃∃p,b,q,m,n. p●𝗔◗b●𝗟◗q = r &
+ ⊗b ϵ 𝐁 & b ϵ 𝐂❨Ⓣ,m❩ & q ϵ 𝐂❨Ⓕ,n❩ & r◖𝗱↑n ϵ t1 &
+ t1[⋔r←🠡[𝐮❨↑(m+n)❩](t1⋔(p◖𝗦))] ⇔ t2
+.
+
+interpretation
+ "balanced focused reduction with immediate updating (prototerm)"
+ 'BlackRightArrowIBF t1 r t2 = (ibfr r t1 t2).
λt1,t2.
∃∃p,q,n. p●𝗔◗𝗟◗q = r &
q ϵ 𝐂❨Ⓕ,n❩ & r◖𝗱↑n ϵ t1 &
- t1[⋔r←↑[𝐮❨↑n❩](t1⋔(p◖𝗦))] ⇔ t2
+ t1[⋔r←🠡[𝐮❨↑n❩](t1⋔(p◖𝗦))] ⇔ t2
.
interpretation
(* Constructions with lift **************************************************)
theorem ifr_lift_bi (f) (t1) (t2) (r):
- t1 ➡𝐢𝐟[r] t2 → ↑[f]t1 ➡𝐢𝐟[↑[f]r] ↑[f]t2.
+ t1 ➡𝐢𝐟[r] t2 → 🠡[f]t1 ➡𝐢𝐟[🠡[f]r] 🠡[f]t2.
#f #t1 #t2 #r
* #p #q #n #Hr #Hn #Ht1 #Ht2 destruct
-@(ex4_3_intro … (↑[f]p) (↑[↑[p◖𝗔◖𝗟]f]q) ((↑[p●𝗔◗𝗟◗q]f)@§❨n❩))
+@(ex4_3_intro … (🠡[f]p) (🠡[🠢[f](p◖𝗔◖𝗟)]q) (🠢[f](p●𝗔◗𝗟◗q)@§❨n❩))
[ -Hn -Ht1 -Ht2 //
| -Ht1 -Ht2
/2 width=1 by lift_path_rmap_closed_L/
(* Constructions with lift for preterm **************************************)
lemma lift_term_fsubst_sn (f) (t) (u) (p):
- (↑[f]t)[⋔(↑[f]p)←↑[↑[p]f]u] ⊆ ↑[f](t[⋔p←u]).
+ (🠡[f]t)[⋔(🠡[f]p)←🠡[🠢[f]p]u] ⊆ 🠡[f](t[⋔p←u]).
#f #t #u #p #ql * *
[ #rl * #r #Hr #H1 #H2 destruct
>lift_path_append
qed-.
lemma lift_term_fsubst_dx (f) (t) (u) (p):
- ↑[f](t[⋔p←u]) ⊆ (↑[f]t)[⋔(↑[f]p)←↑[↑[p]f]u].
+ 🠡[f](t[⋔p←u]) ⊆ (🠡[f]t)[⋔(🠡[f]p)←🠡[🠢[f]p]u].
#f #t #u #p #ql * #q * *
[ #r #Hu #H1 #H2 destruct
@or_introl @ex2_intro
qed-.
lemma lift_term_fsubst (f) (t) (u) (p):
- (↑[f]t)[⋔(↑[f]p)←↑[↑[p]f]u] ⇔ ↑[f](t[⋔p←u]).
+ (🠡[f]t)[⋔(🠡[f]p)←🠡[🠢[f]p]u] ⇔ 🠡[f](t[⋔p←u]).
/3 width=1 by lift_term_fsubst_sn, conj, lift_term_fsubst_dx/ qed.
(* LIFT FOR PROTOTERM *******************************************************)
lemma lift_term_iref_pap_sn (f) (t:prototerm) (k:pnat):
- (𝛕f@⧣❨k❩.↑[⇂*[k]f]t) ⊆ ↑[f](𝛕k.t).
+ (𝛕f@⧣❨k❩.🠡[⇂*[k]f]t) ⊆ 🠡[f](𝛕k.t).
#f #t #k #p * #q * #r #Hr #H1 #H2 destruct
@(ex2_intro … (𝗱k◗𝗺◗r))
/2 width=1 by in_comp_iref/
qed-.
lemma lift_term_iref_pap_dx (f) (t) (k:pnat):
- ↑[f](𝛕k.t) ⊆ 𝛕f@⧣❨k❩.↑[⇂*[k]f]t.
+ 🠡[f](𝛕k.t) ⊆ 𝛕f@⧣❨k❩.🠡[⇂*[k]f]t.
#f #t #k #p * #q #Hq #H0 destruct
elim (in_comp_inv_iref … Hq) -Hq #p #H0 #Hp destruct
<lift_path_d_sn <lift_path_m_sn
qed-.
lemma lift_term_iref_pap (f) (t) (k:pnat):
- (𝛕f@⧣❨k❩.↑[⇂*[k]f]t) ⇔ ↑[f](𝛕k.t).
+ (𝛕f@⧣❨k❩.🠡[⇂*[k]f]t) ⇔ 🠡[f](𝛕k.t).
/3 width=1 by conj, lift_term_iref_pap_sn, lift_term_iref_pap_dx/
qed.
lemma lift_term_iref_nap (f) (t) (n):
- (𝛕↑(f@§❨n❩).↑[⇂*[↑n]f]t) ⇔ ↑[f](𝛕↑n.t).
+ (𝛕↑(f@§❨n❩).🠡[⇂*[↑n]f]t) ⇔ 🠡[f](𝛕↑n.t).
#f #t #n
>tr_pap_succ_nap //
qed.
lemma lift_term_iref_uni (t) (n) (k):
- (𝛕(k+n).t) ⇔ ↑[𝐮❨n❩](𝛕k.t).
+ (𝛕(k+n).t) ⇔ 🠡[𝐮❨n❩](𝛕k.t).
#t #n #k
@(subset_eq_trans … (lift_term_iref_pap …))
<tr_uni_pap >nsucc_pnpred <tr_tls_succ_uni
rec definition lift_path (f) (p) on p: path ≝
match p with
[ list_empty ⇒ (𝐞)
-| list_lcons l q ⇒ (lift_path f q)◖↑[↑[q]f]l
+| list_lcons l q ⇒ (lift_path f q)◖🠡[🠢[f]q]l
].
interpretation
"lift (path)"
- 'UpArrow f l = (lift_path f l).
+ 'UpTriangleArrow f l = (lift_path f l).
(* Basic constructions ******************************************************)
lemma lift_path_empty (f):
- (𝐞) = ↑[f]𝐞.
+ (𝐞) = 🠡[f]𝐞.
// qed.
lemma lift_path_rcons (f) (p) (l):
- (↑[f]p)◖↑[↑[p]f]l = ↑[f](p◖l).
+ (🠡[f]p)◖🠡[🠢[f]p]l = 🠡[f](p◖l).
// qed.
lemma lift_path_d_dx (f) (p) (k):
- (↑[f]p)◖𝗱((↑[p]f)@⧣❨k❩) = ↑[f](p◖𝗱k).
+ (🠡[f]p)◖𝗱(🠢[f]p@⧣❨k❩) = 🠡[f](p◖𝗱k).
// qed.
lemma lift_path_m_dx (f) (p):
- (↑[f]p)◖𝗺 = ↑[f](p◖𝗺).
+ (🠡[f]p)◖𝗺 = 🠡[f](p◖𝗺).
// qed.
lemma lift_path_L_dx (f) (p):
- (↑[f]p)◖𝗟 = ↑[f](p◖𝗟).
+ (🠡[f]p)◖𝗟 = 🠡[f](p◖𝗟).
// qed.
lemma lift_path_A_dx (f) (p):
- (↑[f]p)◖𝗔 = ↑[f](p◖𝗔).
+ (🠡[f]p)◖𝗔 = 🠡[f](p◖𝗔).
// qed.
lemma lift_path_S_dx (f) (p):
- (↑[f]p)◖𝗦 = ↑[f](p◖𝗦).
+ (🠡[f]p)◖𝗦 = 🠡[f](p◖𝗦).
// qed.
(* Constructions with path_append *******************************************)
lemma lift_path_append (f) (p) (q):
- (↑[f]p)●(↑[↑[p]f]q) = ↑[f](p●q).
+ (🠡[f]p)●(🠡[🠢[f]p]q) = 🠡[f](p●q).
#f #p #q elim q -q //
#l #q #IH
<lift_path_rcons <lift_path_rcons
(* Constructions with path_lcons ********************************************)
lemma lift_path_lcons (f) (p) (l):
- (↑[f]l)◗↑[↑[l]f]p = ↑[f](l◗p).
+ (🠡[f]l)◗🠡[🠢[f]l]p = 🠡[f](l◗p).
#f #p #l
<lift_path_append //
qed.
lemma lift_path_d_sn (f) (p) (k:pnat):
- (𝗱(f@⧣❨k❩)◗↑[⇂*[k]f]p) = ↑[f](𝗱k◗p).
+ (𝗱(f@⧣❨k❩)◗🠡[⇂*[k]f]p) = 🠡[f](𝗱k◗p).
// qed.
lemma lift_path_m_sn (f) (p):
- (𝗺◗↑[f]p) = ↑[f](𝗺◗p).
+ (𝗺◗🠡[f]p) = 🠡[f](𝗺◗p).
// qed.
lemma lift_path_L_sn (f) (p):
- (𝗟◗↑[⫯f]p) = ↑[f](𝗟◗p).
+ (𝗟◗🠡[⫯f]p) = 🠡[f](𝗟◗p).
// qed.
lemma lift_path_A_sn (f) (p):
- (𝗔◗↑[f]p) = ↑[f](𝗔◗p).
+ (𝗔◗🠡[f]p) = 🠡[f](𝗔◗p).
// qed.
lemma lift_path_S_sn (f) (p):
- (𝗦◗↑[f]p) = ↑[f](𝗦◗p).
+ (𝗦◗🠡[f]p) = 🠡[f](𝗦◗p).
// qed.
(* Basic inversions *********************************************************)
lemma lift_path_inv_empty (f) (p):
- (𝐞) = ↑[f]p → 𝐞 = p.
+ (𝐞) = 🠡[f]p → 𝐞 = p.
#f * // #p #l
<lift_path_rcons #H0 destruct
qed-.
lemma lift_path_inv_rcons (f) (p2) (q1) (l1):
- q1◖l1 = ↑[f]p2 →
- ∃∃q2,l2. q1 = ↑[f]q2 & l1 = ↑[↑[q2]f]l2 & q2◖l2 = p2.
+ q1◖l1 = 🠡[f]p2 →
+ ∃∃q2,l2. q1 = 🠡[f]q2 & l1 = 🠡[🠢[f]q2]l2 & q2◖l2 = p2.
#f * [| #l2 #q2 ] #q1 #l1
[ <lift_path_empty
| <lift_path_rcons
(* Inversions with path_append **********************************************)
lemma lift_path_inv_append_sn (f) (q1) (r1) (p2):
- q1●r1 = ↑[f]p2 →
- ∃∃q2,r2. q1 = ↑[f]q2 & r1 = ↑[↑[q2]f]r2 & q2●r2 = p2.
+ q1●r1 = 🠡[f]p2 →
+ ∃∃q2,r2. q1 = 🠡[f]q2 & r1 = 🠡[🠢[f]q2]r2 & q2●r2 = p2.
#f #q1 #r1 elim r1 -r1 [| #l1 #r1 #IH ] #p2
[ <list_append_empty_sn #H0 destruct
/2 width=5 by ex3_2_intro/
(* Main inversions **********************************************************)
theorem lift_path_inj (f) (p1) (p2):
- ↑[f]p1 = ↑[f]p2 → p1 = p2.
+ 🠡[f]p1 = 🠡[f]p2 → p1 = p2.
#f #p1 elim p1 -p1 [| #l1 #q1 #IH ] #p2
[ <lift_path_empty #H0
<(lift_path_inv_empty … H0) -H0 //
(* Constructions with tr_after **********************************************)
lemma lift_path_after (g) (f) (p):
- ↑[g]↑[f]p = ↑[g∘f]p.
+ 🠡[g]🠡[f]p = 🠡[g∘f]p.
#g #f #p elim p -p //
#l #p #IH
<lift_path_rcons <lift_path_rcons <IH -IH //
(* Constructions with pcc ***************************************************)
lemma lift_path_closed (o) (f) (q) (n):
- q ϵ 𝐂❨o,n❩ → ↑[f]q ϵ 𝐂❨o,↑[q]f@❨n❩❩.
+ q ϵ 𝐂❨o,n❩ → 🠡[f]q ϵ 𝐂❨o,🠢[f]q@❨n❩❩.
#o #f #q #n #H0 elim H0 -q -n //
#q #n [ #k #Ho ] #_ #IH
/2 width=1 by pcc_m_dx, pcc_L_dx, pcc_A_dx, pcc_S_dx/
qed.
lemma lift_path_rmap_closed (o) (f) (p) (q) (n):
- q ϵ 𝐂❨o,n❩ → ↑[↑[p]f]q ϵ 𝐂❨o,↑[p●q]f@❨n❩❩.
+ q ϵ 𝐂❨o,n❩ → 🠡[🠢[f]p]q ϵ 𝐂❨o,🠢[f](p●q)@❨n❩❩.
/2 width=1 by lift_path_closed/
qed.
lemma lift_path_rmap_closed_L (o) (f) (p) (q) (n):
- q ϵ 𝐂❨o,n❩ → ↑[↑[p◖𝗟]f]q ϵ 𝐂❨o,↑[p●𝗟◗q]f@§❨n❩❩.
+ q ϵ 𝐂❨o,n❩ → 🠡[🠢[f](p◖𝗟)]q ϵ 𝐂❨o,🠢[f](p●𝗟◗q)@§❨n❩❩.
#o #f #p #q #n #Hq
-lapply (lift_path_closed … (↑[p◖𝗟]f) … Hq) #Hq0
+lapply (lift_path_closed … (🠢[f](p◖𝗟)) … Hq) #Hq0
lapply (pcc_L_sn … Hq) -Hq #Hq1
lapply (lift_path_rmap_closed … f p … Hq1) -Hq1
<lift_path_L_sn >lift_rmap_L_dx #Hq1
(* Constructions with path_eq ***********************************************)
lemma lift_path_eq_repl (p):
- stream_eq_repl … (λf1,f2. ↑[f1]p = ↑[f2]p).
+ stream_eq_repl … (λf1,f2. 🠡[f1]p = 🠡[f2]p).
#p elim p -p //
#l #p #IH #f1 #f2 #Hf
<lift_path_rcons <lift_path_rcons
(* Constructions with tr_id *************************************************)
lemma lift_path_id (p):
- p = ↑[𝐢]p.
+ p = 🠡[𝐢]p.
#p elim p -p //
#l #p #IH
<lift_path_rcons //
(* Constructions with list_length *******************************************)
lemma lift_path_length (f) (p):
- ❘p❘ = ❘↑[f]p❘.
+ ❘p❘ = ❘🠡[f]p❘.
#f #p elim p -p // * [ #k ] #p #IH
[ <lift_path_d_dx
| <lift_path_m_dx
(* Constructions with proper condition for path *****************************)
lemma lift_path_proper (f) (p):
- p ϵ 𝐏 → ↑[f]p ϵ 𝐏.
+ p ϵ 𝐏 → 🠡[f]p ϵ 𝐏.
#f *
[ #H0 elim (ppc_inv_empty … H0)
| * [ #k ] #p #_
(* Inversions with proper condition for path ********************************)
lemma lift_path_inv_proper (f) (p):
- ↑[f]p ϵ 𝐏 → p ϵ 𝐏.
+ 🠡[f]p ϵ 𝐏 → p ϵ 𝐏.
#f * //
#H0 elim (ppc_inv_empty … H0)
qed-.
(* Constructions with structure *********************************************)
lemma structure_lift_path (f) (p):
- ⊗p = ⊗↑[f]p.
+ ⊗p = ⊗🠡[f]p.
#f #p elim p -p //
* [ #k ] #p #IH //
[ <lift_path_d_dx <structure_d_dx <structure_d_dx //
qed.
lemma lift_path_structure (f) (p):
- ⊗p = ↑[f]⊗p.
+ ⊗p = 🠡[f]⊗p.
#f #p elim p -p //
* [ #k ] #p #IH //
qed.
(* Constructions with tr_uni ************************************************)
lemma lift_path_d_sn_uni (p) (n) (k):
- (𝗱(k+n)◗p) = ↑[𝐮❨n❩](𝗱k◗p).
+ (𝗱(k+n)◗p) = 🠡[𝐮❨n❩](𝗱k◗p).
#p #n #k
<lift_path_d_sn <tr_uni_pap >nsucc_pnpred
<tr_tls_succ_uni //
interpretation
"lift (prototerm)"
- 'UpArrow f t = (subset_ext_f1 ? ? (lift_path f) t).
+ 'UpTriangleArrow f t = (subset_ext_f1 ? ? (lift_path f) t).
(* Basic constructions ******************************************************)
lemma in_comp_lift_path_term (f) (t) (p):
- p ϵ t → ↑[f]p ϵ ↑[f]t.
+ p ϵ t → 🠡[f]p ϵ 🠡[f]t.
/2 width=1 by subset_in_ext_f1_dx/
qed.
(* Constructions with tr_after **********************************************)
lemma lift_term_after (f) (g) (t):
- ↑[g]↑[f]t ⇔ ↑[g∘f]t.
+ 🠡[g]🠡[f]t ⇔ 🠡[g∘f]t.
#f #g #t @subset_eq_trans
[
| @subset_inclusion_ext_f1_compose
(* Constructions with subset_equivalence ************************************)
lemma lift_term_eq_repl_sn (f1) (f2) (t):
- f1 ≗ f2 → ↑[f1]t ⇔ ↑[f2]t.
+ f1 ≗ f2 → 🠡[f1]t ⇔ 🠡[f2]t.
/3 width=1 by subset_equivalence_ext_f1_exteq, lift_path_eq_repl/
qed.
lemma lift_term_eq_repl_dx (f) (t1) (t2):
- t1 ⇔ t2 → ↑[f]t1 ⇔ ↑[f]t2.
+ t1 ⇔ t2 → 🠡[f]t1 ⇔ 🠡[f]t2.
/2 width=1 by subset_equivalence_ext_f1_bi/
qed.
lemma lift_term_grafted_sn (f) (t) (p):
- ↑[↑[p]f](t⋔p) ⊆ (↑[f]t)⋔(↑[f]p).
+ 🠡[🠢[f]p](t⋔p) ⊆ (🠡[f]t)⋔(🠡[f]p).
#f #t #p #q * #r #Hr #H0 destruct
/2 width=3 by ex2_intro/
qed-.
lemma lift_term_grafted_dx (f) (t) (p):
- (↑[f]t)⋔(↑[f]p) ⊆ ↑[↑[p]f](t⋔p).
+ (🠡[f]t)⋔(🠡[f]p) ⊆ 🠡[🠢[f]p](t⋔p).
#f #t #p #q * #r #Hr #H0
elim (lift_path_inv_append_sn … (sym_eq … H0)) -H0
#p0 #q0 #Hp0 #Hq0 #H0 destruct
qed-.
lemma lift_term_grafted (f) (t) (p):
- ↑[↑[p]f](t⋔p) ⇔ (↑[f]t)⋔(↑[f]p).
+ 🠡[🠢[f]p](t⋔p) ⇔ (🠡[f]t)⋔(🠡[f]p).
/3 width=1 by lift_term_grafted_sn, lift_term_grafted_dx, conj/ qed.
lemma lift_term_grafted_S (f) (t) (p):
- ↑[↑[p]f](t⋔(p◖𝗦)) ⇔ (↑[f]t)⋔((↑[f]p)◖𝗦).
+ 🠡[🠢[f]p](t⋔(p◖𝗦)) ⇔ (🠡[f]t)⋔((🠡[f]p)◖𝗦).
// qed.
(* Constructions with tr_id *************************************************)
lemma lift_term_id_sn (t):
- t ⊆ ↑[𝐢]t.
+ t ⊆ 🠡[𝐢]t.
#t #p #Hp
>(lift_path_id p)
/2 width=1 by in_comp_lift_path_term/
qed-.
lemma lift_term_id_dx (t):
- ↑[𝐢]t ⊆ t.
+ 🠡[𝐢]t ⊆ t.
#t #p * #q #Hq #H destruct //
qed-.
lemma lift_term_id (t):
- t ⇔ ↑[𝐢]t.
+ t ⇔ 🠡[𝐢]t.
/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/
qed.
(* Constructions with proper condition for path *****************************)
lemma lift_term_proper (f) (t):
- t ϵ 𝐏 → ↑[f]t ϵ 𝐏.
+ t ϵ 𝐏 → 🠡[f]t ϵ 𝐏.
#f #t #Ht #p * #q #Hq #H0 destruct
@lift_path_proper @Ht -Ht // (**) (* auto fails *)
qed.
(* Inversions with proper condition for path ********************************)
lemma lift_term_inv_proper (f) (t):
- ↑[f]t ϵ 𝐏 → t ϵ 𝐏.
+ 🠡[f]t ϵ 𝐏 → t ϵ 𝐏.
#f #t #Ht #p #Hp
@(lift_path_inv_proper f)
@Ht -Ht @in_comp_lift_path_term // (**) (* auto fails *)
rec definition lift_rmap (f) (p) on p: tr_map ≝
match p with
[ list_empty ⇒ f
-| list_lcons l q ⇒ ↑[l](lift_rmap f q)
+| list_lcons l q ⇒ 🠢[lift_rmap f q]l
].
interpretation
"lift (relocation map)"
- 'UpArrow p f = (lift_rmap f p).
+ 'RightTriangleArrow f p = (lift_rmap f p).
(* Basic constructions ******************************************************)
lemma lift_rmap_empty (f):
- f = ↑[𝐞]f.
+ f = 🠢[f]𝐞.
// qed.
lemma lift_rmap_rcons (f) (p) (l):
- ↑[l]↑[p]f = ↑[p◖l]f.
+ 🠢[🠢[f]p]l = 🠢[f](p◖l).
// qed.
lemma lift_rmap_d_dx (f) (p) (k:pnat):
- ⇂*[k](↑[p]f) = ↑[p◖𝗱k]f.
+ ⇂*[k](🠢[f]p) = 🠢[f](p◖𝗱k).
// qed.
lemma lift_rmap_m_dx (f) (p):
- ↑[p]f = ↑[p◖𝗺]f.
+ 🠢[f]p = 🠢[f](p◖𝗺).
// qed.
lemma lift_rmap_L_dx (f) (p):
- (⫯↑[p]f) = ↑[p◖𝗟]f.
+ (⫯🠢[f]p) = 🠢[f](p◖𝗟).
// qed.
lemma lift_rmap_A_dx (f) (p):
- ↑[p]f = ↑[p◖𝗔]f.
+ 🠢[f]p = 🠢[f](p◖𝗔).
// qed.
lemma lift_rmap_S_dx (f) (p):
- ↑[p]f = ↑[p◖𝗦]f.
+ 🠢[f]p = 🠢[f](p◖𝗦).
// qed.
(* Constructions with path_append *******************************************)
lemma lift_rmap_append (p) (q) (f):
- ↑[q]↑[p]f = ↑[p●q]f.
+ 🠢[🠢[f]p]q = 🠢[f](p●q).
#p #q elim q -q //
qed.
(* Constructions with path_lcons ********************************************)
lemma lift_rmap_lcons (f) (p) (l):
- ↑[p]↑[l]f = ↑[l◗p]f.
+ 🠢[🠢[f]l]p = 🠢[f](l◗p).
// qed.
lemma lift_rmap_d_sn (f) (p) (k:pnat):
- ↑[p](⇂*[k]f) = ↑[𝗱k◗p]f.
+ 🠢[⇂*[k]f]p = 🠢[f](𝗱k◗p).
// qed.
lemma lift_rmap_m_sn (f) (p):
- ↑[p]f = ↑[𝗺◗p]f.
+ 🠢[f]p = 🠢[f](𝗺◗p).
// qed.
lemma lift_rmap_L_sn (f) (p):
- ↑[p](⫯f) = ↑[𝗟◗p]f.
+ 🠢[⫯f]p = 🠢[f](𝗟◗p).
// qed.
lemma lift_rmap_A_sn (f) (p):
- ↑[p]f = ↑[𝗔◗p]f.
+ 🠢[f]p = 🠢[f](𝗔◗p).
// qed.
lemma lift_rmap_S_sn (f) (p):
- ↑[p]f = ↑[𝗦◗p]f.
+ 🠢[f]p = 🠢[f](𝗦◗p).
// qed.
(* Constructions with tr_after **********************************************)
lemma lift_rmap_after (g) (f) (p:path):
- ↑[↑[f]p]g∘↑[p]f = ↑[p](g∘f).
+ 🠢[g]🠡[f]p∘🠢[f]p = 🠢[g∘f]p.
#g #f #p elim p -p //
#l #p #IH
<lift_rmap_rcons //
lemma tls_plus_lift_rmap_closed (o) (f) (q) (n):
q ϵ 𝐂❨o,n❩ →
- ∀m. ⇂*[m]f ≗ ⇂*[m+n]↑[q]f.
+ ∀m. ⇂*[m]f ≗ ⇂*[m+n]🠢[f]q.
#o #f #q #n #Hq elim Hq -q -n //
#q #n #_ #IH #m
<nplus_succ_dx <stream_tls_swap //
lemma tls_lift_rmap_closed (o) (f) (q) (n):
q ϵ 𝐂❨o,n❩ →
- f ≗ ⇂*[n]↑[q]f.
+ f ≗ ⇂*[n]🠢[f]q.
#o #f #q #n #H0
/2 width=2 by tls_plus_lift_rmap_closed/
qed-.
lemma tls_succ_lift_rmap_append_L_closed_dx (o) (f) (p) (q) (n):
q ϵ 𝐂❨o,n❩ →
- ↑[p]f ≗ ⇂*[↑n]↑[p●𝗟◗q]f.
+ 🠢[f]p ≗ ⇂*[↑n]🠢[f](p●𝗟◗q).
#o #f #p #q #n #Hq
/3 width=2 by tls_lift_rmap_closed, pcc_L_sn/
qed-.
(* Constructions with lift_eq ***********************************************)
lemma lift_rmap_eq_repl (p):
- stream_eq_repl … (λf1,f2. ↑[p]f1 ≗ ↑[p]f2).
+ stream_eq_repl … (λf1,f2. 🠢[f1]p ≗ 🠢[f2]p).
#p elim p -p //
#l #p #IH #f1 #f2 #Hf
/3 width=1 by prelift_rmap_eq_repl/
qed.
lemma tls_lift_rmap_d_dx (f) (p) (n) (k):
- ⇂*[n+k]↑[p]f ≗ ⇂*[n]↑[p◖𝗱k]f.
+ ⇂*[n+k]🠢[f]p ≗ ⇂*[n]🠢[f](p◖𝗱k).
#f #p #n #k
<lift_rmap_d_dx >nrplus_inj_dx >nrplus_inj_sn //
qed.
(* Constructions with tr_id *************************************************)
lemma lift_rmap_id (p):
- (𝐢) = ↑[p]𝐢.
+ (𝐢) = 🠢[𝐢]p.
#p elim p -p //
qed.
(* Constructions with tr_pap ************************************************)
lemma lift_rmap_pap_d_dx (f) (p) (k) (h):
- ↑[p]f@⧣❨h+k❩ = ↑[p◖𝗱k]f@⧣❨h❩+↑[p]f@⧣❨k❩.
+ 🠢[f]p@⧣❨h+k❩ = 🠢[f](p◖𝗱k)@⧣❨h❩+🠢[f]p@⧣❨k❩.
// qed.
(* *)
(**************************************************************************)
-include "delayed_updating/notation/functions/uparrow_2.ma".
+include "delayed_updating/notation/functions/uptrianglearrow_2.ma".
include "delayed_updating/syntax/label.ma".
include "ground/relocation/tr_pap_pap.ma".
interpretation
"prelift (label)"
- 'UpArrow f l = (prelift_label f l).
+ 'UpTriangleArrow f l = (prelift_label f l).
(* Basic constructions ******************************************************)
lemma prelift_label_d (f) (k):
- (𝗱(f@⧣❨k❩)) = ↑[f]𝗱k.
+ (𝗱(f@⧣❨k❩)) = 🠡[f]𝗱k.
// qed.
lemma prelift_label_m (f):
- (𝗺) = ↑[f]𝗺.
+ (𝗺) = 🠡[f]𝗺.
// qed.
lemma prelift_label_L (f):
- (𝗟) = ↑[f]𝗟.
+ (𝗟) = 🠡[f]𝗟.
// qed.
lemma prelift_label_A (f):
- (𝗔) = ↑[f]𝗔.
+ (𝗔) = 🠡[f]𝗔.
// qed.
lemma prelift_label_S (f):
- (𝗦) = ↑[f]𝗦.
+ (𝗦) = 🠡[f]𝗦.
// qed.
(* Basic inversions *********************************************************)
lemma prelift_label_inv_d_sn (f) (l) (k1):
- (𝗱k1) = ↑[f]l →
+ (𝗱k1) = 🠡[f]l →
∃∃k2. k1 = f@⧣❨k2❩ & 𝗱k2 = l.
#f * [ #k ] #k1
[ <prelift_label_d
qed-.
lemma prelift_label_inv_m_sn (f) (l):
- (𝗺) = ↑[f]l → 𝗺 = l.
+ (𝗺) = 🠡[f]l → 𝗺 = l.
#f * [ #k ]
[ <prelift_label_d
| <prelift_label_m
qed-.
lemma prelift_label_inv_L_sn (f) (l):
- (𝗟) = ↑[f]l → 𝗟 = l.
+ (𝗟) = 🠡[f]l → 𝗟 = l.
#f * [ #k ]
[ <prelift_label_d
| <prelift_label_m
qed-.
lemma prelift_label_inv_A_sn (f) (l):
- (𝗔) = ↑[f]l → 𝗔 = l.
+ (𝗔) = 🠡[f]l → 𝗔 = l.
#f * [ #k ]
[ <prelift_label_d
| <prelift_label_m
qed-.
lemma prelift_label_inv_S_sn (f) (l):
- (𝗦) = ↑[f]l → 𝗦 = l.
+ (𝗦) = 🠡[f]l → 𝗦 = l.
#f * [ #k ]
[ <prelift_label_d
| <prelift_label_m
(* Main inversions **********************************************************)
theorem prelift_label_inj (f) (l1) (l2):
- ↑[f]l1 = ↑[f]l2 → l1 = l2.
+ 🠡[f]l1 = 🠡[f]l2 → l1 = l2.
#f * [ #k1 ] #l2 #Hl
[ elim (prelift_label_inv_d_sn … Hl) -Hl #k2 #Hk #H0 destruct
>(tr_pap_inj ???? Hk) -Hk //
(* Constructions with tr_after **********************************************)
lemma prelift_label_after (g) (f) (l):
- ↑[g]↑[f]l = ↑[g∘f]l.
+ 🠡[g]🠡[f]l = 🠡[g∘f]l.
#g #f * [ #k ] //
qed.
(* constructions with tr_map_eq *********************************************)
lemma prelift_label_eq_repl (l):
- stream_eq_repl … (λf1,f2. ↑[f1]l = ↑[f2]l).
+ stream_eq_repl … (λf1,f2. 🠡[f1]l = 🠡[f2]l).
* //
#k #f1 #f2 #Hf
<prelift_label_d <prelift_label_d
(* Constructions with tr_id *************************************************)
lemma prelift_label_id (l):
- l = ↑[𝐢]l.
+ l = 🠡[𝐢]l.
* [ #k ] //
qed.
(* *)
(**************************************************************************)
-include "delayed_updating/notation/functions/uparrow_2.ma".
+include "delayed_updating/notation/functions/righttrianglearrow_2.ma".
include "delayed_updating/syntax/label.ma".
include "ground/relocation/tr_pn.ma".
include "ground/lib/stream_tls.ma".
interpretation
"prelift (relocation map)"
- 'UpArrow l f = (prelift_rmap f l).
+ 'RightTriangleArrow f l = (prelift_rmap f l).
(* Basic constructions ******************************************************)
lemma prelift_rmap_d (f) (k:pnat):
- ⇂*[k]f = ↑[𝗱k]f.
+ ⇂*[k]f = 🠢[f]𝗱k.
// qed.
lemma prelift_rmap_m (f):
- f = ↑[𝗺]f.
+ f = 🠢[f]𝗺.
// qed.
lemma prelift_rmap_L (f):
- (⫯f) = ↑[𝗟]f.
+ (⫯f) = 🠢[f]𝗟.
// qed.
lemma prelift_rmap_A (f):
- f = ↑[𝗔]f.
+ f = 🠢[f]𝗔.
// qed.
lemma prelift_rmap_S (f):
- f = ↑[𝗦]f.
+ f = 🠢[f]𝗦.
// qed.
(* Constructions with tr_after **********************************************)
lemma prelift_rmap_after (g) (f) (l):
- ↑[↑[f]l]g∘↑[l]f = ↑[l](g∘f).
+ 🠢[g]🠡[f]l∘🠢[f]l = 🠢[g∘f]l.
#g #f * [ #k ] //
<prelift_rmap_d //
qed.
(* constructions with tr_map_eq *********************************************)
lemma prelift_rmap_eq_repl (l):
- stream_eq_repl … (λf1,f2. ↑[l]f1 ≗ ↑[l]f2).
+ stream_eq_repl … (λf1,f2. 🠢[f1]l ≗ 🠢[f2]l).
* //
qed.
(* Constructions with tr_id *************************************************)
lemma prelift_rmap_id (l):
- (𝐢) = ↑[l]𝐢.
+ (𝐢) = 🠢[𝐢]l.
* [ #k ] //
qed.
(* Constructions with lift_path *********************************************)
lemma preunwind2_lift_rmap_after (g) (f) (l):
- ▶[g]↑[f]l∘↑[l]f ≗ ▶[g∘f]l.
+ ▶[g]🠡[f]l∘🠢[f]l ≗ ▶[g∘f]l.
#g #f * // #k
<prelift_label_d <prelift_rmap_d <preunwind2_rmap_d <preunwind2_rmap_d
@(stream_eq_trans … (tr_compose_assoc …))
(* Constructions with lift_path *********************************************)
lemma lift_unwind2_path_after (g) (f) (p):
- ↑[g]▼[f]p = ▼[g∘f]p.
+ 🠡[g]▼[f]p = ▼[g∘f]p.
#g #f * // * [ #k ] #p //
<unwind2_path_d_dx <unwind2_path_d_dx <lift_path_d_dx
<lift_path_structure >tr_compose_pap
qed.
lemma unwind2_lift_path_after (g) (f) (p):
- ▼[g]↑[f]p = ▼[g∘f]p.
+ ▼[g]🠡[f]p = ▼[g∘f]p.
#g #f * // * [ #k ] #p
[ <unwind2_path_d_dx <unwind2_path_d_dx
<structure_lift_path >tr_compose_pap
(* Constructions with lift_prototerm ****************************************)
lemma lift_unwind2_term_after (f1) (f2) (t):
- ↑[f2]▼[f1]t ⇔ ▼[f2∘f1]t.
+ 🠡[f2]▼[f1]t ⇔ ▼[f2∘f1]t.
#f1 #f2 #t @subset_eq_trans
[| @subset_inclusion_ext_f1_compose ]
@subset_equivalence_ext_f1_exteq #p
qed.
lemma unwind2_lift_term_after (f1) (f2) (t):
- ▼[f2]↑[f1]t ⇔ ▼[f2∘f1]t.
+ ▼[f2]🠡[f1]t ⇔ ▼[f2∘f1]t.
#f1 #f2 #t @subset_eq_trans
[| @subset_inclusion_ext_f1_compose ]
@subset_equivalence_ext_f1_exteq #p
include "delayed_updating/syntax/path_closed.ma".
include "delayed_updating/syntax/path_depth.ma".
include "ground/relocation/xap.ma".
+include "ground/lib/stream_tls_plus.ma".
include "ground/lib/stream_eq_eq.ma".
include "ground/arith/nat_le_plus.ma".
include "ground/arith/nat_le_pred.ma".
/2 width=2 by unwind2_rmap_push_closed_nap/
qed-.
+lemma unwind2_rmap_append_closed_true_dx_xap_depth (f) (p) (q) (n):
+ q ϵ 𝐂❨Ⓣ,n❩ → ♭q = ▶[f](p●q)@❨n❩.
+#f #p #q #n #Hq elim Hq -q -n //
+#q #n #k #Ho #_ #IH
+<unwind2_rmap_d_dx <tr_compose_xap
+>Ho // <tr_uni_xap_succ <Ho //
+qed-.
+
lemma tls_succ_plus_unwind2_rmap_push_closed (o) (f) (q) (n):
q ϵ 𝐂❨o,n❩ →
∀m. ⇂*[m]f ≗ ⇂*[↑(m+n)]▶[⫯f]q.
#o #f #q #n #Hq elim Hq -q -n //
-#q #n [ #k #_ ] #_ #IH #m
-[ @(stream_eq_trans … (tls_unwind2_rmap_d_dx …))
- >nrplus_inj_dx >nrplus_inj_sn >nsucc_unfold //
-| <unwind2_rmap_L_dx <nplus_succ_dx //
-]
+#q #n #k #_ #_ #IH #m
+@(stream_eq_trans … (tls_unwind2_rmap_d_dx …))
+>nrplus_inj_dx >nrplus_inj_sn >nsucc_unfold //
+qed-.
+
+lemma tls_plus_unwind2_rmap_closed_true (f) (q) (n):
+ q ϵ 𝐂❨Ⓣ,n❩ →
+ ∀m. ⇂*[m]f ≗ ⇂*[m+n]▶[f]q.
+#f #q #n #Hq elim Hq -q -n //
+#q #n #k #Ho #_ #IH #m
+>Ho // <nplus_succ_dx
+@(stream_eq_trans … (tls_unwind2_rmap_d_dx …))
+>nrplus_inj_dx >nrplus_inj_sn >nsucc_unfold
+>nplus_succ_dx <Ho //
qed-.
lemma tls_succ_unwind2_rmap_append_closed_Lq_dx (o) (f) (p) (q) (n):
/2 width=2 by tls_succ_plus_unwind2_rmap_push_closed/
qed-.
+lemma tls_unwind2_rmap_append_closed_true_dx (f) (p) (q) (n):
+ q ϵ 𝐂❨Ⓣ,n❩ →
+ ▶[f]p ≗ ⇂*[n]▶[f](p●q).
+/2 width=1 by tls_plus_unwind2_rmap_closed_true/
+qed-.
+
lemma unwind2_rmap_append_closed_Lq_dx_nap_plus (o) (f) (p) (q) (m) (n):
q ϵ 𝐂❨o,n❩ →
▶[f]p@❨m❩+♭q = ▶[f](p●𝗟◗q)@§❨m+n❩.
| /2 width=2 by unwind2_rmap_append_closed_Lq_dx_nap_depth/
]
qed-.
+
+lemma unwind2_rmap_append_closed_bLq_dx_nap_plus (o) (f) (p) (b) (q) (m) (n):
+ b ϵ 𝐂❨Ⓣ,m❩ → q ϵ 𝐂❨o,n❩ →
+ ♭b+♭q = ▶[f](p●b●𝗟◗q)@§❨m+n❩.
+#o #f #p #b #q #m #n #Hb #Hq
+>(unwind2_rmap_append_closed_true_dx_xap_depth f p … Hb) -Hb
+>(unwind2_rmap_append_closed_Lq_dx_nap_plus … Hq) -Hq //
+qed-.
+
+lemma tls_succ_plus_unwind2_rmap_append_closed_bLq_dx (o) (f) (p) (b) (q) (m) (n):
+ b ϵ 𝐂❨Ⓣ,m❩ → q ϵ 𝐂❨o,n❩ →
+ ▶[f]p ≗ ⇂*[↑(m+n)]▶[f](p●b●𝗟◗q).
+#o #f #p #b #q #m #n #Hb #Hq
+>nplus_succ_dx <stream_tls_plus
+@(stream_eq_trans … (tls_unwind2_rmap_append_closed_true_dx … Hb)) -Hb
+@stream_tls_eq_repl
+@(stream_eq_trans … (tls_succ_unwind2_rmap_append_closed_Lq_dx … Hq)) -Hq //
+qed-.
(* Constructions with lift_path *********************************************)
lemma lift_unwind2_rmap_after (g) (f) (p):
- ↑[⊗p]g∘▶[f]p ≗ ▶[g∘f]p.
+ 🠢[g]⊗p∘▶[f]p ≗ ▶[g∘f]p.
#g #f #p elim p -p //
* [ #k ] #p #IH //
[ <unwind2_rmap_d_dx <unwind2_rmap_d_dx
qed.
lemma unwind2_lift_rmap_after (g) (f) (p:path):
- ▶[g]↑[f]p∘↑[p]f ≗ ▶[g∘f]p.
+ ▶[g]🠡[f]p∘🠢[f]p ≗ ▶[g∘f]p.
#g #f #p elim p -p // #l #p #IH
<lift_path_rcons <lift_rmap_rcons <unwind2_rmap_rcons <unwind2_rmap_rcons
@(stream_eq_trans … (preunwind2_lift_rmap_after …))