]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_depth.ma
update in delayed_updating
[helm.git] / matita / matita / contribs / lambdadelta / delayed_updating / syntax / path_depth.ma
index e76c4943b704ce2729085a27fe150c8decb87448..cdf24943af922c8a47c74990fb79e92acf803231 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "delayed_updating/syntax/path.ma".
-include "ground/arith/nat_succ.ma".
+include "ground/arith/nat_plus.ma".
 include "ground/notation/functions/verticalbars_1.ma".
 
 (* DEPTH FOR PATH ***********************************************************)
@@ -23,10 +23,11 @@ match p with
 [ list_empty     ā‡’ šŸŽ
 | list_lcons l q ā‡’
   match l with
-  [ label_node_d _ ā‡’ depth q
-  | label_edge_L   ā‡’ ā†‘(depth q)
-  | label_edge_A   ā‡’ depth q
-  | label_edge_S   ā‡’ depth q
+  [ label_d _ ā‡’ depth q
+  | label_m   ā‡’ depth q
+  | label_L   ā‡’ ā†‘(depth q)
+  | label_A   ā‡’ depth q
+  | label_S   ā‡’ depth q
   ]
 ].
 
@@ -39,14 +40,53 @@ interpretation
 lemma depth_empty: šŸŽ = ā˜šžā˜.
 // qed.
 
-lemma depth_d (q) (n): ā˜qā˜ = ā˜š—±āØnā©ā——qā˜.
+lemma depth_d_sn (q) (n): ā˜qā˜ = ā˜š—±nā——qā˜.
 // qed.
 
-lemma depth_L (q): ā†‘ā˜qā˜ = ā˜š—Ÿā——qā˜.
+lemma depth_m_sn (q): ā˜qā˜ = ā˜š—ŗā——qā˜.
 // qed.
 
-lemma depth_A (q): ā˜qā˜ = ā˜š—”ā——qā˜.
+lemma depth_L_sn (q): ā†‘ā˜qā˜ = ā˜š—Ÿā——qā˜.
 // qed.
 
-lemma depth_S (q): ā˜qā˜ = ā˜š—¦ā——qā˜.
+lemma depth_A_sn (q): ā˜qā˜ = ā˜š—”ā——qā˜.
+// qed.
+
+lemma depth_S_sn (q): ā˜qā˜ = ā˜š—¦ā——qā˜.
+// qed.
+
+(* Main constructions *******************************************************)
+
+theorem depth_append (p1) (p2):
+        ā˜p2ā˜+ā˜p1ā˜ = ā˜p1ā—p2ā˜.
+#p1 elim p1 -p1 //
+* [ #n ] #p1 #IH #p2 <list_append_lcons_sn
+[ <depth_d_sn <depth_d_sn //
+| <depth_m_sn <depth_m_sn //
+| <depth_L_sn <depth_L_sn //
+| <depth_A_sn <depth_A_sn //
+| <depth_S_sn <depth_S_sn //
+]
+qed.
+
+(* Constructions with list_rcons ********************************************)
+
+lemma depth_d_dx (p) (n):
+      ā˜pā˜ = ā˜pā—–š—±nā˜.
+// qed.
+
+lemma depth_m_dx (p):
+      ā˜pā˜ = ā˜pā—–š—ŗā˜.
+// qed.
+
+lemma depth_L_dx (p):
+      ā†‘ā˜pā˜ = ā˜pā—–š—Ÿā˜.
+// qed.
+
+lemma depth_A_dx (p):
+      ā˜pā˜ = ā˜pā—–š—”ā˜.
+// qed.
+
+lemma depth_S_dx (p):
+      ā˜pā˜ = ā˜pā—–š—¦ā˜.
 // qed.