]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc/tc_lfxs.etc
e2f613a8783826950d610f687d31118f5aaa79b3
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc / tc_lfxs.etc
1 (* lfxs_pair_repl_dx is not generic enough **********************************)
2
3 fact tc_lfxs_pair_repl_dx_aux: ∀R. (∀L. reflexive … (R L)) →
4                                ∀I,L1,Y,V1,T. L1.ⓑ{I}V1 ⦻**[R, T] Y →
5                                ∀L2,V2. Y = L2.ⓑ{I}V2 →
6                                L1.ⓑ{I}V1 ⦻**[R, T] L2.ⓑ{I}V1.
7 #R #HR #I #L1 #Y #V1 #T #H elim H -Y
8 /3 width=6 by lfxs_pair_repl_dx, step, inj/
9 #Y #Y2 #HY #HY2 #IH #L2 #V2 #H destruct
10 elim (lfxs_fwd_dx … HY2) #L #V #H destruct
11 lapply (IH ???) [4: |*: // ] -IH #HL1
12 @(step … HL1) -HL1 //
13
14 /3 width=6 by lfxs_pair_repl_dx, step, inj/
15
16 lemma tc_lfxs_pair_repl_dx: ∀R. (∀L. reflexive … (R L)) →
17                             ∀I,L1,L2,V1,V2,T.
18                             L1.ⓑ{I}V1 ⦻**[R, T] L2.ⓑ{I}V2 →
19                             L1.ⓑ{I}V1 ⦻**[R, T] L2.ⓑ{I}V1.
20
21 (* does this really hold? ***************************************************)
22
23 lemma pippo1: ∀R. s_r_confluent1 … R (lfxs R) →
24               s_r_transitive … R (lfxs R) →
25               s_r_transitive … R (tc_lfxs R).
26 #R #H1R #H2R @s_r_trans_LTC2 @s_r_trans_LTC1 //
27 qed-.
28
29 definition s_r_confluent2: ∀A,B. relation2 (A→relation B) (B→relation A) ≝ λA,B,R1,R2.
30                            ∀L2,T1,T2. R1 L2 T1 T2 → ∀L1. R2 T1 L1 L2 → R2 T2 L1 L2.
31
32
33 axiom s_r_conf2_LTC2: ∀R. s_r_transitive … R (lfxs R) →
34                       s_r_confluent1 … R (lfxs R) →
35                       s_r_confluent2 … R (tc_lfxs R).
36 (*
37 #R #H1R #H2R #L2 #T1 #T2 #HT12 #L1 #H
38 @(TC_ind_dx ??????? H) -L1 /3 width=3 by inj/
39 #L1 #L #HL1 #HL2 #IH @(tc_lfxs_step_sn … IH) -IH //
40 *)
41
42 lemma tc_lfxs_inv_zero: ∀R. s_r_confluent1 … R (lfxs R) →
43                         s_r_transitive … R (lfxs R) →
44                         ∀Y1,Y2. Y1 ⦻**[R, #0] Y2 →
45                         (Y1 = ⋆ ∧ Y2 = ⋆) ∨
46                         ∃∃I,L1,L2,V1,V2. L1 ⦻**[R, V1] L2 & LTC … R L1 V1 V2 &
47                                          Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
48 #R #H1R #H2R #Y1 #Y2 #H elim H -Y2
49 [ #Y2 #H elim (lfxs_inv_zero … H) -H *
50   /4 width=9 by ex4_5_intro, inj, or_introl, or_intror, conj/
51 | #Y #Y2 #_ #H elim (lfxs_inv_zero … H) -H *
52   [ #H #H2 * * /3 width=9 by ex4_5_intro, or_introl, or_intror, conj/
53   | #I #L #L2 #V #V2 #HL2 #HV2 #H #H2 * *
54     [ #H1 #H0 destruct
55     | #I0 #L1 #L0 #V1 #V0 #HL10 #HV10 #H1 #H0 destruct
56       @or_intror @ex4_5_intro [6,7: |*: /width=7/ ]
57       [
58       | -HL2 @(trans_TC … HV10) @(pippo1 … HV2) // -V2 @pippo2 //
59       
60       lapply (H2R … HV2 … HL2) -HL2 #HL02
61       /4 width=8 by ex3_5_intro, step, or_intror/
62     ]
63   ]
64
65 qed-.
66
67 (* has no meaning without the former ****************************************)
68
69 lemma lfxs_inv_lref: ∀R,Y1,Y2,i. Y1 ⦻*[R, #⫯i] Y2 →
70                      (Y1 = ⋆ ∧ Y2 = ⋆) ∨
71                      ∃∃I,L1,L2,V1,V2. L1 ⦻*[R, #i] L2 &
72                                       Y1 = L1.ⓑ{I}V1 & Y2 = L2.ⓑ{I}V2.
73 #R #Y1 #Y2 #i * #f #H1 #H2 elim (frees_inv_lref … H1) -H1 *
74 [ #H #_ lapply (lexs_inv_atom1_aux … H2 H) -H2 /3 width=1 by or_introl, conj/
75 | #I1 #L1 #V1 #g #HV1 #HY1 #Hg elim (lexs_inv_push1_aux … H2 … HY1 Hg) -H2 -Hg
76   /4 width=8 by ex3_5_intro, ex2_intro, or_intror/
77 ]
78 qed-.
79
80 (****************************************************************************)
81
82 lemma lfxs_inv_zero_pair_sn: ∀R,I,Y2,L1,V1. L1.ⓑ{I}V1 ⦻*[R, #0] Y2 →
83                              ∃∃L2,V2. L1 ⦻*[R, V1] L2 & R L1 V1 V2 &
84                                       Y2 = L2.ⓑ{I}V2.
85 #R #I #Y2 #L1 #V1 #H elim (lfxs_inv_zero … H) -H *
86 [ #H destruct
87 | #J #Y1 #L2 #X1 #V2 #HV1 #HV12 #H1 #H2 destruct
88   /2 width=5 by ex3_2_intro/
89 ]
90 qed-.
91
92 lemma lfxs_inv_zero_pair_dx: ∀R,I,Y1,L2,V2. Y1 ⦻*[R, #0] L2.ⓑ{I}V2 →
93                              ∃∃L1,V1. L1 ⦻*[R, V1] L2 & R L1 V1 V2 &
94                                       Y1 = L1.ⓑ{I}V1.
95 #R #I #Y1 #L2 #V2 #H elim (lfxs_inv_zero … H) -H *
96 [ #_ #H destruct
97 | #J #L1 #Y2 #V1 #X2 #HV1 #HV12 #H1 #H2 destruct
98   /2 width=5 by ex3_2_intro/
99 ]
100 qed-.
101
102 lemma lfxs_inv_lref_pair_sn: ∀R,I,Y2,L1,V1,i. L1.ⓑ{I}V1 ⦻*[R, #⫯i] Y2
103
104                              ∃∃L2,V2. L1 ⦻*[R, #i] L2 & Y2 = L2.ⓑ{I}V2.
105 #R #I #Y2 #L1 #V1 #i #H elim (lfxs_inv_lref … H) -H *
106 [ #H destruct
107 | #J #Y1 #L2 #X1 #V2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/
108 ]
109 qed-.
110
111 lemma lfxs_inv_lref_pair_dx: ∀R,I,Y1,L2,V2,i. Y1 ⦻*[R, #⫯i] L2.ⓑ{I}V2
112
113                              ∃∃L1,V1. L1 ⦻*[R, #i] L2 & Y1 = L1.ⓑ{I}V1.
114 #R #I #Y1 #L2 #V2 #i #H elim (lfxs_inv_lref … H) -H *
115 [ #_ #H destruct
116 | #J #L1 #Y2 #V1 #X2 #Hi #H1 #H2 destruct /2 width=4 by ex2_2_intro/
117 ]
118 qed-.