]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops.etc
- new component "s_transition" for the restored fqu and fquq
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_new / drops / drops.etc
1 lemma drop_split: ∀L1,L2,l,m2,s. ⬇[s, l, m2] L1 ≡ L2 → ∀m1. m1 ≤ m2 →
2                   ∃∃L. ⬇[s, l, m2 - m1] L1 ≡ L & ⬇[s, l, m1] L ≡ L2.
3 #L1 #L2 #l #m2 #s #H elim H -L1 -L2 -l -m2
4 [ #l #m2 #Hs #m1 #Hm12 @(ex2_intro … (⋆))
5   @drop_atom #H lapply (Hs H) -s #H destruct /2 width=1 by le_n_O_to_eq/
6 | #I #L1 #V #m1 #Hm1 lapply (le_n_O_to_eq … Hm1) -Hm1
7   #H destruct /2 width=3 by ex2_intro/
8 | #I #L1 #L2 #V #m2 #HL12 #IHL12 #m1 @(nat_ind_plus … m1) -m1
9   [ /3 width=3 by drop_drop, ex2_intro/
10   | -HL12 #m1 #_ #Hm12 lapply (le_plus_to_le_r … Hm12) -Hm12
11     #Hm12 elim (IHL12 … Hm12) -IHL12 >minus_plus_plus_l
12     #L #HL1 #HL2 elim (lt_or_ge (|L1|) (m2-m1)) #H0
13     [ elim (drop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct
14       elim (drop_inv_atom1 … HL2) -HL2 #H #_ destruct
15       @(ex2_intro … (⋆)) [ @drop_O1_ge normalize // ]
16       @drop_atom #H destruct
17     | elim (drop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by drop_drop, ex2_intro/
18     ]
19   ]
20 | #I #L1 #L2 #V1 #V2 #l #m2 #_ #HV21 #IHL12 #m1 #Hm12 elim (IHL12 … Hm12) -IHL12
21   #L #HL1 #HL2 elim (lift_split … HV21 l m1) -HV21 /3 width=5 by drop_skip, ex2_intro/
22 ]
23 qed-.
24
25 lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 0.
26 /2 width=5 by drop_inv_length_eq/ qed-.
27
28 fact drop_inv_FT_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 →
29                       ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → l = 0 →
30                       ⬇[Ⓕ, l, m] L1 ≡ K.ⓑ{I}V.
31 #L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m
32 [ #l #m #_ #J #K #W #H destruct
33 | #I #L #V #J #K #W #H destruct //
34 | #I #L1 #L2 #V #m #_ #IHL12 #J #K #W #H1 #H2 destruct
35   /3 width=1 by drop_drop/
36 | #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #J #K #W #_ #_ #H
37   elim (ysucc_inv_O_dx … H)
38 ]
39 qed-.
40
41 lemma drop_inv_FT: ∀I,L,K,V,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡
42 K.ⓑ{I}V.
43 /2 width=5 by drop_inv_FT_aux/ qed.
44
45 lemma drop_inv_gen: ∀I,L,K,V,s,m. ⬇[s, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡
46 K.ⓑ{I}V.
47 #I #L #K #V * /2 width=1 by drop_inv_FT/
48 qed-.
49
50 lemma drop_inv_T: ∀I,L,K,V,s,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[s, 0, m] L
51 ≡ K.ⓑ{I}V.
52 #I #L #K #V * /2 width=1 by drop_inv_FT/
53 qed-.