]> matita.cs.unibo.it Git - helm.git/commitdiff
update in ground
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 26 Mar 2022 16:56:36 +0000 (17:56 +0100)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Sat, 26 Mar 2022 16:56:36 +0000 (17:56 +0100)
+ some additions and corrections

matita/matita/contribs/lambdadelta/ground/relocation/tr_id.ma
matita/matita/contribs/lambdadelta/ground/relocation/tr_id_hdtl.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/ground/relocation/tr_id_tls.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/ground/relocation/tr_uni_hdtl.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/ground/relocation/tr_uni_tls.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/ground/web/ground_src.tbl

index 9b421ee43390951148ccf3ac0c781656c6ead232..4360fcb116e24850c030e9e6535e97eb0eb60a3b 100644 (file)
@@ -23,7 +23,7 @@ interpretation
   "identity element (total relocation streams)"
   'ElementI = (tr_id).
 
-(* Basic constructions (specific) *******************************************)
+(* Basic constructions ******************************************************)
 
 lemma tr_id_unfold: ā«Æš¢ = š¢.
 <(stream_unfold ā€¦ (š¢)) in āŠ¢ (???%); //
diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_id_hdtl.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_id_hdtl.ma
new file mode 100644 (file)
index 0000000..57f5b4a
--- /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 "ground/relocation/tr_pn_hdtl.ma".
+include "ground/relocation/tr_id.ma".
+
+(* IDENTITY ELEMENT FOR TOTAL RELOCATION MAPS *******************************)
+
+(* Constructions with stream_hd *********************************************)
+
+lemma tr_hd_id:
+      (šŸ) = ā‡ƒš¢.
+// qed.
+
+(* Constructions with stream_tl *********************************************)
+
+lemma tr_tl_id:
+      (š¢) = ā‡‚š¢.
+// qed.
diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_id_tls.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_id_tls.ma
new file mode 100644 (file)
index 0000000..20906fe
--- /dev/null
@@ -0,0 +1,26 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/stream_tls.ma".
+include "ground/relocation/tr_id_hdtl.ma".
+
+(* IDENTITY ELEMENT FOR TOTAL RELOCATION MAPS *******************************)
+
+(* Constructions with stream_tls ********************************************)
+
+lemma tr_tls_id (n):
+      (š¢) = ā‡‚*[n]š¢.
+#n @(nat_ind_succ ā€¦ n) -n //
+#n #IH <stream_tls_swap //
+qed.
diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_uni_hdtl.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_uni_hdtl.ma
new file mode 100644 (file)
index 0000000..2ad69e3
--- /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 "ground/lib/stream_hdtl.ma".
+include "ground/relocation/tr_uni.ma".
+
+(* UNIFORM ELEMENTS FOR TOTAL RELOCATION MAPS *******************************)
+
+(* Constructions with stream_hd *********************************************)
+
+lemma tr_hd_uni (n):
+      ā†‘n = ā‡ƒš®āØnā©.
+// qed.
+
+(* Constructions with stream_tl *********************************************)
+
+lemma tr_tl_uni (n):
+      (š¢) = ā‡‚š®āØnā©.
+// qed.
diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_uni_tls.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_uni_tls.ma
new file mode 100644 (file)
index 0000000..f022eb8
--- /dev/null
@@ -0,0 +1,24 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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/relocation/tr_id_tls.ma".
+include "ground/relocation/tr_uni_hdtl.ma".
+
+(* UNIFORM ELEMENTS FOR TOTAL RELOCATION MAPS *******************************)
+
+(* Constructions with stream_tls ********************************************)
+
+lemma tr_tls_succ_uni (m) (n):
+      (š¢) = ā‡‚*[ā†‘m]š®āØnā©.
+// qed.
index 333416ec298927d746a8cfc40249550e3ace83ef..9ee7ad91a158eacbe6edf238c9ffeb3c7af1ef40 100644 (file)
@@ -30,8 +30,8 @@ table {
         }
       ]
       [ { "total relocation" * } {
-          [ "tr_uni ( š®āØ?ā© )" "tr_uni_eq" "tr_uni_pap" "tr_uni_compose" * ]
-          [ "tr_id ( š¢ ) " "tr_id_pushs" "tr_id_pap" "tr_id_compose" * ]
+          [ "tr_uni ( š®āØ?ā© )" "tr_uni_eq" "tr_uni_hdtl" "tr_uni_tls" "tr_uni_pap" "tr_uni_compose" * ]
+          [ "tr_id ( š¢ ) " "tr_id_hdtl" "tr_id_pushs" "tr_id_tls" "tr_id_pap" "tr_id_compose" * ]
           [ "tr_compose ( ?āˆ˜? )" "tr_compose_eq" "tr_compose_tls" "tr_compose_pn" "tr_compose_pushs" "tr_compose_pap" "tr_compose_compose" * ]
           [ "tr_pap ( ?@āØ?ā© )" "tr_pap_eq" "tr_pap_tls" "tr_pap_pat" "tr_pap_pn" "tr_pap_pushs" "tr_pap_pap" * ]
           [ "tr_pushs ( ā«Æ*[?]? )" * ]