]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground/etc/relocation/pstream.etc
update in delayed_updating
[helm.git] / matita / matita / contribs / lambdadelta / ground / etc / relocation / pstream.etc
index b641364f7e5d08737e9abc5c16f1301d42330672..8827703c5f62b3c7daed862894aa89a697cd6323 100644 (file)
@@ -18,26 +18,6 @@ include "ground/arith/pnat.ma".
 
 (* RELOCATION P-STREAM ******************************************************)
 
-definition gr_map: Type[0] ≝ stream pnat.
-
-definition gr_push: gr_map → gr_map ≝ λf. 𝟏⨮f.
-
-interpretation "push (pstream)" 'UpSpoon f = (gr_push f).
-
-definition gr_next: gr_map → gr_map.
-* #p #f @(↑p⨮f)
-defined.
-
-interpretation "next (pstream)" 'UpArrow f = (gr_next f).
-
-(* Basic properties *********************************************************)
-
-lemma gr_push_unfold: ∀f. 𝟏⨮f = ⫯f.
-// qed.
-
-lemma gr_next_unfold: ∀f,p. (↑p)⨮f = ↑(p⨮f).
-// qed.
-
 (* Basic inversion lemmas ***************************************************)
 
 lemma eq_inv_gr_push_bi: injective ? ? gr_push.
@@ -56,22 +36,6 @@ lemma eq_inv_gr_next_bi: injective ? ? gr_next.
 * #p1 #f1 * #p2 #f2 <gr_next_unfold <gr_next_unfold #H destruct //
 qed-.
 
-lemma push_inv_seq_sn: ∀f,g,p. p⨮g = ⫯f → ∧∧ 𝟏 = p & g = f.
-#f #g #p <gr_push_unfold #H destruct /2 width=1 by conj/
-qed-.
-
-lemma push_inv_seq_dx: ∀f,g,p. ⫯f = p⨮g → ∧∧ 𝟏 = p & g = f.
-#f #g #p <gr_push_unfold #H destruct /2 width=1 by conj/
-qed-.
-
-lemma next_inv_seq_sn: ∀f,g,p. p⨮g = ↑f → ∃∃q. q⨮g = f & ↑q = p.
-* #q #f #g #p <gr_next_unfold #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-lemma next_inv_seq_dx: ∀f,g,p. ↑f = p⨮g → ∃∃q. q⨮g = f & ↑q = p.
-* #q #f #g #p <gr_next_unfold #H destruct /2 width=3 by ex2_intro/
-qed-.
-
 lemma case_prop (Q:predicate gr_map):
       (∀f. Q (⫯f)) → (∀f. Q (↑f)) → ∀f. Q f.
 #Q #H1 #H2 * * //