X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Fnstream.ma;h=c5ccb7373a489400abfad21e4e23e1170713cb11;hb=e8fb201bad04ec30867659c2d42ef45a4b6c3393;hp=ccd8857a8a7b254f69586f516d17c1496c06842a;hpb=750305d95b8f6bb40b5be0e9dfd05d42b256f2a1;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream.ma index ccd8857a8..c5ccb7373 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream.ma @@ -12,9 +12,86 @@ (* *) (**************************************************************************) -include "ground_2/lib/arith.ma". -include "ground_2/lib/streams.ma". +include "ground_2/notation/functions/upspoon_1.ma". +include "ground_2/lib/stream_tls.ma". (* RELOCATION N-STREAM ******************************************************) -definition nstream: Type[0] ≝ stream nat. +definition rtmap: Type[0] ≝ stream nat. + +definition push: rtmap → rtmap ≝ λf. 0⨮f. + +interpretation "push (nstream)" 'UpSpoon f = (push f). + +definition next: rtmap → rtmap. +* #n #f @(↑n⨮f) +defined. + +interpretation "next (nstream)" 'UpArrow f = (next f). + +(* Basic properties *********************************************************) + +lemma push_rew: ∀f. 0⨮f = ⫯f. +// qed. + +lemma next_rew: ∀f,n. (↑n)⨮f = ↑(n⨮f). +// qed. + +(* Basic inversion lemmas ***************************************************) + +lemma injective_push: injective ? ? push. +#f1 #f2 normalize #H destruct // +qed-. + +lemma discr_push_next: ∀f1,f2. ⫯f1 = ↑f2 → ⊥. +#f1 * #n2 #f2 normalize #H destruct +qed-. + +lemma discr_next_push: ∀f1,f2. ↑f1 = ⫯f2 → ⊥. +* #n1 #f1 #f2 normalize #H destruct +qed-. + +lemma injective_next: injective ? ? next. +* #n1 #f1 * #n2 #f2 normalize #H destruct // +qed-. + +lemma push_inv_seq_sn: ∀f,g,n. n⨮g = ⫯f → 0 = n ∧ g = f. +#f #g #n