]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/nstream.ma
minor corrections and updates
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / nstream.ma
index aaa1ba9b6551f67099861a199fdf067f9169e558..c5ccb7373a489400abfad21e4e23e1170713cb11 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/notation/functions/lift_1.ma".
-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 rtmap: Type[0] ≝ stream nat.
 
-definition push: rtmap → rtmap ≝ λf. 0@f.
+definition push: rtmap → rtmap ≝ λf. 0f.
 
-interpretation "push (nstream)" 'Lift f = (push f).
+interpretation "push (nstream)" 'UpSpoon f = (push f).
 
 definition next: rtmap → rtmap.
-* #n #f @(⫯n@f)
-qed.
+* #n #f @(â\86\91n⨮f)
+defined.
 
-interpretation "next (nstream)" 'Successor f = (next f).
+interpretation "next (nstream)" 'UpArrow f = (next f).
 
 (* Basic properties *********************************************************)
 
-lemma push_rew: ∀f. 0@f = ↑f.
+lemma push_rew: ∀f. 0⨮f = ⫯f.
 // qed.
 
-lemma next_rew: â\88\80f,n. (⫯n)@f = â«¯(n@f).
+lemma next_rew: â\88\80f,n. (â\86\91n)⨮f = â\86\91(n⨮f).
 // qed.
 
 (* Basic inversion lemmas ***************************************************)
@@ -44,11 +43,11 @@ lemma injective_push: injective ? ? push.
 #f1 #f2 normalize #H destruct //
 qed-.
 
-lemma discr_push_next: â\88\80f1,f2. â\86\91f1 = â«¯f2 → ⊥.
+lemma discr_push_next: â\88\80f1,f2. â«¯f1 = â\86\91f2 → ⊥.
 #f1 * #n2 #f2 normalize #H destruct
 qed-.
 
-lemma discr_next_push: â\88\80f1,f2. â«¯f1 = â\86\91f2 → ⊥.
+lemma discr_next_push: â\88\80f1,f2. â\86\91f1 = â«¯f2 → ⊥.
 * #n1 #f1 #f2 normalize #H destruct
 qed-.
 
@@ -56,35 +55,43 @@ 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.
+lemma push_inv_seq_sn: ∀f,g,n. n⨮g = ⫯f → 0 = n ∧ g = f.
 #f #g #n <push_rew #H destruct /2 width=1 by conj/
 qed-.
 
-lemma push_inv_seq_dx: â\88\80f,g,n. â\86\91f = n@g → 0 = n ∧ g = f.
+lemma push_inv_seq_dx: â\88\80f,g,n. â«¯f = n⨮g → 0 = n ∧ g = f.
 #f #g #n <push_rew #H destruct /2 width=1 by conj/
 qed-.
 
-lemma next_inv_seq_sn: ∀f,g,n. n@g = ⫯f → ∃∃m. m@g = f & ⫯m = n.
+lemma next_inv_seq_sn: ∀f,g,n. n⨮g = ↑f → ∃∃m. m⨮g = f & ↑m = n.
 * #m #f #g #n <next_rew #H destruct /2 width=3 by ex2_intro/
 qed-.
 
-lemma next_inv_seq_dx: â\88\80f,g,n. â«¯f = n@g â\86\92 â\88\83â\88\83m. m@g = f & â«¯m = n.
+lemma next_inv_seq_dx: â\88\80f,g,n. â\86\91f = n⨮g â\86\92 â\88\83â\88\83m. m⨮g = f & â\86\91m = n.
 * #m #f #g #n <next_rew #H destruct /2 width=3 by ex2_intro/
 qed-.
 
 lemma case_prop: ∀R:predicate rtmap.
-                 (â\88\80f. R (â\86\91f)) â\86\92 (â\88\80f. R (⫯f)) → ∀f. R f.
+                 (â\88\80f. R (⫯f)) â\86\92 (â\88\80f. R (â\86\91f)) → ∀f. R f.
 #R #H1 #H2 * * //
 qed-.
 
 lemma case_type0: ∀R:rtmap→Type[0].
-                  (â\88\80f. R (â\86\91f)) â\86\92 (â\88\80f. R (⫯f)) → ∀f. R f.
+                  (â\88\80f. R (⫯f)) â\86\92 (â\88\80f. R (â\86\91f)) → ∀f. R f.
 #R #H1 #H2 * * //
 qed-.
 
-lemma iota_push: â\88\80R,a,b,f. a f = case_type0 R a b (â\86\91f).
+lemma iota_push: â\88\80R,a,b,f. a f = case_type0 R a b (⫯f).
 // qed.
 
-lemma iota_next: â\88\80R,a,b,f. b f = case_type0 R a b (⫯f).
+lemma iota_next: â\88\80R,a,b,f. b f = case_type0 R a b (â\86\91f).
 #R #a #b * //
 qed.
+
+(* Specific properties ******************************************************)
+
+lemma tl_push: ∀f. f = ⫰⫯f.
+// qed.
+
+lemma tl_next: ∀f. ⫰f = ⫰↑f.
+* // qed.