]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/static_2/relocation/sex_tc.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / relocation / sex_tc.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/lib/star.ma".
16 include "static_2/relocation/sex.ma".
17
18 (* GENERIC ENTRYWISE EXTENSION OF CONTEXT-SENSITIVE REALTIONS FOR TERMS *****)
19
20 definition s_rs_transitive_isid: relation (relation3 lenv bind bind) ≝ λRN,RP.
21                                  ∀f. 𝐈⦃f⦄ → s_rs_transitive … RP (λ_.sex RN RP f).
22
23 (* Properties with transitive closure ***************************************)
24
25 lemma sex_tc_refl: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
26                    ∀f. reflexive … (TC … (sex RN RP f)).
27 /3 width=1 by sex_refl, TC_reflexive/ qed.
28
29 lemma sex_tc_next_sn: ∀RN,RP. c_reflexive … RN →
30                       ∀f,I2,L1,L2. TC … (sex RN RP f) L1 L2 → ∀I1. RN L1 I1 I2 → 
31                       TC … (sex RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}).
32 #RN #RP #HRN #f #I2 #L1 #L2 #H @(TC_ind_dx ??????? H) -L1
33 /3 width=3 by sex_next, TC_strap, inj/
34 qed.
35
36 lemma sex_tc_next_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
37                       ∀f,I1,I2,L1. (CTC … RN) L1 I1 I2 → ∀L2. L1 ⪤[RN, RP, f] L2 →
38                       TC … (sex RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}).
39 #RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2
40 /4 width=5 by sex_refl, sex_next, step, inj/
41 qed.
42
43 lemma sex_tc_push_sn: ∀RN,RP. c_reflexive … RP →
44                       ∀f,I2,L1,L2. TC … (sex RN RP f) L1 L2 → ∀I1. RP L1 I1 I2 → 
45                       TC … (sex RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}).
46 #RN #RP #HRP #f #I2 #L1 #L2 #H @(TC_ind_dx ??????? H) -L1
47 /3 width=3 by sex_push, TC_strap, inj/
48 qed.
49
50 lemma sex_tc_push_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
51                       ∀f,I1,I2,L1. (CTC … RP) L1 I1 I2 → ∀L2. L1 ⪤[RN, RP, f] L2 →
52                       TC … (sex RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}).
53 #RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2
54 /4 width=5 by sex_refl, sex_push, step, inj/
55 qed.
56
57 lemma sex_tc_inj_sn: ∀RN,RP,f,L1,L2. L1 ⪤[RN, RP, f] L2 → L1 ⪤[CTC … RN, RP, f] L2.
58 #RN #RP #f #L1 #L2 #H elim H -f -L1 -L2
59 /3 width=1 by sex_push, sex_next, inj/
60 qed.
61
62 lemma sex_tc_inj_dx: ∀RN,RP,f,L1,L2. L1 ⪤[RN, RP, f] L2 → L1 ⪤[RN, CTC … RP, f] L2.
63 #RN #RP #f #L1 #L2 #H elim H -f -L1 -L2
64 /3 width=1 by sex_push, sex_next, inj/
65 qed.
66
67 (* Main properties with transitive closure **********************************)
68
69 theorem sex_tc_next: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
70                      ∀f,I1,I2,L1. (CTC … RN) L1 I1 I2 → ∀L2. TC … (sex RN RP f) L1 L2 →
71                      TC … (sex RN RP (↑f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}).
72 #RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2
73 /4 width=5 by sex_tc_next_sn, sex_tc_refl, trans_TC/
74 qed.
75
76 theorem sex_tc_push: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
77                      ∀f,I1,I2,L1. (CTC … RP) L1 I1 I2 → ∀L2. TC … (sex RN RP f) L1 L2 →
78                      TC … (sex RN RP (⫯f)) (L1.ⓘ{I1}) (L2.ⓘ{I2}).
79 #RN #RP #HRN #HRP #f #I1 #I2 #L1 #H elim H -I2
80 /4 width=5 by sex_tc_push_sn, sex_tc_refl, trans_TC/
81 qed.
82
83 (* Basic_2A1: uses: TC_lpx_sn_ind *)
84 theorem sex_tc_step_dx: ∀RN,RP. s_rs_transitive_isid RN RP →
85                         ∀f,L1,L. L1 ⪤[RN, RP, f] L → 𝐈⦃f⦄ →
86                         ∀L2. L ⪤[RN, CTC … RP, f] L2 → L1⪤ [RN, CTC … RP, f] L2.
87 #RN #RP #HRP #f #L1 #L #H elim H -f -L1 -L
88 [ #f #_ #Y #H -HRP >(sex_inv_atom1 … H) -Y // ]
89 #f #I1 #I #L1 #L #HL1 #HI1 #IH #Hf #Y #H
90 [ elim (isid_inv_next … Hf) -Hf //
91 | lapply (isid_inv_push … Hf ??) -Hf [3: |*: // ] #Hf
92   elim (sex_inv_push1 … H) -H #I2 #L2 #HL2 #HI2 #H destruct
93   @sex_push [ /2 width=1 by/ ] -L2 -IH
94   @(TC_strap … HI1) -HI1
95   @(HRP … HL1) // (**) (* auto fails *)
96 ]
97 qed-.
98
99 (* Advanced properties ******************************************************)
100
101 lemma sex_tc_dx: ∀RN,RP. s_rs_transitive_isid RN RP →
102                  ∀f. 𝐈⦃f⦄ → ∀L1,L2. TC … (sex RN RP f) L1 L2 → L1 ⪤[RN, CTC … RP, f] L2.
103 #RN #RP #HRP #f #Hf #L1 #L2 #H @(TC_ind_dx ??????? H) -L1
104 /3 width=3 by sex_tc_step_dx, sex_tc_inj_dx/
105 qed.
106
107 (* Advanced inversion lemmas ************************************************)
108
109 lemma sex_inv_tc_sn: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
110                      ∀f,L1,L2. L1 ⪤[CTC … RN, RP, f] L2 → TC … (sex RN RP f) L1 L2.
111 #RN #RP #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2
112 /2 width=1 by sex_tc_next, sex_tc_push_sn, sex_atom, inj/
113 qed-.
114
115 lemma sex_inv_tc_dx: ∀RN,RP. c_reflexive … RN → c_reflexive … RP →
116                      ∀f,L1,L2. L1 ⪤[RN, CTC … RP, f] L2 → TC … (sex RN RP f) L1 L2.
117 #RN #RP #HRN #HRP #f #L1 #L2 #H elim H -f -L1 -L2
118 /2 width=1 by sex_tc_push, sex_tc_next_sn, sex_atom, inj/
119 qed-.