]> matita.cs.unibo.it Git - helm.git/commitdiff
update in delayed_updating
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 24 Jan 2022 23:30:11 +0000 (00:30 +0100)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 24 Jan 2022 23:30:11 +0000 (00:30 +0100)
+ bugs fixed in depth, dfr, ifr
+ WIP on dfr_lift_bi

matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr.ma
matita/matita/contribs/lambdadelta/delayed_updating/reduction/dfr_ifr.ma
matita/matita/contribs/lambdadelta/delayed_updating/reduction/ifr.ma
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_constructors.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_depth.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_structure_depth.ma
matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_depth.ma
matita/matita/contribs/lambdadelta/delayed_updating/syntax/prototerm_constructors.ma

index 3cdf0d47e9f63a9c1666d6a23b2c036c293dc2ac..e118134dd450a7bfcc2a49c6865b250a62c90d05 100644 (file)
@@ -20,6 +20,7 @@ include "delayed_updating/syntax/path_structure.ma".
 include "delayed_updating/syntax/path_balanced.ma".
 include "delayed_updating/syntax/path_depth.ma".
 include "delayed_updating/notation/relations/black_rightarrow_df_4.ma".
+include "ground/arith/nat_rplus.ma".
 include "ground/xoa/ex_1_2.ma".
 include "ground/xoa/and_4.ma".
 
@@ -28,8 +29,8 @@ include "ground/xoa/and_4.ma".
 definition dfr (p) (q): relation2 prototerm prototerm ≝
            λt1,t2. ∃∃b,n.
            let r ≝ p●𝗔◗b●𝗟◗q in
-           â\88§â\88§ â\8a\97b Ïµ ð\9d\90\81 & â\88\80f. â\9d\98\9d\98 = (â\86\91[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
-              t1[⋔r←𝛗n.(t1⋔(p◖𝗦))] ⇔ t2
+           â\88§â\88§ â\8a\97b Ïµ ð\9d\90\81 & â\88\80f. â\86\91â\9d\98\9d\98 = (â\86\91[q]f)@❨n❩ & r◖𝗱n ϵ t1 &
+              t1[⋔r←𝛗(n+❘b❘).(t1⋔(p◖𝗦))] ⇔ t2
 .
 
 interpretation
index a8d9469ddaa6fe171c11ece64ebff473d74a3429..c555072f6558176f3d9d046ffb611e37fcf778c0 100644 (file)
@@ -26,7 +26,7 @@ lemma dfr_lift_bi (f) (p) (q) (t1) (t2): t1 ϵ 𝐓 →
       t1 ➡𝐝𝐟[p,q] t2 → ↑[f]t1 ➡𝐟[⊗p,⊗q] ↑[f]t2.
 #f #p #q #t1 #t2 #H0t1
 * #b #n * #Hb #Hn  #Ht1 #Ht2
-@(ex1_2_intro … (⊗b) (❘⊗q❘)) @and4_intro
+@(ex1_2_intro â\80¦ (â\8a\97b) (â\86\91â\9d\98â\8a\97\9d\98)) @and4_intro
 [ //
 | #g <lift_rmap_structure <depth_structure
   >tr_pushs_swap <tr_pap_pushs_le //
index ffa4f44bd0ea968f66f6aaf46653295c126ef62c..b188a7730dc1b10ae5fd10b664390d98bb3963ad 100644 (file)
@@ -26,8 +26,8 @@ include "ground/xoa/and_4.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 &
-              t1[⋔r←↑[𝐮❨n❩](t1⋔(p◖𝗦))] ⇔ t2
+           ∧∧ ⊗b ϵ 𝐁 & ∀f. ❘𝗟◗q❘ = (↑[q]⫯f)@❨n❩ & r◖𝗱n ϵ t1 &
+              t1[⋔r←↑[𝐮❨❘b●𝗟◗q❘❩](t1⋔(p◖𝗦))] ⇔ t2
 .
 
 interpretation
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_constructors.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_constructors.ma
new file mode 100644 (file)
index 0000000..0202dc9
--- /dev/null
@@ -0,0 +1,56 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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".
+
+axiom lift_path_after (p) (f1) (f2):
+      ↑[f2]↑[f1]p = ↑[f2∘f1]p.
+(*
+#p @(path_ind_lift … p) -p // [ #n | #n #l #p | #p ] #IH #f1 #f2
+[ <lift_path_d_empty_sn <lift_path_d_empty_sn
+| <lift_path_d_lcons_sn <lift_path_d_lcons_sn
+| <lift_path_L_sn <lift_path_L_sn <lift_path_L_sn
+*)
+include "delayed_updating/substitution/lift_prototerm.ma".
+
+axiom lift_term_after (t) (f1) (f2):
+      ↑[f2]↑[f1]t ⇔ ↑[f2∘f1]t.
+
+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).
+#f #t #n #p * #q #Hq #H0 destruct
+@(ex2_intro … (𝗱n◗𝗺◗q))
+/2 width=1 by in_comp_iref/
+qed-.
+
+lemma lift_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_lift_bi/
+qed-.
+
+lemma lift_iref_after (f) (t) (n:pnat):
+      ↑[f∘𝐮❨n❩]t ⇔ ↑[f](𝛗n.t).
+/3 width=1 by conj, lift_iref_after_sn, lift_iref_after_dx/
+qed.
+
+lemma lift_iref (f) (t) (n:pnat):
+      ↑[f]↑[𝐮❨n❩]t ⇔ ↑[f](𝛗n.t).
+/3 width=3 by lift_term_after, subset_eq_trans/
+qed.
diff --git a/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_depth.ma b/matita/matita/contribs/lambdadelta/delayed_updating/substitution/lift_depth.ma
new file mode 100644 (file)
index 0000000..b88464a
--- /dev/null
@@ -0,0 +1,30 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/path_depth.ma".
+include "ground/relocation/tr_pushs.ma".
+
+(* LIFT FOR PATH ***********************************************************)
+
+(* Basic constructions with depth ******************************************)
+
+lemma pippo (p) (f):
+      (⫯*[❘p❘]f) ∘ (↑[p]𝐢) = ↑[p]f.
+#p elim p -p
+[ #f <lift_rmap_empty <lift_rmap_empty <tr_pushs_zero
+| * [ #n ] #p #IH #f //
+  <lift_rmap_d_sn <lift_rmap_d_sn <depth_d
+  @(trans_eq … (IH …)) -IH
+      
index ee42bfb2437ba841ce544d243d5596a471a44fb4..228e132bfb70216cc6f18f2bbce74d578a5707a4 100644 (file)
@@ -23,11 +23,10 @@ include "ground/arith/nat_pred_succ.ma".
 (* Basic constructions with structure and depth ****************************)
 
 lemma lift_rmap_structure (p) (f):
-      (⫯*[â\86\93â\9d\98\9d\98]f) = â\86\91\8a\97p]f.
+      (⫯*[❘p❘]f) = ↑[⊗p]f.
 #p elim p -p //
 * [ #n ] #p #IH #f //
-[ <lift_rmap_L_sn <IH -IH >tr_pushs_swap //
-| <lift_rmap_A_sn //
+[ <lift_rmap_A_sn //
 | <lift_rmap_S_sn //
 ]
 qed.
index 88952c2bc0b4b4edcd60a53c291f1c85a7a4866c..1634678af36ecc77d57b90df7e0b6dddeb4fbab0 100644 (file)
 (**************************************************************************)
 
 include "delayed_updating/syntax/path.ma".
+include "ground/arith/nat_succ.ma".
 include "ground/notation/functions/verticalbars_1.ma".
 
 (* DEPTH FOR PATH ***********************************************************)
 
-rec definition depth (p) on p: pnat ≝
+rec definition depth (p) on p: nat ≝
 match p with
-[ list_empty     â\87\92 ð\9d\9f\8f
+[ list_empty     â\87\92 ð\9d\9f\8e
 | list_lcons l q ⇒
   match l with
   [ label_d _ ⇒ depth q
@@ -36,7 +37,7 @@ interpretation
 
 (* Basic constructions ******************************************************)
 
-lemma depth_empty: ð\9d\9f\8f = ❘𝐞❘.
+lemma depth_empty: ð\9d\9f\8e = ❘𝐞❘.
 // qed.
 
 lemma depth_d (q) (n): ❘q❘ = ❘𝗱n◗q❘.
index fd86ef975a39b862aaf05ac99dc5efd298fb10ed..35a79b0399624d79444c8ba3243003e8b1eb5cb6 100644 (file)
@@ -55,8 +55,21 @@ interpretation
   "application (prototerm)"
   'At u t = (prototerm_node_2 label_S label_A u t).
 
+(* Basic constructions *******************************************************)
+
+lemma in_comp_iref (t) (q) (n):
+      q ϵ t → 𝗱n◗𝗺◗q ϵ 𝛗n.t.
+/2 width=3 by ex2_intro/ qed.
+
 (* Basic Inversions *********************************************************)
 
+lemma in_comp_inv_iref (t) (p) (n):
+      p ϵ 𝛗n.t →
+      ∃∃q. 𝗱n◗𝗺◗q = p & q ϵ t.
+#t #p #n * #q #Hq #Hp
+/2 width=3 by ex2_intro/
+qed-.
+(*
 lemma prototerm_in_root_inv_lcons_oref:
       ∀p,l,n. l◗p ϵ ▵#n →
       ∧∧ 𝗱n = l & 𝐞 = p.
@@ -98,3 +111,4 @@ lemma prototerm_in_root_inv_lcons_appl:
 <list_append_lcons_sn #H0 destruct
 /4 width=2 by ex_intro, or_introl, or_intror, conj/
 qed-.
+*)