]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/static/frees_fqus.ma
basic_2: stronger supclosure allows better inversion lemmas
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / frees_fqus.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/relocation/drops_weight.ma".
16 include "basic_2/s_computation/fqus_weight.ma".
17 include "basic_2/static/frees.ma".
18
19 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
20
21 (* Properties with star-iterated supclosure *********************************)
22
23 lemma frees_fqus_drops: ∀f1,G,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 →
24                         ∀L2,T2. ⦃G, L1, T1⦄ ⊐* ⦃G, L2, T2⦄ →
25                         ∀I,n. ⬇*[n] L1 ≡ L2.ⓑ{I}T2 →
26                         ∃∃f2. L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 & f2 ⊆ ⫱*[⫯n] f1.
27 #f1 #G #L1 #T1 #H elim H -f1 -L1 -T1
28 [ #f1 #J #Hf1 #L2 #T2 #H12 #I #n #HL12
29   elim (fqus_inv_atom1 … H12) -H12 #H1 #H2 #H3 destruct
30   lapply (drops_fwd_lw … HL12) -HL12 #HL12
31   elim (lt_le_false … HL12) -HL12 //
32 | #f1 #J #L1 #V1 #s #Hf1 #IH #L2 #T2 #H12
33   elim (fqus_inv_sort1 … H12) -H12 [ * | #H12 #I * ]
34   [ -IH -Hf1 #H1 #H2 #H3 #I #n #HL12 destruct
35     lapply (drops_fwd_lw … HL12) -HL12 #HL12
36     elim (lt_le_false … HL12) -HL12 //
37   | -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
38     #H destruct <(fqus_inv_refl_atom3 … H12) -H12 /2 width=3 by sle_refl, ex2_intro/
39   | -Hf1 #I #HL12 lapply (drops_inv_drop1 … HL12) -HL12
40     #HL12 elim (IH … H12 … HL12) -IH -H12 -HL12 /3 width=3 by ex2_intro/
41   ]
42 | #f1 #J #L1 #V1 #Hf1 #IH #L2 #T2 #H12
43   elim (fqus_inv_zero1 … H12) -H12 [ * | #H12 #I * ]
44   [ -IH -Hf1 #H1 #H2 #H3 #I #n #HL12 destruct
45     lapply (drops_fwd_lw … HL12) -HL12 #HL12
46     elim (lt_le_false … HL12) -HL12 //
47   | -IH -H12 #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
48      #H destruct /3 width=3 by sle_refl, ex2_intro/
49   | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
50     #HL12 elim (IH … H12 … HL12) -IH -H12 -HL12 /3 width=3 by ex2_intro/
51   ]
52 | #f1 #J #L1 #V1 #i #Hf1 #IH #L2 #T2 #H12
53   elim (fqus_inv_lref1 … H12) -H12 [ * | #H12 #I * ]
54   [ -IH -Hf1 #H1 #H2 #H3 #I #n #HL12 destruct
55     lapply (drops_fwd_lw … HL12) -HL12 #HL12
56     elim (lt_le_false … HL12) -HL12 //
57   | -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
58     #H destruct <(fqus_inv_refl_atom3 … H12) -H12 /2 width=3 by sle_refl, ex2_intro/
59   | -Hf1 #I #HL12 lapply (drops_inv_drop1 … HL12) -HL12
60     #HL12 elim (IH … H12 … HL12) -IH -H12 -HL12 /3 width=3 by ex2_intro/
61   ]
62 | #f1 #J #L1 #V1 #l #Hf1 #IH #L2 #T2 #H12
63   elim (fqus_inv_gref1 … H12) -H12 [ * | #H12 #I * ]
64   [ -IH -Hf1 #H1 #H2 #H3 #I #n #HL12 destruct
65     lapply (drops_fwd_lw … HL12) -HL12 #HL12
66     elim (lt_le_false … HL12) -HL12 //
67   | -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
68     #H destruct <(fqus_inv_refl_atom3 … H12) -H12 /2 width=3 by sle_refl, ex2_intro/
69   | -Hf1 #I #HL12 lapply (drops_inv_drop1 … HL12) -HL12
70     #HL12 elim (IH … H12 … HL12) -IH -H12 -HL12 /3 width=3 by ex2_intro/
71   ]
72 | #f1V #f1T #f1 #p #J #L1 #V #T #_ #_ #Hf1 #IHV #IHT #L2 #T2 #H12 #I #n #HL12
73   elim (fqus_inv_bind1 … H12) -H12 [ * |*: #H12 ]
74   [ -IHV -IHT -Hf1 #H1 #H2 #H3 destruct
75     lapply (drops_fwd_lw … HL12) -HL12 #HL12
76     elim (lt_le_false … HL12) -HL12 //
77   | -IHT elim (IHV … H12 … HL12) -IHV -H12 -HL12
78     /4 width=6 by sor_tls, sor_sle_sn, ex2_intro/
79   | -IHV elim (IHT … H12 I (⫯n)) -IHT -H12 /2 width=1 by drops_drop/ -HL12
80     <tls_xn /4 width=6 by ex2_intro, sor_tls, sor_sle_dx/
81   ]
82 | #f1V #f1T #f1 #J #L1 #V #T #_ #_ #Hf1 #IHV #IHT #L2 #T2 #H12 #I #n #HL12
83   elim (fqus_inv_flat1 … H12) -H12 [ * |*: #H12 ]
84   [ -IHV -IHT -Hf1 #H1 #H2 #H3 destruct
85     lapply (drops_fwd_lw … HL12) -HL12 #HL12
86     elim (lt_le_false … HL12) -HL12 //
87   | -IHT elim (IHV … H12 … HL12) -IHV -H12 -HL12
88     /4 width=6 by sor_tls, sor_sle_sn, ex2_intro/
89   | -IHV elim (IHT … H12 … HL12) -IHT -H12 -HL12
90     /4 width=6 by ex2_intro, sor_tls, sor_sle_dx/
91   ]
92 ]
93 qed-.