]> matita.cs.unibo.it Git - helm.git/commitdiff
update in delayed_updating
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 15 Jun 2022 22:47:11 +0000 (00:47 +0200)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 15 Jun 2022 22:47:11 +0000 (00:47 +0200)
+ additions for lift
+ advances in dfr_lift_bi
+ some corrections
+ some parked files removed

34 files changed:
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_constructors.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_after.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_proper.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_uni.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_proper.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_constructors.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_fsubst.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_preterm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm_eq.etc [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr.ma
matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_lift.ma
matita/matita/contribs/lambdadelta/delayed_updating/reduction/ifr.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/fsubst_eq.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_constructors.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_id.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_head.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_id.ma [deleted file]
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_uni.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_prototerm.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_prototerm_id.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_prototerm_proper.ma
matita/matita/contribs/lambdadelta/delayed_updating/syntax/bdd_term.ma
matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_head.ma
matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma
matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors_eq.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_eq.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_equivalence.ma [deleted file]
matita/matita/contribs/lambdadelta/ground/relocation/nap.ma
matita/matita/contribs/lambdadelta/ground/relocation/xap.ma

diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift.etc
deleted file mode 100644 (file)
index 6933e9d..0000000
+++ /dev/null
@@ -1,158 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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/notation/functions/uparrow_4.ma".
-include "delayed_updating/notation/functions/uparrow_2.ma".
-include "delayed_updating/syntax/path.ma".
-include "ground/relocation/tr_id_pap.ma".
-
-(* LIFT FOR PATH ************************************************************)
-
-definition lift_continuation (A:Type[0]) ≝
-           tr_map → path → A.
-
-rec definition lift_gen (A:Type[0]) (k:lift_continuation A) (f) (p) on p ≝
-match p with
-[ list_empty     ⇒ k f (𝐞)
-| list_lcons l q ⇒
-  match l with
-  [ 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
-  | label_S   ⇒ lift_gen (A) (λg,p. k g (𝗦◗p)) f q
-  ]
-].
-
-interpretation
-  "lift (gneric)"
-  'UpArrow A k f p = (lift_gen A k f p).
-
-definition proj_path: lift_continuation … ≝
-           λf,p.p.
-
-definition proj_rmap: lift_continuation … ≝
-           λf,p.f.
-
-interpretation
-  "lift (path)"
-  'UpArrow f p = (lift_gen ? proj_path f p).
-
-interpretation
-  "lift (relocation map)"
-  'UpArrow p f = (lift_gen ? proj_rmap f p).
-
-(* Basic constructions ******************************************************)
-
-lemma lift_empty (A) (k) (f):
-      k f (𝐞) = ↑{A}❨k, f, 𝐞❩.
-// qed.
-
-lemma lift_d_sn (A) (k) (p) (n) (f):
-      ↑❨(λg,p. k g (𝗱(f@⧣❨n❩)◗p)), 𝐢, p❩ = ↑{A}❨k, f, 𝗱n◗p❩.
-// qed.
-
-lemma lift_m_sn (A) (k) (p) (f):
-      ↑❨(λg,p. k g (𝗺◗p)), f, p❩ = ↑{A}❨k, f, 𝗺◗p❩.
-// qed.
-
-lemma lift_L_sn (A) (k) (p) (f):
-      ↑❨(λg,p. k g (𝗟◗p)), ⫯f, p❩ = ↑{A}❨k, f, 𝗟◗p❩.
-// qed.
-
-lemma lift_A_sn (A) (k) (p) (f):
-      ↑❨(λg,p. k g (𝗔◗p)), f, p❩ = ↑{A}❨k, f, 𝗔◗p❩.
-// qed.
-
-lemma lift_S_sn (A) (k) (p) (f):
-      ↑❨(λg,p. k g (𝗦◗p)), f, p❩ = ↑{A}❨k, f, 𝗦◗p❩.
-// qed.
-
-(* Basic constructions with proj_path ***************************************)
-
-lemma lift_path_empty (f):
-      (𝐞) = ↑[f]𝐞.
-// qed.
-
-(* Basic constructions with proj_rmap ***************************************)
-
-lemma lift_rmap_empty (f):
-      f = ↑[𝐞]f.
-// qed.
-
-lemma lift_rmap_d_sn (f) (p) (n):
-      ↑[p]𝐢 = ↑[𝗱n◗p]f.
-// qed.
-
-lemma lift_rmap_m_sn (f) (p):
-      ↑[p]f = ↑[𝗺◗p]f.
-// qed.
-
-lemma lift_rmap_L_sn (f) (p):
-      ↑[p](⫯f) = ↑[𝗟◗p]f.
-// qed.
-
-lemma lift_rmap_A_sn (f) (p):
-      ↑[p]f = ↑[𝗔◗p]f.
-// qed.
-
-lemma lift_rmap_S_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_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 //
-]
-qed.
-
-(* Advanced constructions with proj_rmap and path_rcons *********************)
-
-lemma lift_rmap_d_dx (f) (p) (n):
-      (𝐢) = ↑[p◖𝗱n]f.
-// qed.
-
-lemma lift_rmap_m_dx (f) (p):
-      ↑[p]f = ↑[p◖𝗺]f.
-// qed.
-
-lemma lift_rmap_L_dx (f) (p):
-      (⫯↑[p]f) = ↑[p◖𝗟]f.
-// qed.
-
-lemma lift_rmap_A_dx (f) (p):
-      ↑[p]f = ↑[p◖𝗔]f.
-// qed.
-
-lemma lift_rmap_S_dx (f) (p):
-      ↑[p]f = ↑[p◖𝗦]f.
-// qed.
-
-lemma lift_rmap_pap_d_dx (f) (p) (n) (m):
-      m = ↑[p◖𝗱n]f@⧣❨m❩.
-// qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_constructors.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_constructors.etc
deleted file mode 100644 (file)
index 4eeb2ab..0000000
+++ /dev/null
@@ -1,47 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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_uni.ma".
-include "delayed_updating/syntax/prototerm_constructors.ma".
-
-(* LIFT FOR PROTOTERM *******************************************************)
-
-lemma lift_iref_bi (t1) (t2) (n):
-      t1 ⇔ t2 → 𝛗n.t1 ⇔ 𝛗n.t2.
-/2 width=1 by subset_equivalence_ext_f1_bi/
-qed.
-
-lemma lift_iref_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/
-qed-.
-
-lemma lift_iref_dx (f) (t) (n:pnat):
-      ↑[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
-/2 width=1 by in_comp_iref/
-qed-.
-
-lemma lift_iref (f) (t) (n:pnat):
-      (𝛗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).
-// qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_eq.etc
deleted file mode 100644 (file)
index f6759bd..0000000
+++ /dev/null
@@ -1,164 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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.ma".
-include "ground/relocation/tr_pap_eq.ma".
-include "ground/relocation/tr_pn_eq.ma".
-
-(* 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.
-
-interpretation
-  "extensional equivalence (lift continuation)"
-  'RingEq A k1 k2 = (lift_exteq A k1 k2).
-
-(* Constructions with lift_exteq ********************************************)
-
-lemma lift_eq_repl (A) (p) (k1) (k2):
-      k1 ≗{A} k2 → stream_eq_repl … (λf1,f2. ↑❨k1, f1, p❩ = ↑❨k2, f2, p❩).
-#A #p elim p -p [| * [ #n ] #q #IH ]
-#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=1 by stream_eq_refl/
-| /3 width=1 by/
-| /3 width=1 by tr_push_eq_repl/
-| /3 width=1 by/
-| /3 width=1 by/
-]
-qed-.
-
-(* Advanced constructions ***************************************************)
-
-lemma lift_lcons_alt (A) (k) (f) (p) (l): k ≗ k →
-      ↑❨λg,p2. k g (l◗p2), f, p❩ = ↑{A}❨λg,p2. k g ((l◗𝐞)●p2), f, p❩.
-#A #k #f #p #l #Hk
-@lift_eq_repl // #g1 #g2 #p2 #Hg @Hk -Hk // (**) (* auto fail *)
-qed.
-
-lemma lift_append_rcons_sn (A) (k) (f) (p1) (p) (l): k ≗ k →
-      ↑❨λg,p2. k g (p1●l◗p2), f, p❩ = ↑{A}❨λg,p2. k g (p1◖l●p2), f, p❩.
-#A #k #f #p1 #p #l #Hk
-@lift_eq_repl // #g1 #g2 #p2 #Hg
-<list_append_rcons_sn @Hk -Hk // (**) (* auto fail *)
-qed.
-
-(* Advanced constructions with proj_path ************************************)
-
-lemma proj_path_proper:
-      proj_path ≗ proj_path.
-// qed.
-
-lemma lift_path_eq_repl (p):
-      stream_eq_repl … (λf1,f2. ↑[f1]p = ↑[f2]p).
-/2 width=1 by lift_eq_repl/ qed.
-
-lemma lift_path_append_sn (p) (f) (q):
-      q●↑[f]p = ↑❨(λg,p. proj_path g (q●p)), f, p❩.
-#p elim p -p // * [ #n ] #p #IH #f #q
-[ <lift_d_sn <lift_d_sn
-| <lift_m_sn <lift_m_sn
-| <lift_L_sn <lift_L_sn
-| <lift_A_sn <lift_A_sn
-| <lift_S_sn <lift_S_sn
-] 
->lift_lcons_alt // >lift_append_rcons_sn //
-<IH <IH -IH <list_append_rcons_sn //
-qed.
-
-lemma lift_path_lcons (f) (p) (l):
-      l◗↑[f]p = ↑❨(λg,p. proj_path g (l◗p)), f, p❩.
-#f #p #l
->lift_lcons_alt <lift_path_append_sn //
-qed.
-
-lemma lift_path_d_sn (f) (p) (n):
-      (𝗱(f@⧣❨n❩)◗↑[𝐢]p) = ↑[f](𝗱n◗p).
-// qed.
-
-lemma lift_path_m_sn (f) (p):
-      (𝗺◗↑[f]p) = ↑[f](𝗺◗p).
-// qed.
-
-lemma lift_path_L_sn (f) (p):
-      (𝗟◗↑[⫯f]p) = ↑[f](𝗟◗p).
-// qed.
-
-lemma lift_path_A_sn (f) (p):
-      (𝗔◗↑[f]p) = ↑[f](𝗔◗p).
-// qed.
-
-lemma lift_path_S_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 //
-* [ #n1 ] #p1 #IH #f
-[ <lift_path_d_sn <lift_path_d_sn <IH //
-| <lift_path_m_sn <lift_path_m_sn <IH //
-| <lift_path_L_sn <lift_path_L_sn <IH //
-| <lift_path_A_sn <lift_path_A_sn <IH //
-| <lift_path_S_sn <lift_path_S_sn <IH //
-]
-qed.
-
-lemma lift_path_d_dx (n) (p) (f):
-      (↑[f]p)◖𝗱((↑[p]f)@⧣❨n❩) = ↑[f](p◖𝗱n).
-#n #p #f <lift_path_append //
-qed.
-
-lemma lift_path_m_dx (p) (f):
-      (↑[f]p)◖𝗺 = ↑[f](p◖𝗺).
-#p #f <lift_path_append //
-qed.
-
-lemma lift_path_L_dx (p) (f):
-      (↑[f]p)◖𝗟 = ↑[f](p◖𝗟).
-#p #f <lift_path_append //
-qed.
-
-lemma lift_path_A_dx (p) (f):
-      (↑[f]p)◖𝗔 = ↑[f](p◖𝗔).
-#p #f <lift_path_append //
-qed.
-
-lemma lift_path_S_dx (p) (f):
-      (↑[f]p)◖𝗦 = ↑[f](p◖𝗦).
-#p #f <lift_path_append //
-qed.
-
-(* COMMENT 
-
-(* Advanced constructions with proj_rmap and stream_tls *********************)
-
-lemma lift_rmap_tls_d_dx (f) (p) (m) (n):
-      ⇂*[m+n]↑[p]f ≗ ⇂*[m]↑[p◖𝗱n]f.
-#f #p #m #n
-<lift_rmap_d_dx >nrplus_inj_dx
-/2 width=1 by tr_tls_compose_uni_dx/
-qed.
-*)
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_after.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_after.etc
deleted file mode 100644 (file)
index 938d215..0000000
+++ /dev/null
@@ -1,28 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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".
-
-(* LIFT FOR PATH ************************************************************)
-
-(* Constructions with tr_after **********************************************)
-
-lemma lift_path_after (p) (f1) (f2):
-      ↑[f2]↑[f1]p = ↑[f2∘f1]p.
-#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.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_proper.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_proper.etc
deleted file mode 100644 (file)
index 0eef978..0000000
+++ /dev/null
@@ -1,42 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_uni.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_path_uni.etc
deleted file mode 100644 (file)
index 9a62cb2..0000000
+++ /dev/null
@@ -1,24 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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_uni_pap.ma".
-
-(* LIFT FOR PATH ************************************************************)
-
-(* Constructions with tr_uni ************************************************)
-
-lemma lift_path_d_sn_uni (p) (m) (n):
-      (𝗱(n+m)◗p) = ↑[𝐮❨m❩](𝗱(n)◗p).
-// qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm.etc
deleted file mode 100644 (file)
index 6555b5d..0000000
+++ /dev/null
@@ -1,30 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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.ma".
-include "delayed_updating/syntax/prototerm.ma".
-include "ground/lib/subset_ext.ma".
-
-(* LIFT FOR PROTOTERM *******************************************************)
-
-interpretation
-  "lift (prototerm)"
-  'UpArrow f t = (subset_ext_f1 ? ? (lift_gen ? proj_path f) t).
-
-(* Basic constructions ******************************************************)
-
-lemma in_comp_lift_bi (f) (p) (t):
-      p ϵ t → ↑[f]p ϵ ↑[f]t.
-/2 width=1 by subset_in_ext_f1_dx/
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_eq.etc
deleted file mode 100644 (file)
index 0b3ba39..0000000
+++ /dev/null
@@ -1,57 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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_path_after.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.
-
-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.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_proper.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/lift_i/lift_prototerm_proper.etc
deleted file mode 100644 (file)
index a9b1621..0000000
+++ /dev/null
@@ -1,36 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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-.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_constructors.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_constructors.etc
deleted file mode 100644 (file)
index 93ddb2e..0000000
+++ /dev/null
@@ -1,42 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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/unwind2/unwind_prototerm_eq.ma".
-include "delayed_updating/syntax/prototerm_constructors.ma".
-
-(* UNWIND FOR PROTOTERM *****************************************************)
-
-lemma unwind_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/
-qed-.
-
-lemma unwind_iref_after_dx (f) (t) (n:pnat):
-      ▼[f](𝛗n.t) ⊆ ▼[f∘𝐮❨n❩]t.
-#f #t #n #p * #q #Hq #H0 destruct
-elim (in_comp_inv_iref … Hq) -Hq #p #Hp #Ht destruct
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_iref_after (f) (t) (n:pnat):
-      ▼[f∘𝐮❨n❩]t ⇔ ▼[f](𝛗n.t).
-/3 width=1 by conj, unwind_iref_after_sn, unwind_iref_after_dx/
-qed.
-
-lemma unwind_iref (f) (t) (n:pnat):
-      ▼[f]▼[𝐮❨n❩]t ⇔ ▼[f](𝛗n.t).
-/3 width=3 by unwind_term_after, subset_eq_trans/
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_fsubst.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_fsubst.etc
deleted file mode 100644 (file)
index fedc665..0000000
+++ /dev/null
@@ -1,61 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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/unwind2/unwind_prototerm_eq.ma".
-include "delayed_updating/unwind2/unwind_structure.ma".
-include "delayed_updating/substitution/fsubst.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/syntax/prototerm_proper.ma".
-
-(* UNWIND FOR PROTOTERM *****************************************************)
-
-(* Constructions with fsubst ************************************************)
-
-lemma unwind_fsubst_sn (f) (t) (u) (p): u ϵ 𝐏 →
-      (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u] ⊆ ▼[f](t[⋔p←u]).
-#f #t #u #p #Hu #ql * *
-[ #rl * #r #Hr #H1 #H2 destruct
-  >unwind_append_proper_dx
-  /4 width=5 by in_comp_unwind_bi, in_ppc_comp_trans, or_introl, ex2_intro/
-| * #q #Hq #H1 #H0
-  @(ex2_intro … H1) @or_intror @conj // *
-  [ <list_append_empty_dx #H2 destruct
-    elim (unwind_path_root f q) #r #_ #Hr /2 width=2 by/
-  | #l #r #H2 destruct
-    @H0 -H0 [| <unwind_append_proper_dx /2 width=3 by ppc_lcons/ ]
-  ]
-]
-qed-.
-
-lemma unwind_fsubst_dx (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
-      ▼[f](t[⋔p←u]) ⊆ (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u].
-#f #t #u #p #Hu #H1p #H2p #ql * #q * *
-[ #r #Hu #H1 #H2 destruct
-  @or_introl @ex2_intro
-  [|| <unwind_append_proper_dx /2 width=5 by in_ppc_comp_trans/ ]
-  /2 width=3 by ex2_intro/
-| #Hq #H0 #H1 destruct
-  @or_intror @conj [ /2 width=1 by in_comp_unwind_bi/ ] *
-  [ <list_append_empty_dx #Hr @(H0 … (𝐞)) -H0
-    <list_append_empty_dx @H2p -H2p /2 width=1 by prototerm_in_comp_root/
-  | #l #r #Hr
-    elim (unwind_inv_append_proper_dx … Hr) -Hr // #s1 #s2 #Hs1 #_ #H1 destruct
-    lapply (H2p … Hs1) -H2p -Hs1 /2 width=2 by ex_intro/
-  ]
-]
-qed-.
-
-lemma unwind_fsubst (f) (t) (u) (p): u ϵ 𝐏 → p ϵ ▵t → t ϵ 𝐓 →
-      (▼[f]t)[⋔(⊗p)←▼[▼[p]f]u] ⇔ ▼[f](t[⋔p←u]).
-/4 width=3 by unwind_fsubst_sn, conj, unwind_fsubst_dx/ qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_preterm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_preterm_eq.etc
deleted file mode 100644 (file)
index d60c54c..0000000
+++ /dev/null
@@ -1,68 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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_equivalence.ma".
-include "delayed_updating/syntax/path_structure_inner.ma".
-include "delayed_updating/syntax/preterm.ma".
-include "delayed_updating/unwind2/unwind_structure.ma".
-include "delayed_updating/unwind2/unwind_prototerm.ma".
-
-(* UNWIND FOR PRETERM *******************************************************)
-
-(* Constructions with subset_equivalence ************************************)
-
-lemma unwind_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
-<unwind_append_inner_sn //
-qed-.
-
-lemma unwind_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 (unwind_inv_append_inner_sn … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-/2 width=1 by in_comp_unwind_bi/
-qed-.
-
-lemma unwind_grafted (f) (t) (p): p ϵ 𝐈 → p ϵ ▵t → t ϵ 𝐓 →
-      ▼[▼[p]f](t⋔p) ⇔ (▼[f]t)⋔(⊗p).
-/3 width=1 by unwind_grafted_sn, conj, unwind_grafted_dx/ qed.
-
-
-lemma unwind_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 (unwind_inv_append_proper_dx … (sym_eq … H0)) -H0 //
-#p0 #q0 #Hp0 #Hq0 #H0 destruct
-<(Ht … Hp0) [|*: /2 width=2 by ex_intro/ ] -p
-elim (unwind_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_unwind_bi/
-qed-.
-
-lemma unwind_grafted_S (f) (t) (p): p ϵ ▵t → t ϵ 𝐓 →
-      ▼[▼[p]f](t⋔(p◖𝗦)) ⇔ (▼[f]t)⋔((⊗p)◖𝗦).
-#f #t #p #Hp #Ht
-@conj
-[ >unwind_rmap_S_dx >structure_S_dx
-  @unwind_grafted_sn //
-| /2 width=1 by unwind_grafted_S_dx/
-]
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm.etc
deleted file mode 100644 (file)
index 1a5d46b..0000000
+++ /dev/null
@@ -1,30 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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/unwind2/unwind.ma".
-include "delayed_updating/syntax/prototerm.ma".
-include "ground/lib/subset_ext.ma".
-
-(* UNWIND FOR PROTOTERM *****************************************************)
-
-interpretation
-  "unwind (prototerm)"
-  'BlackDownTriangle f t = (subset_ext_f1 ? ? (unwind_gen ? proj_path f) t).
-
-(* Basic constructions ******************************************************)
-
-lemma in_comp_unwind_bi (f) (p) (t):
-      p ϵ t → ▼[f]p ϵ ▼[f]t.
-/2 width=1 by subset_in_ext_f1_dx/
-qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm_eq.etc b/matita/matita/contribs/lambdadelta/delayed_updating/etc/unwind2/unwind_prototerm_eq.etc
deleted file mode 100644 (file)
index 15b8cdb..0000000
+++ /dev/null
@@ -1,40 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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/unwind2/unwind_eq.ma".
-include "delayed_updating/unwind2/unwind_prototerm.ma".
-
-(* UNWIND FOR PROTOTERM *****************************************************)
-
-(* Constructions with subset_equivalence ************************************)
-
-lemma unwind_term_eq_repl_sn (f1) (f2) (t):
-      f1 ≗ f2 → ▼[f1]t ⇔ ▼[f2]t.
-/3 width=1 by subset_equivalence_ext_f1_exteq, unwind_path_eq_repl/
-qed.
-
-lemma unwind_term_eq_repl_dx (f) (t1) (t2):
-      t1 ⇔ t2 → ▼[f]t1 ⇔ ▼[f]t2.
-/2 width=1 by subset_equivalence_ext_f1_bi/
-qed.
-
-lemma unwind_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.
index 0a4b7cf1f8b56238f652e21fd943e6a9e3a85d8d..738553c7de027e4165309cd8bebc1668aa9c07dc 100644 (file)
@@ -14,7 +14,7 @@
 
 include "delayed_updating/substitution/fsubst.ma".
 include "delayed_updating/syntax/prototerm_constructors.ma".
-include "delayed_updating/syntax/prototerm_equivalence.ma".
+include "delayed_updating/syntax/prototerm_eq.ma".
 include "delayed_updating/syntax/path_head.ma".
 include "delayed_updating/syntax/path_depth.ma".
 include "delayed_updating/syntax/path_reverse.ma".
index e33b454bed039864ef99debb5b0fe171d4397ee3..3c6e4d6d9a205f19501869d3ac612b76d52de34c 100644 (file)
@@ -42,7 +42,7 @@ lemma pippo (f) (r):
 *)
 
 theorem dfr_lift_bi (f) (p) (q) (t1) (t2): (*t1 ϵ 𝐓 → *)
-        t1 ➡𝐝𝐟[p,q] t2 → ↑[f]t1 ➡𝐟[↑[f]p,↑[↑[p◖𝗔◖𝗟]f]q] ↑[f]t2.
+        t1 â\9e¡ð\9d\90\9dð\9d\90\9f[p,q] t2 â\86\92 â\86\91[f]t1 â\9e¡ð\9d\90\9dð\9d\90\9f\86\91[f]p,â\86\91\86\91[pâ\97\96ð\9d\97\94â\97\96ð\9d\97\9f]f]q] â\86\91[f]t2.
 #f #p #q #t1 #t2
 * #n * #H1n #Ht1 #Ht2
 @(ex_intro … ((↑[p●𝗔◗𝗟◗q]f)@⧣❨n❩)) @and3_intro
@@ -50,15 +50,11 @@ theorem dfr_lift_bi (f) (p) (q) (t1) (t2): (*t1 ϵ 𝐓 → *)
   <lift_rmap_L_dx >lift_path_L_sn
   >list_append_rcons_sn in H1n; <reverse_append #H1n
   <(lift_path_head … 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
+| lapply (in_comp_lift_path_term f … Ht1) -Ht2 -Ht1 -H1n
+  <lift_path_d_dx #Ht1 //
+| lapply (lift_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 …))
index 87d855281d53c7a51c69608eee0700a69100ecff..a63368e1eb644b277899865c0a6ca8f43ca868a6 100644 (file)
@@ -14,7 +14,7 @@
 
 include "delayed_updating/unwind/unwind2_prototerm.ma".
 include "delayed_updating/substitution/fsubst.ma".
-include "delayed_updating/syntax/prototerm_equivalence.ma".
+include "delayed_updating/syntax/prototerm_eq.ma".
 include "delayed_updating/syntax/path_head.ma".
 include "delayed_updating/syntax/path_reverse.ma".
 include "delayed_updating/notation/relations/black_rightarrow_if_4.ma".
index 035a647a8a8794deca6b1852e33175c5e26b8805..a3a3494191ad176fd4a62cea1cb261669c1fc2a4 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "delayed_updating/substitution/fsubst.ma".
-include "delayed_updating/syntax/prototerm_equivalence.ma".
+include "delayed_updating/syntax/prototerm_eq.ma".
 
 (* Constructions with subset_equivalence ************************************)
 
index 1cdcd4650d1b990de616a87dd79a2f8e202b818d..2fae75c26603af11f644ca666774478249c613bd 100644 (file)
 
 include "delayed_updating/substitution/lift_prototerm_id.ma".
 include "delayed_updating/substitution/lift_path_uni.ma".
-include "delayed_updating/syntax/prototerm_constructors.ma".
+include "delayed_updating/syntax/prototerm_constructors_eq.ma".
 
 (* LIFT FOR PROTOTERM *******************************************************)
 
-lemma lift_iref_bi (t1) (t2) (n):
-      t1 ⇔ t2 → 𝛗n.t1 ⇔ 𝛗n.t2.
-/2 width=1 by subset_equivalence_ext_f1_bi/
-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
@@ -34,7 +29,7 @@ lemma lift_iref_dx (f) (t) (n:pnat):
       ↑[f](𝛗n.t) ⊆ 𝛗f@⧣❨n❩.↑[⇂*[n]f]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/
+/3 width=1 by in_comp_iref, in_comp_lift_path_term/
 qed-.
 
 lemma lift_iref (f) (t) (n:pnat):
@@ -47,5 +42,5 @@ lemma lift_iref_uni (t) (m) (n):
 #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/
+/3 width=1 by iref_eq_repl, lift_term_id/
 qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_id.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_id.ma
new file mode 100644 (file)
index 0000000..c2376b5
--- /dev/null
@@ -0,0 +1,38 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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 proj_path and 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.
+
+(* Constructions with proj_rmap and tr_id ***********************************)
+
+lemma lift_rmap_id (p):
+      (𝐢) = ↑[p]𝐢.
+#p elim p -p //
+* [ #n ] #p #IH //
+qed.
index 126b691c6ff341101e2ea4a00a871d619dcecfb1..71b11f5d6d4807c35414dcc16ecc3f9cf9b0a699 100644 (file)
@@ -1,16 +1,25 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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_head.ma".
 include "delayed_updating/syntax/path_reverse.ma".
 include "ground/relocation/xap.ma".
 
-axiom tr_xap_succ_pos (f) (n):
-      ↑↓(f@❨↑n❩) = f@❨↑n❩.
-
-axiom tr_xap_plus (n1) (n2) (f):
-      (⇂*[n2]f)@❨n1❩+f@❨n2❩ = f@❨n1+n2❩.
+(* LIFT FOR PATH ************************************************************)
 
-axiom eq_inv_path_empty_head (p) (n):
-      (𝐞) = ↳[n]p → 𝟎 = n.
+(* Constructions with head for path *****************************************)
 
 lemma lift_path_head (f) (p) (q) (n):
       pᴿ = ↳[n](pᴿ●qᴿ) →
@@ -25,7 +34,34 @@ lemma lift_path_head (f) (p) (q) (n):
     [ <reverse_rcons <path_head_d_sn #H0
       elim (eq_inv_list_lcons_bi ????? H0) -H0 #_ #H0
       <list_append_assoc <lift_rmap_d_dx <lift_path_d_dx <reverse_rcons
-      <tr_xap_succ_pos <path_head_d_sn >tr_xap_succ_pos
-      <lift_path_d_dx >lift_rmap_append <reverse_rcons  
-      @eq_f2 // <(IH … H0) -IH -H0
-      @eq_f2 // <tr_xap_plus //
+      <tr_xap_succ_nap <path_head_d_sn >tr_xap_succ_nap
+      <lift_path_d_dx >lift_rmap_append <reverse_rcons
+      <(IH … H0) -IH -H0 <tr_xap_plus //
+    | <reverse_rcons <path_head_m_sn #H0
+      elim (eq_inv_list_lcons_bi ????? H0) -H0 #_ #H0
+      <list_append_assoc <lift_rmap_m_dx <lift_path_m_dx <reverse_rcons
+      <tr_xap_succ_nap <path_head_m_sn >tr_xap_succ_nap
+      <lift_path_m_dx <reverse_rcons
+      <(IH … H0) -IH -H0 //
+    | <reverse_rcons <path_head_L_sn #H0
+      elim (eq_inv_list_lcons_bi ????? H0) -H0 #_ #H0
+      <list_append_assoc <lift_rmap_L_dx <lift_path_L_dx <reverse_rcons
+      <tr_xap_succ_nap <path_head_L_sn >tr_xap_succ_nap
+      <lift_path_L_dx <reverse_rcons
+      <(IH … H0) -IH -H0 //
+    | <reverse_rcons <path_head_A_sn #H0
+      elim (eq_inv_list_lcons_bi ????? H0) -H0 #_ #H0
+      <list_append_assoc <lift_rmap_A_dx <lift_path_A_dx <reverse_rcons
+      <tr_xap_succ_nap <path_head_A_sn >tr_xap_succ_nap
+      <lift_path_A_dx <reverse_rcons
+      <(IH … H0) -IH -H0 //
+    | <reverse_rcons <path_head_S_sn #H0
+      elim (eq_inv_list_lcons_bi ????? H0) -H0 #_ #H0
+      <list_append_assoc <lift_rmap_S_dx <lift_path_S_dx <reverse_rcons
+      <tr_xap_succ_nap <path_head_S_sn >tr_xap_succ_nap
+      <lift_path_S_dx <reverse_rcons
+      <(IH … H0) -IH -H0 //
+    ]
+  ]
+]
+qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_id.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_path_id.ma
deleted file mode 100644 (file)
index 1966c25..0000000
+++ /dev/null
@@ -1,30 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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.
index f68010b32a4eabfcd6a6d72f99b42988f3ebe574..0e22a3d39917cc97389097f4015fe4b58c85c8ca 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "delayed_updating/substitution/lift_path_id.ma".
+include "delayed_updating/substitution/lift_id.ma".
 include "ground/relocation/tr_uni_pap.ma".
 include "ground/relocation/tr_uni_tls.ma".
 
index 6555b5d9ed0ec36e6d6d1605a92cfe591e740a63..ed5fbe84c57a3c4f41e7966b018d92d0f22882c9 100644 (file)
@@ -24,7 +24,7 @@ interpretation
 
 (* Basic constructions ******************************************************)
 
-lemma in_comp_lift_bi (f) (p) (t):
+lemma in_comp_lift_path_term (f) (p) (t):
       p ϵ t → ↑[f]p ϵ ↑[f]t.
 /2 width=1 by subset_in_ext_f1_dx/
 qed.
index b66b63f46acc907bfaa19432af8bafc2743ca93e..7343319f42e7ff5241817a329331e17776385d9e 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "delayed_updating/substitution/lift_prototerm_eq.ma".
-include "delayed_updating/substitution/lift_path_id.ma".
+include "delayed_updating/substitution/lift_id.ma".
 
 (* LIFT FOR PROTOTERM *******************************************************)
 
@@ -23,7 +23,7 @@ lemma lift_term_id_sn (t):
       t ⊆ ↑[𝐢]t.
 #t #p #Hp
 >(lift_path_id p)
-/2 width=1 by in_comp_lift_bi/
+/2 width=1 by in_comp_lift_path_term/
 qed-.
 
 lemma lift_term_id_dx (t):
index a9b16219f04c4988c46b2af043c604e5905fba26..45d65b3ff9c76d975c552913417d9a2c15273518 100644 (file)
@@ -32,5 +32,5 @@ 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 *)
+@Ht -Ht @in_comp_lift_path_term // (**) (* auto fails *)
 qed-.
index 8aaefad914ec00ca71ebd9468813d62635283e94..6a70f61cb0284030cf8349ffce95619bc520b721 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "delayed_updating/syntax/prototerm_constructors.ma".
-include "delayed_updating/syntax/prototerm_equivalence.ma".
+include "delayed_updating/syntax/prototerm_eq.ma".
 include "delayed_updating/notation/functions/class_d_phi_0.ma".
 include "ground/xoa/or_5.ma".
 include "ground/xoa/ex_3_1.ma".
index 827af44891149cd6c5ec3d6bddbf3e4fc04cbb74..9cdfe620ff67088da01fa37b512b41f1fa3e9e65 100644 (file)
@@ -71,3 +71,20 @@ lemma path_head_A_sn (p) (n):
 lemma path_head_S_sn (p) (n):
       (𝗦◗↳[↑n]p) = ↳[↑n](𝗦◗p).
 // qed.
+
+(* Basic inversions *********************************************************)
+
+lemma eq_inv_path_empty_head (p) (n):
+      (𝐞) = ↳[n]p → 𝟎 = n.
+*
+[ #m <path_head_empty #H0
+  <(eq_inv_empty_labels … H0) -m //
+| * [ #n ] #p #n @(nat_ind_succ … n) -n // #m #_
+  [ <path_head_d_sn
+  | <path_head_m_sn
+  | <path_head_L_sn
+  | <path_head_A_sn
+  | <path_head_S_sn
+  ] #H0 destruct
+]
+qed-.
index df218ebc2f035b9211edf7a199fb64ee1914dcd9..434f1e810cb44fe31d6bc2764bfed79e637bb6fd 100644 (file)
@@ -61,7 +61,7 @@ lemma in_comp_iref (t) (q) (n):
       q ϵ t → 𝗱n◗𝗺◗q ϵ 𝛗n.t.
 /2 width=3 by ex2_intro/ qed.
 
-(* Basic Inversions *********************************************************)
+(* Basic inversions *********************************************************)
 
 lemma in_comp_inv_iref (t) (p) (n):
       p ϵ 𝛗n.t →
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors_eq.ma b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors_eq.ma
new file mode 100644 (file)
index 0000000..88f5672
--- /dev/null
@@ -0,0 +1,25 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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_constructors.ma".
+include "ground/lib/subset_ext_equivalence.ma".
+
+(* CONSTRUCTORS FOR PROTOTERM ***********************************************)
+
+(* Constructions with equivalence for prototerm *****************************)
+
+lemma iref_eq_repl (t1) (t2) (n):
+      t1 ⇔ t2 → 𝛗n.t1 ⇔ 𝛗n.t2.
+/2 width=1 by subset_equivalence_ext_f1_bi/
+qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_eq.ma b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_eq.ma
new file mode 100644 (file)
index 0000000..baffe07
--- /dev/null
@@ -0,0 +1,32 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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.ma".
+include "ground/lib/subset_equivalence.ma".
+
+(* EQUIVALENCE FOR PROTOTERM ************************************************)
+
+(* Constructions with prototerm_root ****************************************)
+
+lemma prototerm_root_incl_repl:
+      ∀t1,t2. t1 ⊆ t2 → ▵t1 ⊆ ▵t2.
+#t1 #t2 #Ht #p * #q #Hq
+/3 width=2 by ex_intro/
+qed.
+
+lemma prototerm_root_eq_repl:
+      ∀t1,t2. t1 ⇔ t2 → ▵t1 ⇔ ▵t2.
+#t1 #t2 * #H1 #H2
+/3 width=3 by conj, prototerm_root_incl_repl/
+qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_equivalence.ma b/matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_equivalence.ma
deleted file mode 100644 (file)
index 414c47d..0000000
+++ /dev/null
@@ -1,32 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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_equivalence.ma".
-include "delayed_updating/syntax/prototerm.ma".
-
-(* EQUIVALENCE FOR PROTOTERM ************************************************)
-
-(* Constructions with prototerm_root ****************************************)
-
-lemma prototerm_root_incl_repl:
-      ∀t1,t2. t1 ⊆ t2 → ▵t1 ⊆ ▵t2.
-#t1 #t2 #Ht #p * #q #Hq
-/3 width=2 by ex_intro/
-qed.
-
-lemma prototerm_root_eq_repl:
-      ∀t1,t2. t1 ⇔ t2 → ▵t1 ⇔ ▵t2.
-#t1 #t2 * #H1 #H2
-/3 width=3 by conj, prototerm_root_incl_repl/
-qed.
index 985b3ae26e292b0be8b77d0db3f7fa7d555e4010..ac32304ed352d6fd996c64900ead586effd36fe2 100644 (file)
@@ -22,6 +22,10 @@ lemma tr_nap_unfold (f) (l):
       ↓(f@⧣❨↑l❩) = f@↑❨l❩.
 // qed.
 
+lemma tr_pap_succ_nap (f) (l):
+      ↑(f@↑❨l❩) = f@⧣❨↑l❩.
+// qed.
+
 lemma tr_compose_nap (f2) (f1) (l):
       f2@↑❨f1@↑❨l❩❩ = (f2∘f1)@↑❨l❩.
 #f2 #f1 #l
index 18bcec6a998e7d2b297258c3bd47bc2ff473464a..5e607ac373e56b79c44d2d545da8168b645504b2 100644 (file)
@@ -1,6 +1,8 @@
+include "ground/arith/nat_rplus_pplus.ma".
 include "ground/relocation/nap.ma".
 include "ground/notation/functions/apply_2.ma".
 include "ground/relocation/tr_compose_pn.ma".
+include "ground/relocation/tr_pap_tls.ma".
 
 definition tr_xap (f) (l:nat): nat ≝
            (⫯f)@↑❨l❩.
@@ -21,6 +23,12 @@ lemma tr_xap_ninj (f) (p):
       ninj (f@⧣❨p❩) = f@❨ninj p❩.
 // qed.
 
+lemma tr_xap_succ_nap (f) (n):
+      ↑(f@↑❨n❩) = f@❨↑n❩.
+#f #n
+<tr_xap_ninj //
+qed.
+
 lemma tr_compose_xap (f2) (f1) (l):
       f2@❨f1@❨l❩❩ = (f2∘f1)@❨l❩.
 #f2 #f1 #l
@@ -53,3 +61,11 @@ lemma tr_xap_pushs_le (f) (n) (m):
 <tr_xap_unfold >tr_pushs_succ <tr_nap_pushs_lt //
 /2 width=1 by nlt_succ_dx/
 qed-.
+
+lemma tr_xap_plus (n1) (n2) (f):
+      (⇂*[n2]f)@❨n1❩+f@❨n2❩ = f@❨n1+n2❩.
+* [| #n1 ] // * [| #n2 ] // #f
+<nrplus_inj_sn <nrplus_inj_dx
+<nrplus_inj_sn <nrplus_inj_dx
+>tr_pap_plus //
+qed.