]> 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 dc462b3da4218a9b9588270c864a8d98e67f7e51..efb7b3810985d3951028493b5c524f1413d6dbfb 100644 (file)
@@ -29,23 +29,15 @@ interpretation "next (nstream)" 'Successor f = (next f).
 
 (* Basic properties *********************************************************)
 
-lemma push_eq_repl: eq_stream_repl … (λf1,f2. ↑f1 ≐ ↑f2).
-/2 width=1 by eq_seq/ qed.
-
-lemma next_eq_repl: eq_stream_repl … (λf1,f2. ⫯f1 ≐ ⫯f2).
-* #n1 #f1 * #n2 #f2 #H elim (eq_stream_inv_seq ????? H) -H
-/2 width=1 by eq_seq/
-qed.
-
-lemma push_rew: ∀f. ↑f = 0@f.
+lemma push_rew: ∀f. 0@f = ↑f.
 // qed.
 
-lemma next_rew: ∀f,n. ⫯(n@f) = (⫯n)@f.
+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 ***************************************************)
 
 lemma injective_push: injective ? ? push.
@@ -64,50 +56,22 @@ lemma injective_next: injective ? ? next.
 * #n1 #f1 * #n2 #f2 normalize #H destruct //
 qed-.
 
-lemma push_inv_seq_sn: ∀f,g,n. n@g = ↑f → n = 0 ∧ g = f.
-#f #g #n >push_rew #H destruct /2 width=1 by conj/
-qed-.
-
-lemma push_inv_seq_dx: ∀f,g,n. ↑f = n@g → n = 0 ∧ 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. n = ⫯m & f = m@g.
-* #m #f #g #n >next_rew #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-lemma next_inv_seq_dx: ∀f,g,n. ⫯f = n@g → ∃∃m. n = ⫯m & f = m@g.
-* #m #f #g #n >next_rew #H destruct /2 width=3 by ex2_intro/
-qed-.
-
-lemma push_inv_dx: ∀x,f. x ≐ ↑f → ∃∃g. x = ↑g & g ≐ f.
-* #m #x #f #H elim (eq_stream_inv_seq ????? H) -H
-/2 width=3 by ex2_intro/
-qed-.
-
-lemma push_inv_sn: ∀f,x. ↑f ≐ x → ∃∃g. x = ↑g & f ≐ g.
-#f #x #H lapply (eq_stream_sym … H) -H
-#H elim (push_inv_dx … H) -H
-/3 width=3 by eq_stream_sym, ex2_intro/
+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_bi: ∀f,g. ↑f ≐ ↑g → f ≐ g.
-#f #g #H elim (push_inv_dx … H) -H
-#x #H #Hg lapply (injective_push … H) -H //
+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-.
 
-lemma next_inv_dx: ∀x,f. x ≐ ⫯f → ∃∃g. x = ⫯g & g ≐ f.
-* #m #x * #n #f #H elim (eq_stream_inv_seq ????? H) -H
-/3 width=5 by eq_seq, ex2_intro/
+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_sn: ∀f,x. ⫯f ≐ x → ∃∃g. x = ⫯g & f ≐ g.
-#f #x #H lapply (eq_stream_sym … H) -H
-#H elim (next_inv_dx … H) -H
-/3 width=3 by eq_stream_sym, ex2_intro/
+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 next_inv_bi: ∀f,g. ⫯f ≐ ⫯g → f ≐ g.
-#f #g #H elim (next_inv_dx … H) -H
-#x #H #Hg lapply (injective_next … H) -H //
+lemma pn_split: ∀f. (∃g. ↑g = f) ∨ (∃g. ⫯g = f).
+* * /3 width=2 by or_introl, or_intror, ex_intro/
 qed-.