]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops.etc
theory of generic slicing almost completed ....
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_new / drops / drops.etc
1 lemma drop_refl_atom_O2: ∀s,l. ⬇[s, l, O] ⋆ ≡ ⋆.
2 /2 width=1 by drop_atom/ qed.
3
4 (* Basic_1: was by definition: drop_refl *)
5 lemma drop_refl: ∀L,l,s. ⬇[s, l, 0] L ≡ L.
6 #L elim L -L //
7 #L #I #V #IHL #l #s @(nat_ind_plus … l) -l /2 width=1 by drop_pair, drop_skip/
8 qed.
9
10 lemma drop_split: ∀L1,L2,l,m2,s. ⬇[s, l, m2] L1 ≡ L2 → ∀m1. m1 ≤ m2 →
11                   ∃∃L. ⬇[s, l, m2 - m1] L1 ≡ L & ⬇[s, l, m1] L ≡ L2.
12 #L1 #L2 #l #m2 #s #H elim H -L1 -L2 -l -m2
13 [ #l #m2 #Hs #m1 #Hm12 @(ex2_intro … (⋆))
14   @drop_atom #H lapply (Hs H) -s #H destruct /2 width=1 by le_n_O_to_eq/
15 | #I #L1 #V #m1 #Hm1 lapply (le_n_O_to_eq … Hm1) -Hm1
16   #H destruct /2 width=3 by ex2_intro/
17 | #I #L1 #L2 #V #m2 #HL12 #IHL12 #m1 @(nat_ind_plus … m1) -m1
18   [ /3 width=3 by drop_drop, ex2_intro/
19   | -HL12 #m1 #_ #Hm12 lapply (le_plus_to_le_r … Hm12) -Hm12
20     #Hm12 elim (IHL12 … Hm12) -IHL12 >minus_plus_plus_l
21     #L #HL1 #HL2 elim (lt_or_ge (|L1|) (m2-m1)) #H0
22     [ elim (drop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct
23       elim (drop_inv_atom1 … HL2) -HL2 #H #_ destruct
24       @(ex2_intro … (⋆)) [ @drop_O1_ge normalize // ]
25       @drop_atom #H destruct
26     | elim (drop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by drop_drop, ex2_intro/
27     ]
28   ]
29 | #I #L1 #L2 #V1 #V2 #l #m2 #_ #HV21 #IHL12 #m1 #Hm12 elim (IHL12 … Hm12) -IHL12
30   #L #HL1 #HL2 elim (lift_split … HV21 l m1) -HV21 /3 width=5 by drop_skip, ex2_intro/
31 ]
32 qed-.
33
34 (* Basic_2A1: includes: drop_split *)
35 lemma drops_split_trans: ∀L1,L2,t. ⬇*[t] L1 ≡ L2 → ∀t1,t2. t1 ⊚ t2 ≡ t →
36                          ∃∃L. ⬇*[t1] L1 ≡ L & ⬇*[t2] L ≡ L2.
37 #L1 #L2 #t #H elim H -L1 -L2 -t
38 [ #t1 #t2 #H elim (after_inv_empty3 … H) -H
39   /2 width=3 by ex2_intro, drops_atom/
40 | #I #L1 #L2 #V #t #HL12 #IHL12 #t1 #t2 #H elim (after_inv_false3 … H) -H *
41   [ #tl1 #tl2 #H1 #H2 #Ht destruct elim (IHL12 … Ht) -t
42     #tl #H1 #H2
43   | #tl1 #H #Ht destruct elim (IHL12 … Ht) -t
44     /3 width=5 by ex2_intro, drops_drop/
45   ]
46 | #I #L1 #L2 #V1 #V2 #t #_ #HV21 #IHL12 #t1 #t2 #H elim (after_inv_true3 … H) -H
47   #tl1 #tl2 #H1 #H2 #Ht destruct elim (lifts_split_trans … HV21 … Ht) -HV21
48   elim (IHL12 … Ht) -t /3 width=5 by ex2_intro, drops_skip/
49 ]
50 qed-.
51
52 lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 0.
53 /2 width=5 by drop_inv_length_eq/ qed-.
54
55 fact drop_inv_FT_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 →
56                       ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → l = 0 →
57                       ⬇[Ⓕ, l, m] L1 ≡ K.ⓑ{I}V.
58 #L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m
59 [ #l #m #_ #J #K #W #H destruct
60 | #I #L #V #J #K #W #H destruct //
61 | #I #L1 #L2 #V #m #_ #IHL12 #J #K #W #H1 #H2 destruct
62   /3 width=1 by drop_drop/
63 | #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #J #K #W #_ #_ #H
64   elim (ysucc_inv_O_dx … H)
65 ]
66 qed-.
67
68 lemma drop_inv_FT: ∀I,L,K,V,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡
69 K.ⓑ{I}V.
70 /2 width=5 by drop_inv_FT_aux/ qed.
71
72 lemma drop_inv_gen: ∀I,L,K,V,s,m. ⬇[s, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡
73 K.ⓑ{I}V.
74 #I #L #K #V * /2 width=1 by drop_inv_FT/
75 qed-.
76
77 lemma drop_inv_T: ∀I,L,K,V,s,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[s, 0, m] L
78 ≡ K.ⓑ{I}V.
79 #I #L #K #V * /2 width=1 by drop_inv_FT/
80 qed-.