]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma
renaming in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / aaa_drops.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_drops.ma".
16 include "basic_2/s_computation/fqup_weight.ma".
17 include "basic_2/s_computation/fqup_drops.ma".
18 include "basic_2/static/aaa.ma".
19
20 (* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
21
22 (* Advanced properties ******************************************************)
23
24 (* Basic_2A1: was: aaa_lref *)
25 lemma aaa_lref_drops: ∀I,G,K,V,B,i,L. ⬇*[i] L ≘ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ⁝ B → ⦃G, L⦄ ⊢ #i ⁝ B.
26 #I #G #K #V #B #i elim i -i
27 [ #L #H lapply (drops_fwd_isid … H ?) -H //
28   #H destruct /2 width=1 by aaa_zero/
29 | #i #IH #L <uni_succ #H #HB lapply (drops_inv_bind2_isuni_next … H) -H // *
30   #Z #Y #HY #H destruct /3 width=1 by aaa_lref/
31 ]
32 qed.
33
34 (* Advanced inversion lemmas ************************************************)
35
36 (* Basic_2A1: was: aaa_inv_lref *)
37 lemma aaa_inv_lref_drops: ∀G,A,i,L. ⦃G, L⦄ ⊢ #i ⁝ A →
38                           ∃∃I,K,V. ⬇*[i] L ≘ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A.
39 #G #A #i elim i -i
40 [ #L #H elim (aaa_inv_zero … H) -H /3 width=5 by drops_refl, ex2_3_intro/
41 | #i #IH #L #H elim (aaa_inv_lref … H) -H
42   #I #K #H #HA destruct elim (IH … HA) -IH -HA /3 width=5 by drops_drop, ex2_3_intro/
43 ]
44 qed-.
45
46 (* Properties with generic slicing for local environments *******************)
47
48 (* Basic_2A1: includes: aaa_lift *)
49 (* Note: it should use drops_split_trans_pair2 *)
50 lemma aaa_lifts: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀b,f,L2. ⬇*[b, f] L2 ≘ L1 →
51                  ∀T2. ⬆*[f] T1 ≘ T2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
52 @(fqup_wf_ind_eq (Ⓣ)) #G0 #L0 #T0 #IH #G #L1 * *
53 [ #s #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -IH
54   lapply (aaa_inv_sort … H) -H #H destruct
55   >(lifts_inv_sort1 … HX) -HX //
56 | #i1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
57   elim (aaa_inv_lref_drops … H) -H #J #K1 #V1 #HLK1 #HA
58   elim (lifts_inv_lref1 … HX) -HX #i2 #Hf #H destruct
59   lapply (drops_trans … HL21 … HLK1 ??) -HL21 [1,2: // ] #H
60   elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK2 #HY
61   lapply (drops_tls_at … Hf … HY) -HY #HY -Hf
62   elim (drops_inv_skip2 … HY) -HY #Z #K2 #HK21 #HZ #H destruct
63   elim (liftsb_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct
64   /4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_gen/
65 | #l #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -f -IH
66   elim (aaa_inv_gref … H)
67 | #p * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
68   [ elim (aaa_inv_abbr … H) -H #B #HB #HA
69     elim (lifts_inv_bind1 …  HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
70     /5 width=9 by aaa_abbr, drops_skip, ext2_pair/
71   | elim (aaa_inv_abst … H) -H #B #A0 #HB #HA #H0
72     elim (lifts_inv_bind1 … HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
73     /5 width=8 by aaa_abst, drops_skip, ext2_pair/
74   ]
75 | * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
76   [ elim (aaa_inv_appl … H) -H #B #HB #HA
77     elim (lifts_inv_flat1 …  HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
78     /3 width=10 by aaa_appl/
79   | elim (aaa_inv_cast … H) -H #H1A #H2A
80     elim (lifts_inv_flat1 …  HX) -HX #V2 #T2 #HV12 #HT12 #H destruct
81     /3 width=8 by aaa_cast/
82   ]
83 ]
84 qed-.
85
86 (* Inversion lemmas with generic slicing for local environments *************)
87
88 (* Basic_2A1: includes: aaa_inv_lift *)
89 lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀b,f,L1. ⬇*[b, f] L2 ≘ L1 →
90                      ∀T1. ⬆*[f] T1 ≘ T2 → ⦃G, L1⦄ ⊢ T1 ⁝ A.
91 @(fqup_wf_ind_eq (Ⓣ)) #G0 #L0 #T0 #IH #G #L2 * *
92 [ #s #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -b -IH
93   lapply (aaa_inv_sort … H) -H #H destruct
94   >(lifts_inv_sort2 … HX) -HX //
95 | #i2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
96   elim (aaa_inv_lref_drops … H) -H #J #K2 #V2 #HLK2 #HA
97   elim (lifts_inv_lref2 … HX) -HX #i1 #Hf #H destruct
98   lapply (drops_split_div … HL21 (𝐔❴i1❵) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY
99   lapply (drops_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by after_uni_dx/ ] #HY
100   lapply (drops_tls_at … Hf … HY) -HY #HY -Hf
101   elim (drops_inv_skip1 … HY) -HY #Z #K1 #HK21 #HZ #H destruct
102   elim (liftsb_inv_pair_dx … HZ) -HZ #V1 #HV12 #H destruct
103   /4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_F/
104 | #l #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -IH -b -f
105   elim (aaa_inv_gref … H)
106 | #p * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
107   [ elim (aaa_inv_abbr … H) -H #B #HB #HA
108     elim (lifts_inv_bind2 …  HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
109     /5 width=9 by aaa_abbr, drops_skip, ext2_pair/
110   | elim (aaa_inv_abst … H) -H #B #A0 #HB #HA #H0
111     elim (lifts_inv_bind2 …  HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
112     /5 width=8 by aaa_abst, drops_skip, ext2_pair/
113   ]
114 | * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
115   [ elim (aaa_inv_appl … H) -H #B #HB #HA
116     elim (lifts_inv_flat2 …  HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
117     /3 width=10 by aaa_appl/
118   | elim (aaa_inv_cast … H) -H #H1A #H2A
119     elim (lifts_inv_flat2 …  HX) -HX #V1 #T1 #HV12 #HT12 #H destruct
120     /3 width=8 by aaa_cast/
121   ]
122 ]
123 qed-.