]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_preterm_eq.ma
update in delayed_updating
[helm.git] / matita / matita / contribs / lambdadelta / delayed_updating / substitution / lift_preterm_eq.ma
index 91facd7c591ad86571d9aec7146f50eeca7251bf..8143f860051118b38bbde6d5aa2e27ce6c7adb9c 100644 (file)
@@ -17,32 +17,69 @@ 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_S_sn (f) (t) (p):
-      ↑[↑[p]f](t⋔(p◖𝗦)) ⊆ (↑[f]t)⋔((⊗p)◖𝗦).
-#f #t #p #q * #r #Hr #H0 destruct
+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
-<list_append_rcons_sn <list_append_rcons_sn
-<lift_append_proper_dx //
+<lift_append_inner_sn //
 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 //
+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
-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)◖𝗦).
-/3 width=1 by lift_grafted_S_sn, conj, lift_grafted_S_dx/ 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 //
++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
+-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)◖������).
+-/3 width=1 by lift_grafted_S_sn, conj, lift_grafted_S_dx/ 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.
+
+*)