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 &
+ ∧∧ ⊗b ϵ 𝐁 & ∀f. ↑❘q❘ = (↑[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
t1[⋔r←𝛗(n+❘b❘).(t1⋔(p◖𝗦))] ⇔ t2
.
>tr_pushs_swap <tr_pap_pushs_le //
| lapply (in_comp_lift_bi f … Ht1) -Ht1 -H0t1 -Hb -Ht2
<lift_d_empty_dx //
-| lapply (eq_lift_bi f … Ht2) -Ht2 #Ht2
+| lapply (lift_term_eq_repl_dx f … Ht2) -Ht2 #Ht2
@(subset_eq_trans … Ht2) -t2
@(subset_eq_trans … (lift_fsubst …))
[ <structure_append <structure_A_sn <structure_append <structure_L_sn
include "delayed_updating/substitution/fsubst.ma".
include "delayed_updating/substitution/lift_prototerm.ma".
+include "delayed_updating/syntax/prototerm_equivalence.ma".
include "delayed_updating/syntax/path_depth.ma".
include "delayed_updating/syntax/path_structure.ma".
include "delayed_updating/syntax/path_balanced.ma".
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 &
+ ∧∧ ⊗b ϵ 𝐁 & ∀f. ↑❘q❘ = (↑[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
t1[⋔r←↑[𝐮❨❘b●𝗟◗q❘❩](t1⋔(p◖𝗦))] ⇔ t2
.
(**************************************************************************)
include "delayed_updating/substitution/fsubst.ma".
-include "delayed_updating/substitution/lift_prototerm.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".
(* *)
(**************************************************************************)
-include "ground/relocation/tr_compose_compose.ma".
-include "ground/relocation/tr_compose_pn.ma".
-include "delayed_updating/substitution/lift_eq.ma".
-
-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.
-
-include "delayed_updating/substitution/lift_prototerm.ma".
-
-axiom lift_term_after (t) (f1) (f2):
- ↑[f2]↑[f1]t ⇔ ↑[f2∘f1]t.
-
+include "delayed_updating/substitution/lift_prototerm_eq.ma".
include "delayed_updating/syntax/prototerm_constructors.ma".
(* LIFT FOR PROTOTERM *******************************************************)
-lemma lift_iref_after_sn (f) (t) (n:pnat):
- (↑[f∘𝐮❨n❩]t) ⊆ ↑[f](𝛗n.t).
+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))
/2 width=1 by in_comp_iref/
(**************************************************************************)
include "delayed_updating/substitution/lift.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".
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.
include "delayed_updating/substitution/lift.ma".
include "delayed_updating/syntax/prototerm.ma".
-include "ground/lib/subset_ext_equivalence.ma".
+include "ground/lib/subset_ext.ma".
(* LIFT FOR PROTOTERM *******************************************************)
p ϵ t → ↑[f]p ϵ ↑[f]t.
/2 width=1 by subset_in_ext_f1_dx/
qed.
-
-lemma eq_lift_bi (f) (t1) (t2):
- t1 ⇔ t2 → ↑[f]t1 ⇔ ↑[f]t2.
-/2 width=1 by subset_equivalence_ext_f1_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 "ground/lib/subset_ext_equivalence.ma".
+include "delayed_updating/substitution/lift_eq.ma".
+include "delayed_updating/substitution/lift_prototerm.ma".
+
+(* LIFT FOR PROTOTERM *******************************************************)
+
+(* Constructions with subset_equivalence ************************************)
+
+lemma lift_term_eq_repl_sn (f1) (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.
+/2 width=1 by subset_equivalence_ext_f1_bi/
+qed.
+
+lemma lift_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.
include "delayed_updating/syntax/path_structure.ma".
include "delayed_updating/syntax/path_proper.ma".
include "ground/xoa/ex_4_2.ma".
-include "ground/xoa/ex_3_2.ma".
(* LIFT FOR PATH ***********************************************************)