1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "static_2/relocation/lifts_lifts_bind.ma".
16 include "static_2/relocation/drops_weight.ma".
18 (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************)
20 (* Main properties **********************************************************)
22 (* Basic_2A1: includes: drop_conf_ge drop_conf_be drop_conf_le *)
23 theorem drops_conf: ∀b1,f1,L1,L. ⇩*[b1,f1] L1 ≘ L →
24 ∀b2,f,L2. ⇩*[b2,f] L1 ≘ L2 →
25 ∀f2. f1 ⊚ f2 ≘ f → ⇩*[b2,f2] L ≘ L2.
26 #b1 #f1 #L1 #L #H elim H -f1 -L1 -L
27 [ #f1 #_ #b2 #f #L2 #HL2 #f2 #Hf12 elim (drops_inv_atom1 … HL2) -b1 -HL2
28 #H #Hf destruct @drops_atom
29 #H elim (after_inv_isid3 … Hf12) -Hf12 /2 width=1 by/
30 | #f1 #I1 #K1 #K #_ #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_nxx … Hf) -Hf [2,3: // ]
31 #g #Hg #H destruct /3 width=3 by drops_inv_drop1/
32 | #f1 #I1 #I #K1 #K #_ #HI1 #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*:// ]
33 #g2 #g #Hf #H1 #H2 destruct
34 [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, liftsb_div3/
35 | /4 width=3 by drops_inv_drop1, drops_drop/
40 (* Basic_1: was: drop1_trans *)
41 (* Basic_2A1: includes: drop_trans_ge drop_trans_le drop_trans_ge_comm
44 theorem drops_trans: ∀b1,f1,L1,L. ⇩*[b1,f1] L1 ≘ L →
45 ∀b2,f2,L2. ⇩*[b2,f2] L ≘ L2 →
46 ∀f. f1 ⊚ f2 ≘ f → ⇩*[b1∧b2,f] L1 ≘ L2.
47 #b1 #f1 #L1 #L #H elim H -f1 -L1 -L
48 [ #f1 #Hf1 #b2 #f2 #L2 #HL2 #f #Hf elim (drops_inv_atom1 … HL2) -HL2
49 #H #Hf2 destruct @drops_atom #H elim (andb_inv_true_dx … H) -H
50 #H1 #H2 lapply (after_isid_inv_sn … Hf ?) -Hf
51 /3 width=3 by isid_eq_repl_back/
52 | #f1 #I1 #K1 #K #_ #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_nxx … Hf) -Hf
53 /3 width=3 by drops_drop/
54 | #f1 #I1 #I #K1 #K #_ #HI1 #IH #b2 #f2 #L2 #HL2 #f #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*: // ]
55 #g2 #g #Hg #H1 #H2 destruct
56 [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, liftsb_trans/
57 | /4 width=3 by drops_inv_drop1, drops_drop/
62 theorem drops_conf_div: ∀f1,L,K. ⇩*[Ⓣ,f1] L ≘ K → ∀f2. ⇩*[Ⓣ,f2] L ≘ K →
63 𝐔⦃f1⦄ → 𝐔⦃f2⦄ → f1 ≡ f2.
64 #f1 #L #K #H elim H -f1 -L -K
65 [ #f1 #Hf1 #f2 #Hf2 elim (drops_inv_atom1 … Hf2) -Hf2
66 /3 width=1 by isid_inv_eq_repl/
67 | #f1 #I #L #K #Hf1 #IH #f2 elim (pn_split f2) *
68 #g2 #H2 #Hf2 #HU1 #HU2 destruct
69 [ elim (drops_inv_skip1 … Hf2) -IH -HU1 -Hf2 #Y2 #X2 #HY2 #_ #H destruct
70 lapply (drops_fwd_isid … HY2 ?) -HY2 /2 width=3 by isuni_inv_push/ -HU2
71 #H destruct elim (drops_inv_x_bind_xy … Hf1)
72 | /4 width=5 by drops_inv_drop1, isuni_inv_next, eq_next/
74 | #f1 #I1 #I2 #L #K #Hf1 #_ #IH #f2 elim (pn_split f2) *
75 #g2 #H2 #Hf2 #HU1 #HU2 destruct
76 [ elim (drops_inv_skip1 … Hf2) -Hf2 #Y2 #X2 #HY2 #_ #H destruct -Hf1
77 /4 width=5 by isuni_fwd_push, eq_push/
78 | lapply (drops_inv_drop1 … Hf2) -Hf2 -IH -HU2 #Hg2
79 lapply (drops_fwd_isid … Hf1 ?) -Hf1 /2 width=3 by isuni_inv_push/ -HU1
80 #H destruct elim (drops_inv_x_bind_xy … Hg2)
85 (* Advanced properties ******************************************************)
87 (* Basic_2A1: includes: drop_mono *)
88 lemma drops_mono: ∀b1,f,L,L1. ⇩*[b1,f] L ≘ L1 →
89 ∀b2,L2. ⇩*[b2,f] L ≘ L2 → L1 = L2.
90 #b1 #f #L #L1 lapply (after_isid_dx 𝐈𝐝 … f)
91 /3 width=8 by drops_conf, drops_fwd_isid/
94 lemma drops_inv_uni: ∀L,i. ⇩*[Ⓕ,𝐔❴i❵] L ≘ ⋆ → ∀I,K. ⇩*[i] L ≘ K.ⓘ{I} → ⊥.
96 lapply (drops_F … H2) -H2 #H2
97 lapply (drops_mono … H2 … H1) -L -i #H destruct
100 lemma drops_ldec_dec: ∀L,i. Decidable (∃∃K,W. ⇩*[i] L ≘ K.ⓛW).
101 #L #i elim (drops_F_uni L i) [| * * [ #I #K1 | * #W1 #K1 ] ]
102 [4: /3 width=3 by ex1_2_intro, or_introl/
103 |*: #H1L @or_intror * #K2 #W2 #H2L
104 lapply (drops_mono … H2L … H1L) -L #H destruct
108 (* Basic_2A1: includes: drop_conf_lt *)
109 lemma drops_conf_skip1: ∀b2,f,L,L2. ⇩*[b2,f] L ≘ L2 →
110 ∀b1,f1,I1,K1. ⇩*[b1,f1] L ≘ K1.ⓘ{I1} →
112 ∃∃I2,K2. L2 = K2.ⓘ{I2} &
113 ⇩*[b2,f2] K1 ≘ K2 & ⇧*[f2] I2 ≘ I1.
114 #b2 #f #L #L2 #H2 #b1 #f1 #I1 #K1 #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/
118 (* Basic_2A1: includes: drop_trans_lt *)
119 lemma drops_trans_skip2: ∀b1,f1,L1,L. ⇩*[b1,f1] L1 ≘ L →
120 ∀b2,f2,I2,K2. ⇩*[b2,f2] L ≘ K2.ⓘ{I2} →
122 ∃∃I1,K1. L1 = K1.ⓘ{I1} &
123 ⇩*[b1∧b2,f] K1 ≘ K2 & ⇧*[f] I2 ≘ I1.
124 #b1 #f1 #L1 #L #H1 #b2 #f2 #I2 #K2 #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/
129 (* Basic_2A1: includes: drops_conf_div *)
130 lemma drops_conf_div_bind: ∀f1,f2,I1,I2,L,K.
131 ⇩*[Ⓣ,f1] L ≘ K.ⓘ{I1} → ⇩*[Ⓣ,f2] L ≘ K.ⓘ{I2} →
132 𝐔⦃f1⦄ → 𝐔⦃f2⦄ → f1 ≡ f2 ∧ I1 = I2.
133 #f1 #f2 #I1 #I2 #L #K #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 conj/