"identity element (total relocation streams)"
'ElementI = (tr_id).
-(* Basic constructions (specific) *******************************************)
+(* Basic constructions ******************************************************)
lemma tr_id_unfold: ā«Æš¢ = š¢.
<(stream_unfold ā¦ (š¢)) in ā¢ (???%); //
--- /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 "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.
--- /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 "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.
--- /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 "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.
--- /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 "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.
}
]
[ { "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 ( ā«Æ*[?]? )" * ]