]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma
update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / lex_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/relocation/rtmap_id.ma".
16 include "basic_2/syntax/ext2_tc.ma".
17 include "basic_2/relocation/lexs_tc.ma".
18 include "basic_2/relocation/lex.ma".
19
20 alias symbol "subseteq" = "relation inclusion".
21
22 (* GENERIC EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS **************)
23
24 (* Inversion lemmas with transitive closure *********************************)
25
26 (* Basic_2A1: was: lpx_sn_LTC_TC_lpx_sn *)
27 lemma lex_inv_ltc: ∀R. c_reflexive … R →
28                    lex (LTC … R) ⊆ TC … (lex R).
29 #R #HR #L1 #L2 *
30 /5 width=11 by lexs_inv_tc_dx, lexs_co, ext2_inv_tc, ext2_refl, monotonic_TC, ex2_intro/
31 qed-.
32
33 lemma s_rs_transitive_lex_inv_isid: ∀R. s_rs_transitive … R (λ_.lex R) →
34                                     s_rs_transitive_isid cfull (cext2 R).
35 #R #HR #f #Hf #L2 #T1 #T2 #H #L1 #HL12
36 elim (ext2_tc … H) -H
37 [ /3 width=1 by ext2_inv_tc, ext2_unit/
38 | #I #V1 #V2 #HV12
39   @ext2_inv_tc @ext2_pair
40   @(HR … HV12) -HV12 /2 width=3 by ex2_intro/ (**) (* auto fails *)
41 ]
42 qed-.
43
44 (* Properties with transitive closure ***************************************)
45
46 (* Basic_2A1: was: TC_lpx_sn_inv_lpx_sn_LTC *)
47 lemma lex_ltc: ∀R. s_rs_transitive … R (λ_. lex R) →
48                TC … (lex R) ⊆ lex (LTC … R).
49 #R #HR #L1 #L2 #HL12
50 lapply (monotonic_TC … (lexs cfull (cext2 R) 𝐈𝐝) … HL12) -HL12
51 [ #L1 #L2 * /3 width=3 by lexs_eq_repl_fwd, eq_id_inv_isid/
52 | /5 width=9 by s_rs_transitive_lex_inv_isid, lexs_tc_dx, lexs_co, ext2_tc, ex2_intro/
53 ]
54 qed-.
55