--- /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/notation/functions/element_i_0.ma".
+include "ground/relocation/tr_pn.ma".
+
+(* IDENTITY ELEMENT FOR TOTAL RELOCATION MAPS *******************************)
+
+corec definition tr_id: tr_map ≝ ⫯tr_id.
+
+interpretation
+ "identity element (total relocation streams)"
+ 'ElementI = (tr_id).
+
+(* Basic constructions (specific) *******************************************)
+
+lemma tr_id_unfold: ⫯𝐢 = 𝐢.
+<(stream_unfold … (𝐢)) in ⊢ (???%); //
+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/arith/nat_succ.ma".
+include "ground/notation/functions/element_u_1.ma".
+include "ground/relocation/tr_id.ma".
+
+(* UNIFORM ELEMENTS FOR TOTAL RELOCATION MAPS *******************************)
+
+definition tr_uni (n:nat): tr_map ≝ ↑n ⨮ 𝐢.
+
+interpretation
+ "uniform elements (total relocation maps)"
+ 'ElementU n = (tr_uni n).
+
+(* Basic constructions ******************************************************)
+
+lemma tr_uni_zero: 𝐢 = 𝐮❨𝟎❩.
+<tr_id_unfold //
+qed.
+
+lemma tr_uni_succ (n): ↑𝐮❨n❩ = 𝐮❨↑n❩.
+// qed.
}
]
[ { "total relocation" * } {
+ [ "tr_uni ( 𝐮❨?❩ )" * ]
+ [ "tr_id ( 𝐢 ) " * ]
[ "tr_compose ( ?∘? )" "tr_compose_pn" * ]
[ "tr_pap ( ?@❨?❩ )" "tr_pap_eq" "tr_pap_pat" "tr_pap_pn" "tr_pap_pap" * ]
[ "tr_pn ( ⫯? ) ( ↑? )" "tr_pn_hdtl" "tr_pn_tls" * ]