]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma
- advances towards strong normalization
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / drops_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/lifts_lifts.ma".
16 include "basic_2/relocation/drops_weight.ma".
17
18 (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
19
20 (* Main properties on generic relocation ************************************)
21
22 lemma d_liftable2_sn_bi: ∀R. d_liftable2_sn R → d_liftable2_bi R.
23 #R #HR #K #T1 #T2 #HT12 #b #f #L #HLK #U1 #HTU1 #U2 #HTU2
24 elim (HR … HT12 … HLK … HTU1) -HR -K -T1 #X #HTX #HUX
25 <(lifts_mono … HTX … HTU2) -T2 -U2 -b -f //
26 qed-.
27
28 lemma d_deliftable2_sn_bi: ∀R. d_deliftable2_sn R → d_deliftable2_bi R.
29 #R #HR #L #U1 #U2 #HU12 #b #f #K #HLK #T1 #HTU1 #T2 #HTU2
30 elim (HR … HU12 … HLK … HTU1) -HR -L -U1 #X #HUX #HTX
31 <(lifts_inj … HUX … HTU2) -U2 -T2 -b -f //
32 qed-.
33
34 (* Main properties **********************************************************)
35
36 (* Basic_2A1: includes: drop_conf_ge drop_conf_be drop_conf_le *)
37 theorem drops_conf: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
38                     ∀b2,f,L2. ⬇*[b2, f] L1 ≡ L2 →
39                     ∀f2. f1 ⊚ f2 ≡ f → ⬇*[b2, f2] L ≡ L2.
40 #b1 #f1 #L1 #L #H elim H -f1 -L1 -L
41 [ #f1 #_ #b2 #f #L2 #HL2 #f2 #Hf12 elim (drops_inv_atom1 … HL2) -b1 -HL2
42   #H #Hf destruct @drops_atom
43   #H elim (after_inv_isid3 … Hf12) -Hf12 /2 width=1 by/
44 | #f1 #I #K1 #K #V1 #_ #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_nxx … Hf) -Hf [2,3: // ]
45   #g #Hg #H destruct /3 width=3 by drops_inv_drop1/
46 | #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*:// ]
47   #g2 #g #Hf #H1 #H2 destruct
48   [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_div3/
49   | /4 width=3 by drops_inv_drop1, drops_drop/
50   ]
51 ]
52 qed-.
53
54 (* Basic_1: was: drop1_trans *)
55 (* Basic_2A1: includes: drop_trans_ge drop_trans_le drop_trans_ge_comm 
56                         drops_drop_trans
57 *)
58 theorem drops_trans: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
59                      ∀b2,f2,L2. ⬇*[b2, f2] L ≡ L2 →
60                      ∀f. f1 ⊚ f2 ≡ f → ⬇*[b1∧b2, f] L1 ≡ L2.
61 #b1 #f1 #L1 #L #H elim H -f1 -L1 -L
62 [ #f1 #Hf1 #b2 #f2 #L2 #HL2 #f #Hf elim (drops_inv_atom1 … HL2) -HL2
63   #H #Hf2 destruct @drops_atom #H elim (andb_inv_true_dx … H) -H
64   #H1 #H2 lapply (after_isid_inv_sn … Hf ?) -Hf
65   /3 width=3 by isid_eq_repl_back/
66 | #f1 #I #K1 #K #V1 #_ #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_nxx … Hf) -Hf
67   /3 width=3 by drops_drop/
68 | #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*: // ]
69   #g2 #g #Hg #H1 #H2 destruct
70   [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_trans/
71   | /4 width=3 by drops_inv_drop1, drops_drop/
72   ]
73 ]
74 qed-.
75
76 theorem drops_conf_div: ∀f1,L,K. ⬇*[Ⓣ,f1] L ≡ K → ∀f2. ⬇*[Ⓣ,f2] L ≡ K →
77                         𝐔⦃f1⦄ → 𝐔⦃f2⦄ → f1 ≗ f2.
78 #f1 #L #K #H elim H -f1 -L -K
79 [ #f1 #Hf1 #f2 #Hf2 elim (drops_inv_atom1 … Hf2) -Hf2
80   /3 width=1 by isid_inv_eq_repl/
81 | #f1 #I #L #K #V #Hf1 #IH #f2 elim (pn_split f2) *
82   #g2 #H2 #Hf2 #HU1 #HU2 destruct
83   [ elim (drops_inv_skip1 … Hf2) -IH -HU1 -Hf2 #Y2 #X2 #HY2 #_ #H destruct
84     lapply (drops_fwd_isid … HY2 ?) -HY2 /2 width=3 by isuni_inv_push/ -HU2
85     #H destruct elim (drops_inv_x_pair_xy … Hf1)
86   | /4 width=5 by drops_inv_drop1, isuni_inv_next, eq_next/
87   ]
88 | #f1 #I #L #K #V #W #Hf1 #_ #IH #f2 elim (pn_split f2) *
89   #g2 #H2 #Hf2 #HU1 #HU2 destruct
90   [ elim (drops_inv_skip1 … Hf2) -Hf2 #Y2 #X2 #HY2 #_ #H destruct -Hf1
91     /4 width=5 by isuni_fwd_push, eq_push/
92   | lapply (drops_inv_drop1 … Hf2) -Hf2 -IH -HU2 #Hg2
93     lapply (drops_fwd_isid … Hf1 ?) -Hf1 /2 width=3 by isuni_inv_push/ -HU1
94     #H destruct elim (drops_inv_x_pair_xy … Hg2)
95   ]
96 ]
97 qed-.
98
99 (* Advanced properties ******************************************************)
100
101 (* Basic_2A1: includes: drop_mono *)
102 lemma drops_mono: ∀b1,f,L,L1. ⬇*[b1, f] L ≡ L1 →
103                   ∀b2,L2. ⬇*[b2, f] L ≡ L2 → L1 = L2.
104 #b1 #f #L #L1 lapply (after_isid_dx 𝐈𝐝 … f)
105 /3 width=8 by drops_conf, drops_fwd_isid/
106 qed-.
107
108 (* Basic_2A1: includes: drop_conf_lt *)
109 lemma drops_conf_skip1: ∀b2,f,L,L2. ⬇*[b2, f] L ≡ L2 →
110                         ∀b1,f1,I,K1,V1. ⬇*[b1, f1] L ≡ K1.ⓑ{I}V1 →
111                         ∀f2. f1 ⊚ ↑f2 ≡ f →
112                         ∃∃K2,V2. L2 = K2.ⓑ{I}V2 &
113                                  ⬇*[b2, f2] K1 ≡ K2 & ⬆*[f2] V2 ≡ V1.
114 #b2 #f #L #L2 #H2 #b1 #f1 #I #K1 #V1 #H1 #f2 #Hf lapply (drops_conf … H1 … H2 … Hf) -L -Hf
115 #H elim (drops_inv_skip1 … H) -H /2 width=5 by ex3_2_intro/
116 qed-.
117
118 (* Basic_2A1: includes: drop_trans_lt *)
119 lemma drops_trans_skip2: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
120                          ∀b2,f2,I,K2,V2. ⬇*[b2, f2] L ≡ K2.ⓑ{I}V2 →
121                          ∀f. f1 ⊚ f2 ≡ ↑f →
122                          ∃∃K1,V1. L1 = K1.ⓑ{I}V1 &
123                                   ⬇*[b1∧b2, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1.
124 #b1 #f1 #L1 #L #H1 #b2 #f2 #I #K2 #V2 #H2 #f #Hf
125 lapply (drops_trans … H1 … H2 … Hf) -L -Hf
126 #H elim (drops_inv_skip2 … H) -H /2 width=5 by ex3_2_intro/
127 qed-.
128
129 (* Basic_2A1: includes: drops_conf_div *)
130 lemma drops_conf_div_pair: ∀f1,f2,I1,I2,L,K,V1,V2.
131                            ⬇*[Ⓣ,f1] L ≡ K.ⓑ{I1}V1 → ⬇*[Ⓣ,f2] L ≡ K.ⓑ{I2}V2 →
132                            𝐔⦃f1⦄ → 𝐔⦃f2⦄ → ∧∧ f1 ≗ f2 & I1 = I2 & V1 = V2.
133 #f1 #f2 #I1 #I2 #L #K #V1 #V2 #Hf1 #Hf2 #HU1 #HU2
134 lapply (drops_isuni_fwd_drop2 … Hf1) // #H1
135 lapply (drops_isuni_fwd_drop2 … Hf2) // #H2
136 lapply (drops_conf_div … H1 … H2 ??) /2 width=3 by isuni_next/ -H1 -H2 -HU1 -HU2 #H
137 lapply (eq_inv_nn … H ????) -H [5: |*: // ] #H12
138 lapply (drops_eq_repl_back … Hf1 … H12) -Hf1 #H0
139 lapply (drops_mono … H0 … Hf2) -L #H
140 destruct /2 width=1 by and3_intro/
141 qed-.
142
143 lemma drops_inv_uni: ∀L,i. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ → ∀I,K,V. ⬇*[i] L ≡ K.ⓑ{I}V → ⊥.
144 #L #i #H1 #I #K #V #H2
145 lapply (drops_F … H2) -H2 #H2
146 lapply (drops_mono … H2 … H1) -L -i #H destruct
147 qed-.