]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_lift.ma
precommit for rtmap ...
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / nstream_lift.ma
index 07a96e4d570826180da25c2fcdf7d96f8bf424ed..efb7b3810985d3951028493b5c524f1413d6dbfb 100644 (file)
 (**************************************************************************)
 
 include "ground_2/notation/functions/lift_1.ma".
-include "ground_2/relocation/nstream_at.ma".
+include "ground_2/relocation/nstream.ma".
 
 (* RELOCATION N-STREAM ******************************************************)
 
-definition push: nstream → nstream ≝ λt. 0@t.
+definition push: rtmap → rtmap ≝ λf. 0@f.
 
-interpretation "push (nstream)" 'Lift t = (push t).
+interpretation "push (nstream)" 'Lift f = (push f).
 
-definition next: nstream → nstream.
-* #a #t @(⫯a@t)
+definition next: rtmap → rtmap.
+* #n #f @(⫯n@f)
 qed.
 
-interpretation "next (nstream)" 'Successor t = (next t).
+interpretation "next (nstream)" 'Successor f = (next f).
 
-(* Basic properties on push *************************************************)
+(* Basic properties *********************************************************)
 
-lemma push_at_O: ∀t. @⦃0, ↑t⦄ ≡ 0.
+lemma push_rew: ∀f. 0@f = ↑f.
 // qed.
 
-lemma push_at_S: ∀t,i1,i2. @⦃i1, t⦄ ≡ i2 → @⦃⫯i1, ↑t⦄ ≡ ⫯i2.
-/2 width=1 by at_S1/ qed.
+lemma next_rew: ∀f,n. (⫯n)@f = ⫯(n@f).
+// qed.
+(*
+lemma next_rew_sn: ∀f,n1,n2. n1 = ⫯n2 → n1@f = ⫯(n2@f).
+// qed.
+*)
+(* Basic inversion lemmas ***************************************************)
 
-(* Basic inversion lemmas on push *******************************************)
+lemma injective_push: injective ? ? push.
+#f1 #f2 normalize #H destruct //
+qed-.
 
-lemma push_inv_at_S: ∀t,i1,i2. @⦃⫯i1, ↑t⦄ ≡ ⫯i2 → @⦃i1, t⦄ ≡ i2.
-/2 width=1 by at_inv_SOS/ qed-.
+lemma discr_push_next: ∀f1,f2. ↑f1 = ⫯f2 → ⊥.
+#f1 * #n2 #f2 normalize #H destruct
+qed-.
 
-lemma injective_push: injective ? ? push.
-#t1 #t2 normalize #H destruct //
+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-.
 
-(* Basic properties on next *************************************************)
+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 next_at: ∀t,i1,i2. @⦃i1, t⦄ ≡ i2 → @⦃i1, ⫯t⦄ ≡ ⫯i2.
-* /2 width=1 by at_lift/
-qed.
+lemma push_inv_seq_dx: ∀f,g,n. ↑f = n@g → 0 = n ∧ g = f.
+#f #g #n <push_rew #H destruct /2 width=1 by conj/
+qed-.
 
-(* Basic inversion lemmas on next *******************************************)
+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_at: ∀t,i1,i2. @⦃i1, ⫯t⦄ ≡ ⫯i2 → @⦃i1, t⦄ ≡ i2.
-* /2 width=1 by at_inv_xSS/
+lemma next_inv_seq_dx: ∀f,g,n. ⫯f = n@g → ∃∃m. m@g = f & ⫯m = n.
+* #m #f #g #n <next_rew #H destruct /2 width=3 by ex2_intro/
 qed-.
 
-lemma injective_next: injective ? ? next.
-* #a1 #t1 * #a2 #t2 normalize #H destruct //
+lemma pn_split: ∀f. (∃g. ↑g = f) ∨ (∃g. ⫯g = f).
+* * /3 width=2 by or_introl, or_intror, ex_intro/
 qed-.