]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma
- lfpxs based on tc_lfxs
[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_fqup.ma".
17 include "basic_2/static/frees_frees.ma".
18 include "basic_2/static/lfxs.ma".
19
20 (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
21
22 (* Advanced properties ******************************************************)
23
24 lemma lfxs_inv_frees: ∀R,L1,L2,T. L1 ⦻*[R, T] L2 →
25                       ∀f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → L1 ⦻*[R, cfull, f] L2.
26 #R #L1 #L2 #T * /3 width=6 by frees_mono, lexs_eq_repl_back/
27 qed-.
28
29 (* Basic_2A1: uses: llpx_sn_dec *)
30 lemma lfxs_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
31                 ∀L1,L2,T. Decidable (L1 ⦻*[R, T] L2).
32 #R #HR #L1 #L2 #T
33 elim (frees_total L1 T) #f #Hf
34 elim (lexs_dec R cfull HR … L1 L2 f)
35 /4 width=3 by lfxs_inv_frees, cfull_dec, ex2_intro, or_intror, or_introl/
36 qed-.
37
38 lemma lfxs_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
39                           lexs_frees_confluent … R1 cfull →
40                           ∀L1,L2,V. L1 ⦻*[R1, V] L2 → ∀I,T.
41                           ∃∃L. L1 ⦻*[R1, ②{I}V.T] L & L ⦻*[R2, V] L2.
42 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T
43 [ elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg
44   elim (frees_inv_bind … Hg) #y1 #y2 #H #_ #Hy
45 | elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg
46   elim (frees_inv_flat … Hg) #y1 #y2 #H #_ #Hy
47 ]
48 lapply(frees_mono … H … Hf) -H #H1
49 lapply (sor_eq_repl_back1 … Hy … H1) -y1 #Hy
50 lapply (sor_inv_sle_sn … Hy) -y2 #Hfg
51 elim (lexs_sle_split … HR1 HR2 … HL12 … Hfg) -HL12 #L #HL1 #HL2
52 lapply (sle_lexs_trans … HL1 … Hfg) // #H
53 elim (HR … Hf … H) -HR -Hf -H
54 /4 width=7 by sle_lexs_trans, ex2_intro/
55 qed-.
56
57 lemma lfxs_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
58                           lexs_frees_confluent … R1 cfull →
59                           ∀L1,L2,T. L1 ⦻*[R1, T] L2 → ∀I,V.
60                           ∃∃L. L1 ⦻*[R1, ⓕ{I}V.T] L & L ⦻*[R2, T] L2.
61 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V
62 elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg
63 elim (frees_inv_flat … Hg) #y1 #y2 #_ #H #Hy
64 lapply(frees_mono … H … Hf) -H #H2
65 lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy
66 lapply (sor_inv_sle_dx … Hy) -y1 #Hfg
67 elim (lexs_sle_split … HR1 HR2 … HL12 … Hfg) -HL12 #L #HL1 #HL2
68 lapply (sle_lexs_trans … HL1 … Hfg) // #H
69 elim (HR … Hf … H) -HR -Hf -H
70 /4 width=7 by sle_lexs_trans, ex2_intro/
71 qed-.
72
73 (* Main properties **********************************************************)
74
75 (* Basic_2A1: uses: llpx_sn_bind llpx_sn_bind_O *)
76 theorem lfxs_bind: ∀R,p,I,L1,L2,V1,V2,T.
77                    L1 ⦻*[R, V1] L2 → L1.ⓑ{I}V1 ⦻*[R, T] L2.ⓑ{I}V2 →
78                    L1 ⦻*[R, ⓑ{p,I}V1.T] L2.
79 #R #p #I #L1 #L2 #V1 #V2 #T * #f1 #HV #Hf1 * #f2 #HT #Hf2
80 elim (lexs_fwd_pair … Hf2) -Hf2 #Hf2 #_ elim (sor_isfin_ex f1 (⫱f2))
81 /3 width=7 by frees_fwd_isfin, frees_bind, lexs_join, isfin_tl, ex2_intro/
82 qed.
83
84 (* Basic_2A1: llpx_sn_flat *)
85 theorem lfxs_flat: ∀R,I,L1,L2,V,T.
86                    L1 ⦻*[R, V] L2 → L1 ⦻*[R, T] L2 →
87                    L1 ⦻*[R, ⓕ{I}V.T] L2.
88 #R #I #L1 #L2 #V #T * #f1 #HV #Hf1 * #f2 #HT #Hf2 elim (sor_isfin_ex f1 f2)
89 /3 width=7 by frees_fwd_isfin, frees_flat, lexs_join, ex2_intro/
90 qed.
91
92 theorem lfxs_conf: ∀R1,R2.
93                    lexs_frees_confluent R1 cfull →
94                    lexs_frees_confluent R2 cfull →
95                    R_confluent2_lfxs R1 R2 R1 R2 →
96                    ∀T. confluent2 … (lfxs R1 T) (lfxs R2 T).
97 #R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02
98 lapply (frees_mono … Hf1 … Hf) -Hf1 #Hf12
99 lapply (lexs_eq_repl_back … HL01 … Hf12) -f1 #HL01
100 elim (lexs_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ]
101 [ #L #HL1 #HL2
102   elim (HR1 … Hf … HL01) -HL01 #f1 #Hf1 #H1
103   elim (HR2 … Hf … HL02) -HL02 #f2 #Hf2 #H2
104   lapply (sle_lexs_trans … HL1 … H1) // -HL1 -H1 #HL1
105   lapply (sle_lexs_trans … HL2 … H2) // -HL2 -H2 #HL2
106   /3 width=5 by ex2_intro/
107 | #g #I #K0 #V0 #n #HLK0 #Hgf #V1 #HV01 #V2 #HV02 #K1 #HK01 #K2 #HK02
108   elim (frees_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0
109   lapply (sle_lexs_trans … HK01 … H0) // -HK01 #HK01
110   lapply (sle_lexs_trans … HK02 … H0) // -HK02 #HK02
111   elim (HR12 … HV01 … HV02 K1 … K2) /2 width=3 by ex2_intro/
112 ]
113 qed-.
114
115 (* Negated inversion lemmas *************************************************)
116
117 (* Basic_2A1: uses: nllpx_sn_inv_bind nllpx_sn_inv_bind_O *)
118 lemma lfnxs_inv_bind: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
119                       ∀p,I,L1,L2,V,T. (L1 ⦻*[R, ⓑ{p,I}V.T] L2 → ⊥) →
120                       (L1 ⦻*[R, V] L2 → ⊥) ∨ (L1.ⓑ{I}V ⦻*[R, T] L2.ⓑ{I}V → ⊥).
121 #R #HR #p #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
122 /4 width=2 by lfxs_bind, or_intror, or_introl/
123 qed-.
124
125 (* Basic_2A1: uses: nllpx_sn_inv_flat *)
126 lemma lfnxs_inv_flat: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) →
127                       ∀I,L1,L2,V,T. (L1 ⦻*[R, ⓕ{I}V.T] L2 → ⊥) →
128                       (L1 ⦻*[R, V] L2 → ⊥) ∨ (L1 ⦻*[R, T] L2 → ⊥).
129 #R #HR #I #L1 #L2 #V #T #H elim (lfxs_dec … HR L1 L2 V)
130 /4 width=1 by lfxs_flat, or_intror, or_introl/
131 qed-.