]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma
- exclusion binder added in local environments
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lfxs_lfxs.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 "basic_2/relocation/lexs_lexs.ma".
16 include "basic_2/static/frees_drops.ma".
17 include "basic_2/static/lfxs.ma".
18
19 (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
20
21 (* Advanced properties ******************************************************)
22
23 lemma lfxs_inv_frees: ∀R,L1,L2,T. L1 ⪤*[R, T] L2 →
24                       ∀f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ⪤*[cext2 R, cfull, f] L2.
25 #R #L1 #L2 #T * /3 width=6 by frees_mono, lexs_eq_repl_back/
26 qed-.
27
28 (* Basic_2A1: uses: llpx_sn_dec *)
29 lemma lfxs_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
30                 ∀L1,L2,T. Decidable (L1 ⪤*[R, T] L2).
31 #R #HR #L1 #L2 #T
32 elim (frees_total L1 T) #f #Hf
33 elim (lexs_dec (cext2 R) cfull … L1 L2 f)
34 /4 width=3 by lfxs_inv_frees, cfull_dec, ext2_dec, ex2_intro, or_intror, or_introl/
35 qed-.
36
37 lemma lfxs_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
38                           lexs_frees_confluent … (cext2 R1) cfull →
39                           ∀L1,L2,V. L1 ⪤*[R1, V] L2 → ∀I,T.
40                           ∃∃L. L1 ⪤*[R1, ②{I}V.T] L & L ⪤*[R2, V] L2.
41 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T
42 [ elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg
43   elim (frees_inv_bind … Hg) #y1 #y2 #H #_ #Hy
44 | elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg
45   elim (frees_inv_flat … Hg) #y1 #y2 #H #_ #Hy
46 ]
47 lapply(frees_mono … H … Hf) -H #H1
48 lapply (sor_eq_repl_back1 … Hy … H1) -y1 #Hy
49 lapply (sor_inv_sle_sn … Hy) -y2 #Hfg
50 elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2
51 lapply (sle_lexs_trans … HL1 … Hfg) // #H
52 elim (HR … Hf … H) -HR -Hf -H
53 /4 width=7 by sle_lexs_trans, ex2_intro/
54 qed-.
55
56 lemma lfxs_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
57                           lexs_frees_confluent … (cext2 R1) cfull →
58                           ∀L1,L2,T. L1 ⪤*[R1, T] L2 → ∀I,V.
59                           ∃∃L. L1 ⪤*[R1, ⓕ{I}V.T] L & L ⪤*[R2, T] L2.
60 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V
61 elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg
62 elim (frees_inv_flat … Hg) #y1 #y2 #_ #H #Hy
63 lapply(frees_mono … H … Hf) -H #H2
64 lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy
65 lapply (sor_inv_sle_dx … Hy) -y1 #Hfg
66 elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #L #HL1 #HL2
67 lapply (sle_lexs_trans … HL1 … Hfg) // #H
68 elim (HR … Hf … H) -HR -Hf -H
69 /4 width=7 by sle_lexs_trans, ex2_intro/
70 qed-.
71
72 lemma lfxs_bind_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
73                           lexs_frees_confluent … (cext2 R1) cfull →
74                           ∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤*[R1, T] L2 → ∀p.
75                           ∃∃L,V. L1 ⪤*[R1, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤*[R2, T] L2 & R1 L1 V1 V.
76 #R1 #R2 #HR1 #HR2 #HR #I #L1 #L2 #V1 #T * #f #Hf #HL12 #p
77 elim (frees_total L1 (ⓑ{p,I}V1.T)) #g #Hg
78 elim (frees_inv_bind … Hg) #y1 #y2 #_ #H #Hy
79 lapply(frees_mono … H … Hf) -H #H2
80 lapply (tl_eq_repl … H2) -H2 #H2
81 lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy
82 lapply (sor_inv_sle_dx … Hy) -y1 #Hfg
83 lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg
84 elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2
85 lapply (sle_lexs_trans … H … Hfg) // #H0
86 elim (lexs_inv_next1 … H) -H #Z #L #HL1 #H
87 elim (ext2_inv_pair_sn … H) -H #V #HV #H1 #H2 destruct
88 elim (HR … Hf … H0) -HR -Hf -H0
89 /4 width=7 by sle_lexs_trans, ex3_2_intro, ex2_intro/
90 qed-.
91
92 (* Main properties **********************************************************)
93
94 (* Basic_2A1: uses: llpx_sn_bind llpx_sn_bind_O *)
95 theorem lfxs_bind: ∀R,p,I,L1,L2,V1,V2,T.
96                    L1 ⪤*[R, V1] L2 → L1.ⓑ{I}V1 ⪤*[R, T] L2.ⓑ{I}V2 →
97                    L1 ⪤*[R, ⓑ{p,I}V1.T] L2.
98 #R #p #I #L1 #L2 #V1 #V2 #T * #f1 #HV #Hf1 * #f2 #HT #Hf2
99 lapply (lexs_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2))
100 /3 width=7 by frees_fwd_isfin, frees_bind, lexs_join, isfin_tl, ex2_intro/
101 qed.
102
103 (* Basic_2A1: llpx_sn_flat *)
104 theorem lfxs_flat: ∀R,I,L1,L2,V,T.
105                    L1 ⪤*[R, V] L2 → L1 ⪤*[R, T] L2 →
106                    L1 ⪤*[R, ⓕ{I}V.T] L2.
107 #R #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 elim (sor_isfin_ex f1 f2)
108 /3 width=7 by frees_fwd_isfin, frees_flat, lexs_join, ex2_intro/
109 qed.
110
111 theorem lfxs_bind_void: ∀R,p,I,L1,L2,V,T.
112                         L1 ⪤*[R, V] L2 → L1.ⓧ ⪤*[R, T] L2.ⓧ →
113                         L1 ⪤*[R, ⓑ{p,I}V.T] L2.
114 #R #p #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2
115 lapply (lexs_fwd_bind … Hf2) -Hf2 #Hf2 elim (sor_isfin_ex f1 (⫱f2))
116 /3 width=7 by frees_fwd_isfin, frees_bind_void, lexs_join, isfin_tl, ex2_intro/
117 qed.
118
119 theorem lfxs_conf: ∀R1,R2.
120                    lexs_frees_confluent (cext2 R1) cfull →
121                    lexs_frees_confluent (cext2 R2) cfull →
122                    R_confluent2_lfxs R1 R2 R1 R2 →
123                    ∀T. confluent2 … (lfxs R1 T) (lfxs R2 T).
124 #R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02
125 lapply (frees_mono … Hf1 … Hf) -Hf1 #Hf12
126 lapply (lexs_eq_repl_back … HL01 … Hf12) -f1 #HL01
127 elim (lexs_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ]
128 [ #L #HL1 #HL2
129   elim (HR1 … Hf … HL01) -HL01 #f1 #Hf1 #H1
130   elim (HR2 … Hf … HL02) -HL02 #f2 #Hf2 #H2
131   lapply (sle_lexs_trans … HL1 … H1) // -HL1 -H1 #HL1
132   lapply (sle_lexs_trans … HL2 … H2) // -HL2 -H2 #HL2
133   /3 width=5 by ex2_intro/
134 | #g * #I0 [2: #V0 ] #K0 #n #HLK0 #Hgf #Z1 #H1 #Z2 #H2 #K1 #HK01 #K2 #HK02
135   [ elim (ext2_inv_pair_sn … H1) -H1 #V1 #HV01 #H destruct
136     elim (ext2_inv_pair_sn … H2) -H2 #V2 #HV02 #H destruct
137     elim (frees_inv_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0
138     lapply (sle_lexs_trans … HK01 … H0) // -HK01 #HK01
139     lapply (sle_lexs_trans … HK02 … H0) // -HK02 #HK02
140     elim (HR12 … HV01 … HV02 K1 … K2) /3 width=3 by ext2_pair, ex2_intro/
141   | lapply (ext2_inv_unit_sn … H1) -H1 #H destruct
142     lapply (ext2_inv_unit_sn … H2) -H2 #H destruct
143     /3 width=3 by ext2_unit, ex2_intro/
144   ]
145 ]
146 qed-.
147
148 (* Negated inversion lemmas *************************************************)
149
150 (* Basic_2A1: uses: nllpx_sn_inv_bind nllpx_sn_inv_bind_O *)
151 lemma lfnxs_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
152                       ∀p,I,L1,L2,V,T. (L1 ⪤*[R, ⓑ{p,I}V.T] L2 → ⊥) →
153                       (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1.ⓑ{I}V ⪤*[R, T] L2.ⓑ{I}V → ⊥).
154 #R #HR #p #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
155 /4 width=2 by lfxs_bind, or_intror, or_introl/
156 qed-.
157
158 (* Basic_2A1: uses: nllpx_sn_inv_flat *)
159 lemma lfnxs_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
160                       ∀I,L1,L2,V,T. (L1 ⪤*[R, ⓕ{I}V.T] L2 → ⊥) →
161                       (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1 ⪤*[R, T] L2 → ⊥).
162 #R #HR #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
163 /4 width=1 by lfxs_flat, or_intror, or_introl/
164 qed-.
165
166 lemma lfnxs_inv_bind_void: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
167                            ∀p,I,L1,L2,V,T. (L1 ⪤*[R, ⓑ{p,I}V.T] L2 → ⊥) →
168                            (L1 ⪤*[R, V] L2 → ⊥) ∨ (L1.ⓧ ⪤*[R, T] L2.ⓧ → ⊥).
169 #R #HR #p #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
170 /4 width=2 by lfxs_bind_void, or_intror, or_introl/
171 qed-.