]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_id.ma
milestone update in ground_2 and basic_2A
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / nstream_id.ma
index 11307fbd245c0d46c7832c0a4b5bc8e836701959..fcb6cca7208380c55ebe6a51a7cc987882d49284 100644 (file)
@@ -17,17 +17,17 @@ include "ground_2/relocation/rtmap_eq.ma".
 
 (* RELOCATION N-STREAM ******************************************************)
 
-let corec id: rtmap ≝ ↑id.
+corec definition id: rtmap ≝ ⫯id.
 
 interpretation "identity (nstream)"
    'Identity = (id).
 
 (* Basic properties *********************************************************)
 
-lemma id_rew: â\86\91𝐈𝐝 = 𝐈𝐝.
+lemma id_rew: â«¯𝐈𝐝 = 𝐈𝐝.
 <(stream_rew … (𝐈𝐝)) in ⊢ (???%); normalize //
 qed.
 
-lemma id_eq_rew: â\86\91ð\9d\90\88ð\9d\90\9d â\89\97 𝐈𝐝.
+lemma id_eq_rew: â«¯ð\9d\90\88ð\9d\90\9d â\89¡ 𝐈𝐝.
 cases id_rew in ⊢ (??%); //
 qed.