]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma
- more results on relocation
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_sle.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.tcs.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/relocation/rtmap_isid.ma".
16 include "ground_2/relocation/rtmap_isdiv.ma".
17
18 (* RELOCATION MAP ***********************************************************)
19
20 coinductive sle: relation rtmap ≝
21 | sle_push: ∀f1,f2,g1,g2. sle f1 f2 → ↑f1 = g1 → ↑f2 = g2 → sle g1 g2
22 | sle_next: ∀f1,f2,g1,g2. sle f1 f2 → ⫯f1 = g1 → ⫯f2 = g2 → sle g1 g2
23 | sle_weak: ∀f1,f2,g1,g2. sle f1 f2 → ↑f1 = g1 → ⫯f2 = g2 → sle g1 g2
24 .
25
26 interpretation "inclusion (rtmap)"
27    'subseteq t1 t2 = (sle t1 t2).
28
29 (* Basic properties *********************************************************)
30
31 corec lemma sle_refl: ∀f1,f2. f1 ≗ f2 → f1 ⊆ f2.
32 #f1 #f2 * -f1 -f2
33 #f1 #f2 #g1 #g2 #H12 #H1 #H2
34 [ @(sle_push … H1 H2) | @(sle_next … H1 H2) ] -H1 -H2 /2 width=1 by/
35 qed.
36
37 (* Basic inversion lemmas ***************************************************)
38
39 lemma sle_inv_xp: ∀g1,g2. g1 ⊆ g2 → ∀f2. ↑f2 = g2 →
40                   ∃∃f1. f1 ⊆ f2 & ↑f1 = g1.
41 #g1 #g2 * -g1 -g2
42 #f1 #f2 #g1 #g2 #H #H1 #H2 #x2 #Hx2 destruct
43 [ lapply (injective_push … Hx2) -Hx2 /2 width=3 by ex2_intro/ ]
44 elim (discr_push_next … Hx2)
45 qed-.
46
47 lemma sle_inv_nx: ∀g1,g2. g1 ⊆ g2 → ∀f1. ⫯f1 = g1 →
48                   ∃∃f2. f1 ⊆ f2 & ⫯f2 = g2.
49 #g1 #g2 * -g1 -g2
50 #f1 #f2 #g1 #g2 #H #H1 #H2 #x1 #Hx1 destruct
51 [2: lapply (injective_next … Hx1) -Hx1 /2 width=3 by ex2_intro/ ]
52 elim (discr_next_push … Hx1)
53 qed-.
54
55 lemma sle_inv_pn: ∀g1,g2. g1 ⊆ g2 → ∀f1,f2. ↑f1 = g1 → ⫯f2 = g2 → f1 ⊆ f2.
56 #g1 #g2 * -g1 -g2
57 #f1 #f2 #g1 #g2 #H #H1 #H2 #x1 #x2 #Hx1 #Hx2 destruct
58 [ elim (discr_next_push … Hx2)
59 | elim (discr_push_next … Hx1)
60 | lapply (injective_push … Hx1) -Hx1
61   lapply (injective_next … Hx2) -Hx2 //
62 ]
63 qed-.
64
65 (* Advanced inversion lemmas ************************************************)
66
67 lemma sle_inv_pp: ∀g1,g2. g1 ⊆ g2 → ∀f1,f2. ↑f1 = g1 → ↑f2 = g2 → f1 ⊆ f2.
68 #g1 #g2 #H #f1 #f2 #H1 #H2 elim (sle_inv_xp … H … H2) -g2
69 #x1 #H #Hx1 destruct lapply (injective_push … Hx1) -Hx1 //
70 qed-.
71
72 lemma sle_inv_nn: ∀g1,g2. g1 ⊆ g2 → ∀f1,f2.  ⫯f1 = g1 → ⫯f2 = g2 → f1 ⊆ f2.
73 #g1 #g2 #H #f1 #f2 #H1 #H2 elim (sle_inv_nx … H … H1) -g1
74 #x2 #H #Hx2 destruct lapply (injective_next … Hx2) -Hx2 //
75 qed-.
76
77 lemma sle_inv_px: ∀g1,g2. g1 ⊆ g2 → ∀f1. ↑f1 = g1 →
78                   (∃∃f2. f1 ⊆ f2 & ↑f2 = g2) ∨ ∃∃f2. f1 ⊆ f2 & ⫯f2 = g2.
79 #g1 #g2 elim (pn_split g2) * #f2 #H2 #H #f1 #H1
80 [ lapply (sle_inv_pp … H … H1 H2) | lapply (sle_inv_pn … H … H1 H2) ] -H -H1
81 /3 width=3 by ex2_intro, or_introl, or_intror/
82 qed-.
83
84 lemma sle_inv_xn: ∀g1,g2. g1 ⊆ g2 → ∀f2. ⫯f2 = g2 →
85                   (∃∃f1. f1 ⊆ f2 & ↑f1 = g1) ∨ ∃∃f1. f1 ⊆ f2 & ⫯f1 = g1.
86 #g1 #g2 elim (pn_split g1) * #f1 #H1 #H #f2 #H2
87 [ lapply (sle_inv_pn … H … H1 H2) | lapply (sle_inv_nn … H … H1 H2) ] -H -H2
88 /3 width=3 by ex2_intro, or_introl, or_intror/
89 qed-.
90
91 (* Main properties **********************************************************)
92
93 corec theorem sle_trans: Transitive … sle.
94 #f1 #f * -f1 -f
95 #f1 #f #g1 #g #Hf #H1 #H #g2 #H0
96 [ cases (sle_inv_px … H0 … H) * |*: cases (sle_inv_nx … H0 … H) ] -g
97 /3 width=5 by sle_push, sle_next, sle_weak/
98 qed-.
99
100 (* Properties with iteraded push ********************************************)
101
102 lemma sle_pushs: ∀f1,f2. f1 ⊆ f2 → ∀i. ↑*[i] f1 ⊆ ↑*[i] f2.
103 #f1 #f2 #Hf12 #i elim i -i /2 width=5 by sle_push/
104 qed.
105
106 (* Properties with tail *****************************************************)
107
108 lemma sle_px_tl: ∀g1,g2. g1 ⊆ g2 → ∀f1. ↑f1 = g1 → f1 ⊆ ⫱g2.
109 #g1 #g2 #H #f1 #H1 elim (sle_inv_px … H … H1) -H -H1 * //
110 qed.
111
112 lemma sle_xn_tl: ∀g1,g2. g1 ⊆ g2 → ∀f2. ⫯f2 = g2 → ⫱g1 ⊆ f2.
113 #g1 #g2 #H #f2 #H2 elim (sle_inv_xn … H … H2) -H -H2 * //
114 qed.
115
116 lemma sle_tl: ∀f1,f2. f1 ⊆ f2 → ⫱f1 ⊆ ⫱f2.
117 #f1 elim (pn_split f1) * #g1 #H1 #f2 #H
118 [ lapply (sle_px_tl … H … H1) -H //
119 | elim (sle_inv_nx … H … H1) -H //
120 ]
121 qed.
122
123 (* Inversion lemmas with tail ***********************************************)
124
125 lemma sle_inv_tl_sn: ∀f1,f2. ⫱f1 ⊆ f2 → f1 ⊆ ⫯f2.
126 #f1 elim (pn_split f1) * #g1 #H destruct
127 /2 width=5 by sle_next, sle_weak/
128 qed-.
129
130 lemma sle_inv_tl_dx: ∀f1,f2. f1 ⊆ ⫱f2 → ↑f1 ⊆ f2.
131 #f1 #f2 elim (pn_split f2) * #g2 #H destruct
132 /2 width=5 by sle_push, sle_weak/
133 qed-.
134
135 (* Properties with isid *****************************************************)
136
137 corec lemma sle_isid_sn: ∀f1. 𝐈⦃f1⦄ → ∀f2. f1 ⊆ f2.
138 #f1 * -f1
139 #f1 #g1 #Hf1 #H1 #f2 cases (pn_split f2) *
140 /3 width=5 by sle_weak, sle_push/
141 qed.
142
143 (* Inversion lemmas with isid ***********************************************)
144
145 corec lemma sle_inv_isid_dx: ∀f1,f2. f1 ⊆ f2 → 𝐈⦃f2⦄ → 𝐈⦃f1⦄.
146 #f1 #f2 * -f1 -f2
147 #f1 #f2 #g1 #g2 #Hf * * #H
148 [2,3: elim (isid_inv_next … H) // ]
149 lapply (isid_inv_push … H ??) -H
150 /3 width=3 by isid_push/
151 qed-.
152
153 (* Properties with isdiv ****************************************************)
154
155 corec lemma sle_isdiv_dx: ∀f2. 𝛀⦃f2⦄ → ∀f1. f1 ⊆ f2.
156 #f2 * -f2
157 #f2 #g2 #Hf2 #H2 #f1 cases (pn_split f1) *
158 /3 width=5 by sle_weak, sle_next/
159 qed.
160
161 (* Inversion lemmas with isdiv **********************************************)
162
163 corec lemma sle_inv_isdiv_sn: ∀f1,f2. f1 ⊆ f2 → 𝛀⦃f1⦄ → 𝛀⦃f2⦄.
164 #f1 #f2 * -f1 -f2
165 #f1 #f2 #g1 #g2 #Hf * * #H
166 [1,3: elim (isdiv_inv_push … H) // ]
167 lapply (isdiv_inv_next … H ??) -H
168 /3 width=3 by isdiv_next/
169 qed-.