]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/i_static/tc_lfxs_lfeq.ma
64e38fc4892e6a85f12c2113e417babdb751489c
[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/syntax/ext2_tc.ma".
16 include "basic_2/relocation/lexs_tc.ma".
17 include "basic_2/relocation/lex.ma".
18 include "basic_2/static/lfeq_fqup.ma".
19 include "basic_2/static/lfxs_lfxs.ma".
20 include "basic_2/i_static/tc_lfxs_fqup.ma".
21
22 (* ITERATED EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ***)
23
24 lemma tc_lfxs_inv_lex_lfeq: ∀R. c_reflexive … R →
25                             (lexs_frees_confluent (cext2 R) cfull) →
26                             (∀f. 𝐈⦃f⦄ → s_rs_transitive … (cext2 R) (λ_.lexs cfull (cext2 R) f)) →
27                             ∀L1,L2,T. L1 ⪤**[R, T] L2 →
28                             ∃∃L. L1 ⪤[LTC … R] L & L ≡[T] L2.
29 #R #H1R #H2R #H3R #L1 #L2 #T #H
30 @(tc_lfxs_ind_sn … H1R … H) -H -L2
31 [ /4 width=3 by lfeq_refl, lex_refl, inj, ex2_intro/
32 | #L0 #L2 #_ #HL02 * #L * #f0 #Hf0 #HL1 #HL0
33   lapply (lfeq_lfxs_trans … HL0 … HL02) -L0 * #f1 #Hf1 #HL2
34   elim (lexs_sdj_split … ceq_ext … HL2 f0 ?) -HL2
35   [ #L0 #HL0 #HL02 |*: /2 width=1 by ext2_refl, sdj_isid_dx/ ]
36   lapply (lexs_sdj … HL0 f1 ?) /2 width=1 by sdj_isid_sn/ #H
37   elim (H2R … Hf1 … H) -H #f2 #Hf2 #Hf21
38   lapply (sle_lexs_trans … HL02 … Hf21) -f1 // #HL02
39   lapply (lexs_co ?? cfull (LTC … (cext2 R)) … HL1) -HL1 /2 width=1 by ext2_inv_tc/ #HL1  
40   lapply (lexs_inv_tc_dx … HL1) -HL1 /2 width=1 by ext2_refl/ #HL1
41   lapply (step ????? HL1 … HL0) -L #HL10
42   lapply (lexs_tc_dx … H3R … HL10) -HL10 // #HL10
43   lapply (lexs_co … cfull (cext2 (LTC … R)) … HL10) -HL10 /2 width=1 by ext2_tc/ #HL10
44   /3 width=5 by ex2_intro/
45 ]
46 qed-.