]> matita.cs.unibo.it Git - helm.git/blob - matita/contribs/POPLmark/Fsub/part1a_inversion3.ma
tagged 0.5.0-rc1
[helm.git] / matita / contribs / POPLmark / Fsub / part1a_inversion3.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 "Fsub/defn2.ma".
16
17 (*** Lemma A.1 (Reflexivity) ***)
18 theorem JS_Refl : ∀T,G.WFType G T → WFEnv G → G ⊢ T ⊴  T.
19 intros 3; elim H;
20  [1,2,3: autobatch
21  | apply SA_All;
22     [ autobatch
23     | intros; apply (H4 ? H6);
24        [ intro; apply H6; apply (fv_WFT ? ? ? (WFT_Forall ? ? ? H1 H3));
25          simplify; autobatch
26        | autobatch]]]
27 qed.
28
29 (*
30  * A slightly more general variant to lemma A.2.2, where weakening isn't
31  * defined as concatenation of any two disjoint environments, but as
32  * set inclusion.
33  *)
34
35 lemma JS_weakening : ∀G,T,U.G ⊢ T ⊴ U → ∀H.WFEnv H → incl ? G H → H ⊢ T ⊴ U.
36 intros 4; elim H;
37  [1,2,3,4: autobatch depth=4 width=4 size=7
38  | apply (SA_All ? ? ? ? ? (H2 ? H6 H7));intros;
39    apply H4
40      [ intro; autobatch
41      | apply WFE_cons; autobatch
42      | unfold;intros; elim (in_list_cons_case ? ? ? ? H9); destruct; autobatch]]
43 qed.
44
45 lemma JSubtype_inv:
46  ∀G:list bound.∀T1,T:Typ.
47   ∀P:list bound → Typ → Typ → Prop.
48    (∀t. WFEnv G → WFType G t → T=Top → P G t Top) →
49    (∀n. WFEnv G → n ∈ fv_env G → T=TFree n → P G (TFree n) (TFree n)) →
50    (∀n,t1,t.
51     (mk_bound true n t1) ∈ G → G ⊢ t1 ⊴ t → P G t1 t → T=t → P G (TFree n) T) →
52    (∀s1,s2,t1,t2. G ⊢ t1 ⊴ s1 → G ⊢ s2 ⊴ t2 → T=Arrow t1 t2 → P G (Arrow s1 s2) (Arrow t1 t2)) →
53    (∀s1,s2,t1,t2. G ⊢ t1 ⊴ s1 →
54     (∀X. ¬(X ∈ fv_env G) → (mk_bound true X t1)::G ⊢ subst_type_nat s2 (TFree X) O ⊴ subst_type_nat t2 (TFree X) O)
55       → T=Forall t1 t2 → P G (Forall s1 s2) (Forall t1 t2)) →
56     G ⊢ T1 ⊴ T → P G T1 T.
57  intros;
58  generalize in match (refl_eq ? T);
59  generalize in match (refl_eq ? G);
60  elim H5 in ⊢ (? ? ? % → ? ? ? % → %); destruct; 
61   [1,2,3,4: autobatch depth=10 width=10 size=8
62   | apply H4; first [assumption | autobatch]]
63 qed.
64
65 theorem narrowing:∀X,G,G1,U,P,M,N.
66   G1 ⊢ P ⊴ U → (∀G2,T.G2@G1 ⊢ U ⊴ T → G2@G1 ⊢ P ⊴ T) → G ⊢ M ⊴ N →
67   ∀l.G=l@(mk_bound true X U::G1) → l@(mk_bound true X P::G1) ⊢ M ⊴ N.
68 intros 10.elim H2; destruct;
69  [1,2,4: autobatch width=10 depth=10 size=8
70  | elim (decidable_eq_nat X n)
71     [apply (SA_Trans_TVar ? ? ? P); destruct;
72       [ autobatch
73       | rewrite > append_cons; apply H1;
74         lapply (WFE_bound_bound true X t1 U ? ? H3); destruct;
75           [1,3: autobatch
76           | rewrite < append_cons; autobatch
77           ]]
78     | apply (SA_Trans_TVar ? ? ? t1)
79       [ apply (lookup_env_extends ? ? ? ? ? ? ? ? ? ? H3);
80         intro; autobatch
81       | autobatch]]
82  | apply SA_All;
83     [ autobatch
84     | intros;
85       apply (H6 ? ? (mk_bound true X1 t2::l1))
86       [ rewrite > fv_env_extends; autobatch
87       | autobatch]]]
88 qed.
89
90 lemma JS_trans_prova: ∀T,G1.WFType G1 T →
91 ∀G,R,U.incl ? (fv_env G1) (fv_env G) → G ⊢ R ⊴ T → G ⊢ T ⊴ U → G ⊢ R ⊴ U.
92 intros 3;elim H;clear H; try autobatch;
93   [ apply (JSubtype_inv ? ? ? ? ? ? ? ? ? H3); intros; destruct; autobatch
94   | inversion H3; intros; destruct; assumption
95   |*: apply (JSubtype_inv ? ? ? ? ? ? ? ? ? H6); intros; destruct;
96     [1,3: autobatch
97     |*: inversion H7; intros; destruct;
98       [1,2: autobatch depth=4 width=4 size=9
99       | apply SA_Top
100          [ assumption
101          | apply WFT_Forall;
102             [ autobatch
103             | intros;lapply (H8 ? H11);
104               autobatch]]
105       | apply SA_All
106          [ autobatch
107          | intros;apply (H4 X);
108             [intro; autobatch;
109             |intro;  apply H13;apply H5; apply (WFT_to_incl ? ? ? H3);
110              assumption
111             |simplify;autobatch
112             |apply (narrowing X (mk_bound true X t::G) ? ? ? ? ? H9 ? ? [])
113                [intros;apply H2
114                   [unfold;intros;lapply (H5 ? H15);rewrite > fv_append;
115                    autobatch
116                   |apply (JS_weakening ? ? ? H9)
117                      [autobatch
118                      |unfold;intros;autobatch]
119                   |assumption]
120                |*:autobatch]
121             |autobatch]]]]]
122 qed.
123
124 theorem JS_trans : ∀G,T,U,V.G ⊢ T ⊴ U → G ⊢ U ⊴ V → G ⊢ T ⊴ V.
125 intros 5; autobatch.
126 qed.
127
128 theorem JS_narrow : ∀G1,G2,X,P,Q,T,U.
129                        (G2 @ (mk_bound true X Q :: G1)) ⊢ T ⊴ U → G1 ⊢ P ⊴ Q →
130                        (G2 @ (mk_bound true X P :: G1)) ⊢ T ⊴ U.
131 intros; apply (narrowing ? ? ? ? ? ? ? H1 ? H) [|autobatch]
132 intros;apply (JS_trans ? ? ? ? ? H2);apply (JS_weakening ? ? ? H1);
133      [autobatch|unfold;intros;autobatch]
134 qed.