]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_lfeq.ma
- dependences on ceq and ceq_ext fixed
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / i_static / tc_lfxs_lfeq.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/static/lfeq.ma".
16 include "basic_2/static/lfxs_lfxs.ma".
17 include "basic_2/i_static/tc_lfxs_fqup.ma".
18
19 (*
20 axiom cext2_inv_LTC: ∀R,L,I1,I2. cext2 (LTC … R) L I1 I2 → LTC … (cext2 R) L I1 I2.
21
22 #R #L #I1 #I2 * -I1 -I2
23 [ /2 width=1 by ext2_unit, inj/
24 | #I #V1 #V2 #HV12 
25 *)
26
27   
28 (*
29 lemma pippo: ∀RN,RP. (∀L. reflexive … (RP L)) →
30              ∀f,L1,L2. L1 ⪤*[LTC … RN, RP, f] L2 →
31              TC … (lexs RN RP f) L1 L2.
32 #RN #RP #HRP #f #L1 #L2 #H elim H -f -L1 -L2
33 [ /2 width=1 by lexs_atom, inj/ ]
34 #f #I1 #I2 #L1 #L2 #HL12 #HI12 #IH 
35 [ @step [3: 
36 *)
37
38 (*
39 axiom lexs_frees_confluent_LTC_sn: ∀RN,RP. lexs_frees_confluent RN RP →
40                                    lexs_frees_confluent (LTC … RN) RP.
41
42 #RN #RP #HR #f1 #L1 #T #Hf1 #L2 #H  
43 *)
44 (* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***)
45
46 lemma pippo: ∀R. (∀L. reflexive … (R L)) →
47              (lexs_frees_confluent (cext2 R) cfull) →
48              ∀L1,L2,T. L1 ⪤**[R, T] L2 →
49              ∃∃L. L1 ⪤*[LTC … R, T] L & L ≡[T] L2.
50 #R #H1R #H2R #L1 #L2 #T #H
51 @(tc_lfxs_ind_sn … H1R … H) -H -L2
52 [ /4 width=5 by lfxs_refl, inj, ex2_intro/
53 | #L0 #L2 #_ #HL02 * #L * #f1 #Hf1 #HL1 #HL0
54   lapply (lexs_co ??? cfull … (cext2_inv_LTC R) … HL1) -HL1 // #HL1
55   lapply (lfeq_lfxs_trans … HL0 … HL02) -L0 #HL2
56   elim (lexs_frees_confluent_LTC_sn … H2R … Hf1 … HL1) #f2 #Hf2 #Hf21
57   lapply (lfxs_inv_frees … HL2 … Hf2) -HL2 #HL2
58   elim (lexs_sle_split … ceq_ext … HL2 … Hf21) -HL2
59   [ #L0 #HL0 #HL02
60   |*: /2 width=1 by ext2_refl/
61   ]
62   lapply (sle_lexs_trans … HL0 … Hf21) -Hf21 // #H
63   elim (H2R … Hf2 … H) -H #f0 #Hf0 #Hf02
64   lapply (sle_lexs_trans … HL02 … Hf02) -f2 // #HL02
65   @(ex2_intro … L0)
66   [ @(ex2_intro … Hf1)
67   | @(ex2_intro … HL02) //