]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/static/lsubf_frees.ma
improved lsubf allowes to prove lsubf_frees_trans.
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lsubf_frees.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/lsubf.ma".
16
17 (* RESTRICTED REFINEMENT FOR CONTEXT-SENSITIVE FREE VARIABLES ***************)
18
19 (* Properties with context-sensitive free variables *************************)
20
21 lemma lsubf_frees_trans: ∀f2,L2,T. L2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f,L1. ⦃L1, f⦄ ⫃𝐅* ⦃L2, f2⦄ →
22                          ∃∃f1. L1 ⊢ 𝐅*⦃T⦄ ≡ f1 & f1 ⊆ f.
23 #f2 #L2 #T #H elim H -f2 -L2 -T
24 [ #f2 #I #Hf2 #f #L1 #H elim (lsubf_inv_atom2 … H) -H
25   #H #_ destruct /3 width=3 by frees_atom, sle_isid_sn, ex2_intro/
26 | #f2 #I #K2 #W #s #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
27   [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
28   destruct elim (IH … H12) -K2
29   /3 width=3 by frees_sort, sle_inv_tl_dx, ex2_intro/
30 | #f2 #I #K2 #W #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
31   [ #K1 #H elim (sle_inv_nx … H ??) -H [ <tl_next_rew |*: // ]
32     #g2 #_ #H1 #H12 #H2 destruct elim (IH … H12) -K2
33     /3 width=7 by frees_zero, sle_next, ex2_intro/
34   | #g #K1 #V #Hg #Hf #H elim (sle_inv_nx … H ??) -H [ <tl_next_rew |*: // ]
35     #g2 #_ #H1 #H12 #H2 #H3 destruct elim (IH … H12) -K2
36     #f1 #Hf1 elim (sor_isfin_ex … f1 g ??)
37     /4 width=7 by frees_fwd_isfin, frees_flat, frees_zero, sor_inv_sle, sle_next, ex2_intro/
38   ]
39 | #f2 #I #K2 #W #i #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
40   [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
41   destruct elim (IH … H12) -K2
42   /3 width=3 by frees_lref, sle_inv_tl_dx, ex2_intro/
43 | #f2 #I #K2 #W #l #_ #IH #f #L1 #H elim (lsubf_inv_pair2 … H) -H *
44   [ #K1 #_ #H12 #H | #g #K1 #V #Hg #Hf #_ #H12 #H1 #H2 ]
45   destruct elim (IH … H12) -K2
46   /3 width=3 by frees_gref, sle_inv_tl_dx, ex2_intro/
47 | #f2V #f2T #f2 #p #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f #L1 #H12
48   elim (IHV f L1) -IHV [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_sn/ ]
49   elim (IHT (⫯f) (L1.ⓑ{I}V)) -IHT [2: /4 width=4 by lsubf_sle_div, lsubf_pair_nn, sor_inv_sle_dx, sor_inv_tl_dx/ ]
50   -f2V -f2T -f2 -L2 #f1T #HT #Hf1T #f1V #HV #Hf1V elim (sor_isfin_ex … f1V (⫱f1T) ??)
51   /4 width=9 by frees_fwd_isfin, frees_bind, sor_inv_sle, sle_xn_tl, isfin_tl, ex2_intro/
52 | #f2V #f2T #f2 #I #L2 #V #T #_ #_ #Hf2 #IHV #IHT #f #L1 #H12
53   elim (IHV f L1) -IHV [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_sn/ ]
54   elim (IHT f L1) -IHT [2: /3 width=4 by lsubf_sle_div, sor_inv_sle_dx/ ]
55   -f2V -f2T -f2 -L2 #f1T #HT #Hf1T #f1V #HV #Hf1V elim (sor_isfin_ex … f1V f1T ??)
56   /3 width=7 by frees_fwd_isfin, frees_flat, sor_inv_sle, ex2_intro/
57 ]
58 qed-.