]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/static/lfxs_main.ma
444d413ef4230f151039777d22a60aa2bd10f6a6
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lfxs_main.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/lfxs_lfxs.ma".
16 include "basic_2/static/frees_fqup.ma".
17 include "basic_2/static/frees_frees.ma".
18
19 (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
20
21 axiom frees_lexs_conf_sle: ∀RN,RP,f1,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 →
22                            ∀L2. L1 ⦻*[RN, RP, f1] L2 →
23                            ∃∃f2. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 & f2 ⊆ f1.
24
25 theorem lfxs_conf: ∀R. R_confluent2_lfxs R R R R →
26                    ∀T. confluent … (lfxs R T).
27 #R #H1R #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02
28 lapply (frees_mono … Hf1 … Hf) -Hf1 #Hf12
29 lapply (lexs_eq_repl_back … HL01 … Hf12) -f1 #HL01
30 elim (lexs_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ]
31 [ #L #HL1 #HL2
32   elim (frees_lexs_conf_sle … Hf … HL01) -HL01 #f1 #Hf1 #H1
33   elim (frees_lexs_conf_sle … Hf … HL02) -HL02 #f2 #Hf2 #H2
34   lapply (sle_lexs_trans … HL1 … H1) // -HL1 -H1 #HL1
35   lapply (sle_lexs_trans … HL2 … H2) // -HL2 -H2 #HL2
36   /3 width=5 by ex2_intro/
37 | #g #I #K0 #V0 #n #HLK0 #Hgf #V1 #HV01 #V2 #HV02 #K1 #HK01 #K2 #HK02
38   elim (frees_drops_next … Hf … HLK0 … Hgf) -Hf -HLK0 -Hgf #g0 #Hg0 #H0
39   lapply (sle_lexs_trans … HK01 … H0) // -HK01 #HK01
40   lapply (sle_lexs_trans … HK02 … H0) // -HK02 #HK02
41   elim (H1R … HV01 … HV02 K1 … K2) /2 width=3 by ex2_intro/
42 ]
43 qed-.
44
45 (*
46 lemma pippo: ∀R1,R2,RP1,RP2. R_confluent_lfxs R1 R2 RP1 RP2 →
47              lexs_confluent R1 R2 RP1 cfull RP2 cfull.
48 #R1 #R2 #RP1 #RP2 #HR #f #L0 #T0 #T1 #HT01 #T2 #HT02 #L1 #HL01 #L2
49 #HL02
50 *)