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".
t1 ➡𝐝𝐟[p,q] t2 → ▼[f]t1 ➡𝐟[⊗p,⊗q] ▼[f]t2.
#f #p #q #t1 #t2 #H0t1
* #n * #H1n #Ht1 #Ht2
-@(ex_intro … (↑♭⊗q)) @and3_intro
+@(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
+ <unwind2_path_d_dx
>list_append_rcons_sn in H1n; <reverse_append #H1n
lapply (unwind2_rmap_append_pap_closed f … H1n)
<reverse_lcons <depth_L_dx #H2n
@(subset_eq_trans … (unwind2_term_fsubst …))
[ @fsubst_eq_repl [ // | // ]
@(subset_eq_trans … (unwind2_term_iref …))
- @(subset_eq_canc_sn … (lift_term_eq_repl_dx …))
+ @(subset_eq_canc_sn … (unwind2_term_eq_repl_dx …))
[ @unwind2_term_grafted_S /2 width=2 by ex_intro/ | skip ] -Ht1
- @(subset_eq_trans … (unwind2_lift_term_after …))
+ @(subset_eq_trans … (unwind2_term_after …))
@unwind2_term_eq_repl_sn
(* Note: crux of the proof begins *)
@nstream_eq_inv_ext #m
>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 //
+ lapply (tls_unwind2_rmap_append_closed f … H1n) -H1n #H2n
+ <(tr_pap_eq_repl … H2n) -H2n //
(* Note: crux of the proof ends *)
| //
| /2 width=2 by ex_intro/
(* *)
(**************************************************************************)
+include "delayed_updating/unwind/unwind2_prototerm.ma".
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_head.ma".
include "delayed_updating/syntax/path_depth.ma".
λt1,t2. ∃n:pnat.
let r ≝ p●𝗔◗𝗟◗q in
∧∧ (𝗟◗q)ᴿ = ↳[n](rᴿ) & r◖𝗱n ϵ t1 &
- t1[â\8b\94râ\86\90â\86\91[𝐮❨♭(𝗟◗q)❩](t1⋔(p◖𝗦))] ⇔ t2
+ t1[â\8b\94râ\86\90â\96¼[𝐮❨♭(𝗟◗q)❩](t1⋔(p◖𝗦))] ⇔ t2
.
interpretation
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_prototerm_inner.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/prototerm_proper_inner.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".
(* Constructions with unwind ************************************************)
theorem ifr_unwind_bi (f) (p) (q) (t1) (t2):
- t1 ϵ 𝐓 → t1⋔(p◖𝗦) ϵ 𝐏 →
+ t1 ϵ 𝐓 → t1⋔(p◖𝗦) ⧸≬ 𝐈 →
t1 ➡𝐟[p,q] t2 → ▼[f]t1 ➡𝐟[⊗p,⊗q] ▼[f]t2.
-#f #p #q #t1 #t2 #H1t1 #H2t1
+#f #p #q #t1 #t2 #H1t1 #H2t1
* #n * #H1n #Ht1 #Ht2
-@(ex_intro … (↑♭⊗q)) @and3_intro
-[ -H0t1 -Ht1 -Ht2
+@(ex_intro … (↑♭q)) @and3_intro
+[ -H1t1 -H2t1 -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
+| lapply (in_comp_unwind2_path_term f … Ht1) -Ht2 -Ht1 -H1t1 -H2t1
+ <unwind2_path_d_dx
>list_append_rcons_sn in H1n; <reverse_append #H1n
lapply (unwind2_rmap_append_pap_closed f … H1n)
<reverse_lcons <depth_L_dx #H2n
@(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 …))
+ @(subset_eq_canc_dx … (unwind2_term_after …))
+ @(subset_eq_canc_sn … (unwind2_term_eq_repl_dx …))
[ @unwind2_term_grafted_S /2 width=2 by ex_intro/ | skip ] -Ht1
- @(subset_eq_trans … (unwind2_lift_term_after …))
+ @(subset_eq_trans … (unwind2_term_after …))
@unwind2_term_eq_repl_sn
(* Note: crux of the proof begins *)
@nstream_eq_inv_ext #m
>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 //
+ lapply (tls_unwind2_rmap_append_closed f … H1n) -H1n #H2n
+ <(tr_pap_eq_repl … H2n) -H2n //
(* Note: crux of the proof ends *)
-*)
| //
| /2 width=2 by ex_intro/
- | //
+ | @term_proper_outer #H0 (**) (* full auto does not work *)
+ /3 width=2 by unwind2_term_des_inner/
]
]
qed.
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".
+include "ground/relocation/tr_id_pap.ma".
(* LIFT FOR PATH ************************************************************)
[ list_empty ⇒ k f (𝐞)
| list_lcons l q ⇒
match l with
- [ label_d n ⇒ lift_gen (A) (λg,p. k g (𝗱(f@⧣❨n❩)◗p)) (⇂*[n]f) q
+ [ label_d n ⇒ lift_gen (A) (λg,p. k g (𝗱(f@⧣❨n❩)◗p)) (𝐢) 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
// qed.
lemma lift_d_sn (A) (k) (p) (n) (f):
- ↑❨(λg,p. k g (𝗱(f@⧣❨n❩)◗p)), ⇂*[n]f, p❩ = ↑{A}❨k, f, 𝗱n◗p❩.
+ ↑❨(λg,p. k g (𝗱(f@⧣❨n❩)◗p)), 𝐢, p❩ = ↑{A}❨k, f, 𝗱n◗p❩.
// qed.
lemma lift_m_sn (A) (k) (p) (f):
// qed.
lemma lift_rmap_d_sn (f) (p) (n):
- ↑[p](⇂*[ninj n]f) = ↑[𝗱n◗p]f.
+ ↑[p]𝐢 = ↑[𝗱n◗p]f.
// qed.
lemma lift_rmap_m_sn (f) (p):
↑[p]f = ↑[𝗦◗p]f.
// qed.
+(* Advanced cinstructionswith proj_rmap and tr_id ***************************)
+
+lemma lift_rmap_id (p):
+ (𝐢) = ↑[p]𝐢.
+#p elim p -p //
+* [ #n ] #p #IH //
+qed.
+
(* Advanced constructions with proj_rmap and path_append ********************)
lemma lift_rmap_append (p2) (p1) (f):
↑[p2]↑[p1]f = ↑[p1●p2]f.
#p2 #p1 elim p1 -p1 // * [ #n ] #p1 #IH #f //
-[ <lift_rmap_m_sn <lift_rmap_m_sn //
+[ <lift_rmap_d_sn <lift_rmap_d_sn //
+| <lift_rmap_m_sn <lift_rmap_m_sn //
| <lift_rmap_A_sn <lift_rmap_A_sn //
| <lift_rmap_S_sn <lift_rmap_S_sn //
]
(* Advanced constructions with proj_rmap and path_rcons *********************)
lemma lift_rmap_d_dx (f) (p) (n):
- ⇂*[ninj n](↑[p]f) = ↑[p◖𝗱n]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❩+↑[p]f@⧣❨n❩.
+ m = ↑[p◖𝗱n]f@⧣❨m❩.
// qed.
(* *)
(**************************************************************************)
-include "delayed_updating/substitution/lift_prototerm_id.ma".
+include "delayed_updating/substitution/lift_prototerm_eq.ma".
include "delayed_updating/substitution/lift_path_uni.ma".
include "delayed_updating/syntax/prototerm_constructors.ma".
qed.
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))
+ (𝛗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 lift_iref_dx (f) (t) (n:pnat):
- ↑[f](𝛗n.t) ⊆ 𝛗f@⧣❨n❩.↑[⇂*[n]f]t.
+ ↑[f](𝛗n.t) ⊆ 𝛗f@⧣❨n❩.t.
#f #t #n #p * #q #Hq #H0 destruct
elim (in_comp_inv_iref … Hq) -Hq #p #H0 #Hp destruct
-/3 width=1 by in_comp_iref, in_comp_lift_bi/
+/2 width=1 by in_comp_iref/
qed-.
lemma lift_iref (f) (t) (n:pnat):
- (𝛗f@⧣❨n❩.↑[⇂*[n]f]t) ⇔ ↑[f](𝛗n.t).
+ (𝛗f@⧣❨n❩.t) ⇔ ↑[f](𝛗n.t).
/3 width=1 by conj, lift_iref_sn, lift_iref_dx/
qed.
lemma lift_iref_uni (t) (m) (n):
(𝛗(n+m).t) ⇔ ↑[𝐮❨m❩](𝛗n.t).
-#t #m #n
-@(subset_eq_trans … (lift_iref …))
-<tr_uni_pap >nsucc_pnpred <tr_tls_succ_uni
-/3 width=1 by lift_iref_bi, lift_term_id/
-qed.
+// qed.
include "delayed_updating/substitution/lift.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 ************************************************************)
#k1 #k2 #Hk #f1 #f2 #Hf
[ <lift_empty <lift_empty /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 stream_eq_refl/
| /3 width=1 by/
| /3 width=1 by tr_push_eq_repl/
| /3 width=1 by/
qed.
lemma lift_path_d_sn (f) (p) (n):
- (𝗱(f@⧣❨n❩)◗↑[⇂*[n]f]p) = ↑[f](𝗱n◗p).
+ (𝗱(f@⧣❨n❩)◗↑[𝐢]p) = ↑[f](𝗱n◗p).
// qed.
lemma lift_path_m_sn (f) (p):
(𝗦◗↑[f]p) = ↑[f](𝗦◗p).
// qed.
+lemma lift_path_id (p):
+ p = ↑[𝐢]p.
+#p elim p -p //
+* [ #n ] #p #IH //
+[ <lift_path_d_sn //
+| <lift_path_L_sn //
+]
+qed.
+
lemma lift_path_append (p2) (p1) (f):
(↑[f]p1)●(↑[↑[p1]f]p2) = ↑[f](p1●p2).
#p2 #p1 elim p1 -p1 //
]
qed.
-lemma lift_path_d_dx (f) (p) (n):
+lemma lift_path_d_dx (n) (p) (f):
(↑[f]p)◖𝗱((↑[p]f)@⧣❨n❩) = ↑[f](p◖𝗱n).
-#f #p #n <lift_path_append //
+#n #p #f <lift_path_append //
qed.
-lemma lift_path_m_dx (f) (p):
+lemma lift_path_m_dx (p) (f):
(↑[f]p)◖𝗺 = ↑[f](p◖𝗺).
-#f #p <lift_path_append //
+#p #f <lift_path_append //
qed.
-lemma lift_path_L_dx (f) (p):
+lemma lift_path_L_dx (p) (f):
(↑[f]p)◖𝗟 = ↑[f](p◖𝗟).
-#f #p <lift_path_append //
+#p #f <lift_path_append //
qed.
-lemma lift_path_A_dx (f) (p):
+lemma lift_path_A_dx (p) (f):
(↑[f]p)◖𝗔 = ↑[f](p◖𝗔).
-#f #p <lift_path_append //
+#p #f <lift_path_append //
qed.
-lemma lift_path_S_dx (f) (p):
+lemma lift_path_S_dx (p) (f):
(↑[f]p)◖𝗦 = ↑[f](p◖𝗦).
-#f #p <lift_path_append //
+#p #f <lift_path_append //
qed.
(* COMMENT
<lift_rmap_d_dx >nrplus_inj_dx
/2 width=1 by tr_tls_compose_uni_dx/
qed.
-
*)
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 ************************************************************)
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 //
-]
+#p elim p -p [| * ] // #p #IH #f1 #f2
+<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 "ground/lib/list_length.ma".
+
+(* LIFT FOR PATH ************************************************************)
+
+(* Constructions with list_length *******************************************)
+
+lemma lift_path_length (p) (f):
+ ❘p❘ = ❘↑[f]p❘.
+#p elim p -p // * [ #n ] #p #IH #f //
+[ <lift_path_m_sn
+| <lift_path_L_sn
+| <lift_path_A_sn
+| <lift_path_S_sn
+]
+>list_length_lcons >list_length_lcons //
+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".
+
+(* LIFT FOR PATH ************************************************************)
+
+(* Constructions with structure *********************************************)
+
+lemma structure_lift_path (p) (f):
+ ⊗p = ⊗↑[f]p.
+#p elim p -p //
+* [ #n ] #p #IH #f //
+[ <lift_path_m_sn <structure_m_sn <structure_m_sn //
+| <lift_path_L_sn <structure_L_sn <structure_L_sn //
+]
+qed.
+
+lemma lift_path_structure (p) (f):
+ ⊗p = ↑[f]⊗p.
+#p elim p -p //
+* [ #n ] #p #IH #f //
+qed.
(* *)
(**************************************************************************)
-include "delayed_updating/substitution/lift_path_id.ma".
+include "delayed_updating/substitution/lift_eq.ma".
include "ground/relocation/tr_uni_pap.ma".
-include "ground/relocation/tr_uni_tls.ma".
(* LIFT FOR PATH ************************************************************)
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.
+// qed.
| @subset_equivalence_ext_f1_exteq /2 width=5/
]
qed.
+
+lemma lift_term_id_sn (t):
+ t ⊆ ↑[𝐢]t.
+#t #p #Hp
+>(lift_path_id p)
+/2 width=1 by in_comp_lift_bi/
+qed-.
+
+lemma lift_term_id_dx (t):
+ ↑[𝐢]t ⊆ t.
+#t #p * #q #Hq #H destruct //
+qed-.
+
+lemma lift_term_id (t):
+ t ⇔ ↑[𝐢]t.
+/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/
+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_prototerm_eq.ma".
-include "delayed_updating/substitution/lift_path_id.ma".
-
-(* LIFT FOR PROTOTERM *******************************************************)
-
-(* Constructions with tr_id *************************************************)
-
-lemma lift_term_id_sn (t):
- t ⊆ ↑[𝐢]t.
-#t #p #Hp
->(lift_path_id p)
-/2 width=1 by in_comp_lift_bi/
-qed-.
-
-lemma lift_term_id_dx (t):
- ↑[𝐢]t ⊆ t.
-#t #p * #q #Hq #H destruct //
-qed-.
-
-lemma lift_term_id (t):
- t ⇔ ↑[𝐢]t.
-/3 width=2 by lift_term_id_dx, lift_term_id_sn, conj/
-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/syntax/path_inner.ma".
+include "delayed_updating/syntax/path_proper.ma".
+
+(* INNER CONDITION FOR PATH *************************************************)
+
+(* Destructions with proper condition for path ******************************)
+
+lemma path_des_outer_proper (p):
+ p ⧸ϵ 𝐈 → p ϵ 𝐏.
+#p #H1 #H2 destruct
+@H1 -H1 // (**) (* auto fails *)
+qed-.
(* Basic constructions ******************************************************)
lemma ppc_lcons (l) (q): l◗q ϵ 𝐏.
-#l #p #H destruct
+#l #q #H0 destruct
+qed.
+
+lemma ppc_rcons (l) (q): q◖l ϵ 𝐏.
+#l #q #H
+elim (eq_inv_list_empty_rcons ??? H)
qed.
(* Basic inversions ********************************************************)
#H elim (Ht H)
qed.
-(* Basic inversions *********************************************************)
+(* Basic destructions *******************************************************)
-lemma in_ppc_comp_trans (t) (p):
+lemma in_comp_tpc_trans (t) (p):
p ϵ t → t ϵ 𝐏 → p ϵ 𝐏.
#t #p #Hp #Ht
@(Ht … Hp)
qed-.
-lemma tpc_e (t): 𝐞 ϵ t → t ϵ 𝐏 → ⊥.
-/2 width=5 by in_ppc_comp_trans/ qed-.
+(* Basic inversions *********************************************************)
+
+lemma tpc_inv_empty (t):
+ (𝐞) ϵ t → t ϵ 𝐏 → ⊥.
+/2 width=5 by in_comp_tpc_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/syntax/prototerm_proper.ma".
+include "delayed_updating/syntax/path_inner_proper.ma".
+include "ground/lib/subset_overlap.ma".
+
+(* PROPER CONDITION FOR PROTOTERM *******************************************)
+
+(* Constructions with inner condition for prototerm *************************)
+
+lemma term_proper_outer (t):
+ t ⧸≬ 𝐈 → t ϵ 𝐏.
+/4 width=3 by path_des_outer_proper, subset_ol_i/
+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/unwind/unwind2_path_structure.ma".
+include "delayed_updating/syntax/path_inner.ma".
+
+(* UNWIND FOR PATH **********************************************************)
+
+(* Destructions with inner condition for path *******************************)
+
+lemma unwind2_path_des_inner (f) (p):
+ ▼[f]p ϵ 𝐈 → p ϵ 𝐈.
+#f #p @(list_ind_rcons … p) -p //
+#p * [ #n ] #_ //
+<unwind2_path_d_dx #H0
+elim (pic_inv_d_dx … H0)
+qed-.
#f #t #u #p #Hu #ql * *
[ #rl * #r #Hr #H1 #H2 destruct
>unwind2_path_append_proper_dx
- /4 width=5 by in_comp_unwind2_path_term, in_ppc_comp_trans, or_introl, ex2_intro/
+ /4 width=5 by in_comp_unwind2_path_term, in_comp_tpc_trans, or_introl, ex2_intro/
| * #q #Hq #H1 #H0
@(ex2_intro … H1) @or_intror @conj // *
[ <list_append_empty_dx #H2 destruct
#f #t #u #p #Hu #H1p #H2p #ql * #q * *
[ #r #Hu #H1 #H2 destruct
@or_introl @ex2_intro
- [|| <unwind2_path_append_proper_dx /2 width=5 by in_ppc_comp_trans/ ]
+ [|| <unwind2_path_append_proper_dx /2 width=5 by in_comp_tpc_trans/ ]
/2 width=3 by ex2_intro/
| #Hq #H0 #H1 destruct
@or_intror @conj [ /2 width=1 by in_comp_unwind2_path_term/ ] *
--- /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/unwind/unwind2_prototerm.ma".
+include "delayed_updating/unwind/unwind2_path_inner.ma".
+include "ground/lib/subset_overlap.ma".
+
+(* UNWIND FOR PROTOTERM *****************************************************)
+
+(* Destructions with inner condition for path *******************************)
+
+lemma unwind2_term_des_inner (f) (t):
+ ▼[f]t ≬ 𝐈 → t ≬ 𝐈.
+#f #t * #p * #q #Hq #H0 #Hp destruct
+@(subset_ol_i … Hq) -Hq (**) (* auto does not work *)
+@(unwind2_path_des_inner … Hp)
+qed-.