]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_lift.ma
- ground_2: support for relocation updated
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / nstream_lift.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ground_2/notation/functions/lift_1.ma".
16 include "ground_2/relocation/nstream.ma".
17
18 (* RELOCATION N-STREAM ******************************************************)
19
20 definition push: rtmap → rtmap ≝ λf. 0@f.
21
22 interpretation "push (nstream)" 'Lift f = (push f).
23
24 definition next: rtmap → rtmap.
25 * #n #f @(⫯n@f)
26 qed.
27
28 interpretation "next (nstream)" 'Successor f = (next f).
29
30 (* Basic properties *********************************************************)
31
32 lemma push_eq_repl: eq_stream_repl … (λf1,f2. ↑f1 ≐ ↑f2).
33 /2 width=1 by eq_seq/ qed.
34
35 lemma next_eq_repl: eq_stream_repl … (λf1,f2. ⫯f1 ≐ ⫯f2).
36 * #n1 #f1 * #n2 #f2 #H elim (eq_stream_inv_seq ????? H) -H
37 /2 width=1 by eq_seq/
38 qed.
39
40 lemma push_rew: ∀f. ↑f = 0@f.
41 // qed.
42
43 lemma next_rew: ∀f,n. ⫯(n@f) = (⫯n)@f.
44 // qed.
45
46 lemma next_rew_sn: ∀f,n1,n2. n1 = ⫯n2 → n1@f = ⫯(n2@f).
47 // qed.
48
49 (* Basic inversion lemmas ***************************************************)
50
51 lemma injective_push: injective ? ? push.
52 #f1 #f2 normalize #H destruct //
53 qed-.
54
55 lemma discr_push_next: ∀f1,f2. ↑f1 = ⫯f2 → ⊥.
56 #f1 * #n2 #f2 normalize #H destruct
57 qed-.
58
59 lemma discr_next_push: ∀f1,f2. ⫯f1 = ↑f2 → ⊥.
60 * #n1 #f1 #f2 normalize #H destruct
61 qed-.
62
63 lemma injective_next: injective ? ? next.
64 * #n1 #f1 * #n2 #f2 normalize #H destruct //
65 qed-.
66
67 lemma push_inv_seq_sn: ∀f,g,n. n@g = ↑f → n = 0 ∧ g = f.
68 #f #g #n >push_rew #H destruct /2 width=1 by conj/
69 qed-.
70
71 lemma push_inv_seq_dx: ∀f,g,n. ↑f = n@g → n = 0 ∧ g = f.
72 #f #g #n >push_rew #H destruct /2 width=1 by conj/
73 qed-.
74
75 lemma next_inv_seq_sn: ∀f,g,n. n@g = ⫯f → ∃∃m. n = ⫯m & f = m@g.
76 * #m #f #g #n >next_rew #H destruct /2 width=3 by ex2_intro/
77 qed-.
78
79 lemma next_inv_seq_dx: ∀f,g,n. ⫯f = n@g → ∃∃m. n = ⫯m & f = m@g.
80 * #m #f #g #n >next_rew #H destruct /2 width=3 by ex2_intro/
81 qed-.
82
83 lemma push_inv_dx: ∀x,f. x ≐ ↑f → ∃∃g. x = ↑g & g ≐ f.
84 * #m #x #f #H elim (eq_stream_inv_seq ????? H) -H
85 /2 width=3 by ex2_intro/
86 qed-.
87
88 lemma push_inv_sn: ∀f,x. ↑f ≐ x → ∃∃g. x = ↑g & f ≐ g.
89 #f #x #H lapply (eq_stream_sym … H) -H
90 #H elim (push_inv_dx … H) -H
91 /3 width=3 by eq_stream_sym, ex2_intro/
92 qed-.
93
94 lemma push_inv_bi: ∀f,g. ↑f ≐ ↑g → f ≐ g.
95 #f #g #H elim (push_inv_dx … H) -H
96 #x #H #Hg lapply (injective_push … H) -H //
97 qed-.
98
99 lemma next_inv_dx: ∀x,f. x ≐ ⫯f → ∃∃g. x = ⫯g & g ≐ f.
100 * #m #x * #n #f #H elim (eq_stream_inv_seq ????? H) -H
101 /3 width=5 by eq_seq, ex2_intro/
102 qed-.
103
104 lemma next_inv_sn: ∀f,x. ⫯f ≐ x → ∃∃g. x = ⫯g & f ≐ g.
105 #f #x #H lapply (eq_stream_sym … H) -H
106 #H elim (next_inv_dx … H) -H
107 /3 width=3 by eq_stream_sym, ex2_intro/
108 qed-.
109
110 lemma next_inv_bi: ∀f,g. ⫯f ≐ ⫯g → f ≐ g.
111 #f #g #H elim (next_inv_dx … H) -H
112 #x #H #Hg lapply (injective_next … H) -H //
113 qed-.