]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_2A1/frsup/ssta_frsups.etc
update in binaries for λδ
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_2A1 / frsup / ssta_frsups.etc
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/unfold/frsups.ma".
16 include "basic_2/static/ssta.ma".
17
18 (* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************)
19
20 (* Advanced inversion lemmas ************************************************)
21
22 lemma ssta_inv_frsupp: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → ⦃L, U⦄ ⧁+ ⦃L, T⦄ → ⊥.
23 #h #g #L #T #U #l #H elim H -L -T -U -l
24 [ #L #k #l #_ #H
25   elim (frsupp_inv_atom1_frsups … H)
26 | #L #K #V #W #U #i #l #_ #_ #HWU #_ #H
27   elim (lift_frsupp_trans … (⋆) … H … HWU) -U #X #H
28   elim (lift_inv_lref2_be … H ? ?) -H //
29 | #L #K #W #V #U #i #l #_ #_ #HWU #_ #H
30   elim (lift_frsupp_trans … (⋆) … H … HWU) -U #X #H
31   elim (lift_inv_lref2_be … H ? ?) -H //
32 | #a #I #L #V #T #U #l #_ #IHTU #H
33   elim (frsupp_inv_bind1_frsups … H) -H #H [2: /4 width=4/ ] -IHTU
34   lapply (frsups_fwd_fw … H) -H normalize
35   <associative_plus <associative_plus #H
36   elim (le_plus_xySz_x_false … H)
37 | #L #V #T #U #l #_ #IHTU #H
38   elim (frsupp_inv_flat1_frsups … H) -H #H [2: /4 width=4/ ] -IHTU
39   lapply (frsups_fwd_fw … H) -H normalize
40   <associative_plus <associative_plus #H
41   elim (le_plus_xySz_x_false … H)
42 | /3 width=4/
43 ]
44 qed-.
45
46 fact ssta_inv_refl_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → T = U → ⊥.
47 #h #g #L #T #U #l #H elim H -L -T -U -l
48 [ #L #k #l #_ #H
49   lapply (next_lt h k) destruct -H -e0 (**) (* destruct: these premises are not erased *)
50   <e1 -e1 #H elim (lt_refl_false … H)
51 | #L #K #V #W #U #i #l #_ #_ #HWU #_ #H destruct
52   elim (lift_inv_lref2_be … HWU ? ?) -HWU //
53 | #L #K #W #V #U #i #l #_ #_ #HWU #_ #H destruct
54   elim (lift_inv_lref2_be … HWU ? ?) -HWU //
55 | #a #I #L #V #T #U #l #_ #IHTU #H destruct /2 width=1/
56 | #L #V #T #U #l #_ #IHTU #H destruct /2 width=1/
57 | #L #W #T #U #l #HTU #_ #H destruct
58   elim (ssta_inv_frsupp … HTU ?) -HTU /2 width=1/
59 ]
60 qed-.
61
62 lemma ssta_inv_refl: ∀h,g,T,L,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, T⦄ → ⊥.
63 /2 width=8 by ssta_inv_refl_aux/ qed-.
64
65 lemma ssta_inv_frsups: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → ⦃L, U⦄ ⧁* ⦃L, T⦄ → ⊥.
66 #h #g #L #T #U #L #HTU #H elim (frsups_inv_all … H) -H
67 [ * #_ #H destruct /2 width=6 by ssta_inv_refl/
68 | /2 width=8 by ssta_inv_frsupp/
69 ]
70 qed-.