--- /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/unwind/unwind2_constructors.ma".
+include "delayed_updating/unwind/unwind2_preterm_fsubst.ma".
+include "delayed_updating/unwind/unwind2_preterm_eq.ma".
+include "delayed_updating/unwind/unwind2_prototerm_lift.ma".
+include "delayed_updating/unwind/unwind2_rmap_head.ma".
+*)
+include "delayed_updating/substitution/fsubst_eq.ma".
+include "delayed_updating/substitution/lift_prototerm_eq.ma".
+
+include "delayed_updating/syntax/prototerm_proper_constructors.ma".
+
+(* DELAYED FOCUSED REDUCTION ************************************************)
+
+(* Constructions with lift **************************************************)
+
+theorem dfr_lift_bi (f) (p) (q) (t1) (t2): (*t1 ϵ 𝐓 → *)
+ t1 ➡𝐝𝐟[p,q] t2 → ↑[f]t1 ➡𝐟[↑[f]p,↑[↑[p◖𝗔◖𝗟]f]q] ↑[f]t2.
+#f #p #q #t1 #t2
+* #n * #H1n #Ht1 #Ht2
+@(ex_intro … ((↑[p●𝗔◗𝗟◗q]f)@⧣❨n❩)) @and3_intro
+[ -Ht1 -Ht2
+ >H1n >path_head_structure_depth <H1n -H1n //
+| lapply (in_comp_unwind2_path_term f … Ht1) -Ht2 -Ht1 -H0t1
+ <unwind2_path_d_dx <depth_structure
+ >list_append_rcons_sn in H1n; <reverse_append #H1n
+ lapply (unwind2_rmap_append_pap_closed f … H1n)
+ <reverse_lcons <depth_L_dx #H2n
+ lapply (eq_inv_ninj_bi … H2n) -H2n #H2n <H2n -H2n -H1n #Ht1 //
+| lapply (unwind2_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
+ @(subset_eq_trans … Ht2) -t2
+ @(subset_eq_trans … (unwind2_term_fsubst …))
+ [ @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 … (unwind2_lift_term_after …))
+ @unwind2_term_eq_repl_sn
+(* Note: crux of the proof begins *)
+ @nstream_eq_inv_ext #m
+ <tr_compose_pap <tr_compose_pap
+ <tr_uni_pap <tr_uni_pap <tr_pap_plus
+ >list_append_rcons_sn in H1n; <reverse_append #H1n
+ lapply (unwind2_rmap_append_pap_closed f … H1n) #H2n
+ >nrplus_inj_dx in ⊢ (???%); <H2n -H2n
+ lapply (tls_unwind2_rmap_append_closed f … H1n) #H2n
+ <(tr_pap_eq_repl … H2n) -H2n -H1n //
+(* Note: crux of the proof ends *)
+ | //
+ | /2 width=2 by ex_intro/
+ | //
+ ]
+]
+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/ifr.ma".
+
+include "delayed_updating/unwind/unwind2_constructors.ma".
+include "delayed_updating/unwind/unwind2_preterm_fsubst.ma".
+include "delayed_updating/unwind/unwind2_preterm_eq.ma".
+include "delayed_updating/unwind/unwind2_prototerm_lift.ma".
+include "delayed_updating/unwind/unwind2_rmap_head.ma".
+
+include "delayed_updating/substitution/fsubst_eq.ma".
+include "delayed_updating/substitution/lift_prototerm_eq.ma".
+
+include "delayed_updating/syntax/prototerm_proper_constructors.ma".
+include "delayed_updating/syntax/path_head_structure.ma".
+include "delayed_updating/syntax/path_structure_depth.ma".
+include "delayed_updating/syntax/path_structure_reverse.ma".
+include "delayed_updating/syntax/path_depth_reverse.ma".
+
+(* IMMEDIATE FOCUSED REDUCTION **********************************************)
+
+(* Constructions with unwind ************************************************)
+
+theorem ifr_unwind_bi (f) (p) (q) (t1) (t2):
+ t1 ϵ 𝐓 → t1⋔(p◖𝗦) ϵ 𝐏 →
+ t1 ➡𝐟[p,q] t2 → ▼[f]t1 ➡𝐟[⊗p,⊗q] ▼[f]t2.
+#f #p #q #t1 #t2 #H1t1 #H2t1
+* #n * #H1n #Ht1 #Ht2
+@(ex_intro … (↑♭⊗q)) @and3_intro
+[ -H0t1 -Ht1 -Ht2
+ >structure_L_sn >structure_reverse
+ >H1n >path_head_structure_depth <H1n -H1n //
+| lapply (in_comp_unwind2_path_term f … Ht1) -Ht2 -Ht1 -H0t1
+ <unwind2_path_d_dx <depth_structure
+ >list_append_rcons_sn in H1n; <reverse_append #H1n
+ lapply (unwind2_rmap_append_pap_closed f … H1n)
+ <reverse_lcons <depth_L_dx #H2n
+ lapply (eq_inv_ninj_bi … H2n) -H2n #H2n <H2n -H2n -H1n #Ht1 //
+| lapply (unwind2_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
+ @(subset_eq_trans … Ht2) -t2
+ @(subset_eq_trans … (unwind2_term_fsubst …))
+ [ @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 … (unwind2_lift_term_after …))
+ @unwind2_term_eq_repl_sn
+(* Note: crux of the proof begins *)
+ @nstream_eq_inv_ext #m
+ <tr_compose_pap <tr_compose_pap
+ <tr_uni_pap <tr_uni_pap <tr_pap_plus
+ >list_append_rcons_sn in H1n; <reverse_append #H1n
+ lapply (unwind2_rmap_append_pap_closed f … H1n) #H2n
+ >nrplus_inj_dx in ⊢ (???%); <H2n -H2n
+ lapply (tls_unwind2_rmap_append_closed f … H1n) #H2n
+ <(tr_pap_eq_repl … H2n) -H2n -H1n //
+(* Note: crux of the proof ends *)
+*)
+ | //
+ | /2 width=2 by ex_intro/
+ | //
+ ]
+]
+qed.
include "ground/relocation/tr_uni.ma".
include "ground/relocation/tr_pap_tls.ma".
-(* LIFT FOR PATH ***********************************************************)
+(* LIFT FOR PATH ************************************************************)
definition lift_continuation (A:Type[0]) ≝
tr_map → path → A.
+++ /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.
(**************************************************************************)
include "delayed_updating/substitution/lift_prototerm_id.ma".
-include "delayed_updating/substitution/lift_uni.ma".
+include "delayed_updating/substitution/lift_path_uni.ma".
include "delayed_updating/syntax/prototerm_constructors.ma".
(* LIFT FOR PROTOTERM *******************************************************)
include "ground/relocation/tr_pn_eq.ma".
include "ground/lib/stream_tls_eq.ma".
-(* LIFT FOR PATH ***********************************************************)
+(* LIFT FOR PATH ************************************************************)
definition lift_exteq (A): relation2 (lift_continuation A) (lift_continuation A) ≝
λk1,k2. ∀f1,f2,p. f1 ≗ f2 → k1 f1 p = k2 f2 p.
<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 "delayed_updating/substitution/lift_eq.ma".
-include "ground/relocation/tr_id_pap.ma".
-include "ground/relocation/tr_id_tls.ma".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Constructions with tr_id ************************************************)
-
-lemma lift_path_id (p):
- p = ↑[𝐢]p.
-#p elim p -p //
-* [ #n ] #p #IH //
-[ <lift_path_d_sn //
-| <lift_path_L_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 "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.
--- /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_id_pap.ma".
+include "ground/relocation/tr_id_tls.ma".
+
+(* LIFT FOR PATH ************************************************************)
+
+(* Constructions with tr_id *************************************************)
+
+lemma lift_path_id (p):
+ p = ↑[𝐢]p.
+#p elim p -p //
+* [ #n ] #p #IH //
+[ <lift_path_d_sn //
+| <lift_path_L_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_proper.ma".
+
+(* LIFT FOR PATH ************************************************************)
+
+(* Constructions with proper condition for path *****************************)
+
+lemma lift_path_proper (f) (p):
+ p ϵ 𝐏 → ↑[f]p ϵ 𝐏.
+#f *
+[ #H0 elim (ppc_inv_empty … H0)
+| * [ #n ] #p #_
+ [ <lift_path_d_sn /2 width=3 by ppc_lcons/
+ | <lift_path_m_sn /2 width=3 by ppc_lcons/
+ | <lift_path_L_sn /2 width=3 by ppc_lcons/
+ | <lift_path_A_sn /2 width=3 by ppc_lcons/
+ | <lift_path_S_sn /2 width=3 by ppc_lcons/
+ ]
+]
+qed.
+
+(* Inversions with proper condition for path ********************************)
+
+lemma lift_path_inv_proper (f) (p):
+ ↑[f]p ϵ 𝐏 → p ϵ 𝐏.
+#f * //
+#H0 elim (ppc_inv_empty … H0)
+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_path_id.ma".
+include "ground/relocation/tr_uni_pap.ma".
+include "ground/relocation/tr_uni_tls.ma".
+
+(* LIFT FOR PATH ************************************************************)
+
+(* Constructions with tr_uni ************************************************)
+
+lemma lift_path_d_sn_uni (p) (m) (n):
+ (𝗱(n+m)◗p) = ↑[𝐮❨m❩](𝗱(n)◗p).
+#p #m #n
+<lift_path_d_sn <tr_uni_pap >nsucc_pnpred
+<tr_tls_succ_uni //
+qed.
(**************************************************************************)
include "ground/lib/subset_ext_equivalence.ma".
-include "delayed_updating/substitution/lift_after.ma".
+include "delayed_updating/substitution/lift_path_after.ma".
include "delayed_updating/substitution/lift_prototerm.ma".
(* LIFT FOR PROTOTERM *******************************************************)
(**************************************************************************)
include "delayed_updating/substitution/lift_prototerm_eq.ma".
-include "delayed_updating/substitution/lift_id.ma".
+include "delayed_updating/substitution/lift_path_id.ma".
+
+(* LIFT FOR PROTOTERM *******************************************************)
+
+(* Constructions with tr_id *************************************************)
lemma lift_term_id_sn (t):
t ⊆ ↑[𝐢]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/lift_prototerm.ma".
+include "delayed_updating/substitution/lift_path_proper.ma".
+include "delayed_updating/syntax/prototerm_proper.ma".
+
+(* LIFT FOR PROTOTERM *******************************************************)
+
+(* Constructions with proper condition for path *****************************)
+
+lemma lift_term_proper (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 #Ht #p #Hp
+@(lift_path_inv_proper f)
+@Ht -Ht @in_comp_lift_bi // (**) (* auto fails *)
+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_id.ma".
-include "ground/relocation/tr_uni_pap.ma".
-include "ground/relocation/tr_uni_tls.ma".
-
-(* LIFT FOR PATH ***********************************************************)
-
-(* Constructions with tr_uni ***********************************************)
-
-lemma lift_path_d_sn_uni (p) (m) (n):
- (𝗱(n+m)◗p) = ↑[𝐮❨m❩](𝗱(n)◗p).
-#p #m #n
-<lift_path_d_sn <tr_uni_pap >nsucc_pnpred
-<tr_tls_succ_uni //
-qed.