]> matita.cs.unibo.it Git - helm.git/blob - matita/contribs/POPLmark/Fsub/part1a_inversion.ma
tagged 0.5.0-rc1
[helm.git] / matita / contribs / POPLmark / Fsub / part1a_inversion.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 set "baseuri" "cic:/matita/Fsub/part1a_inversion/".
16 include "Fsub/defn.ma".
17
18 (*** Lemma A.1 (Reflexivity) ***)
19 theorem JS_Refl : ∀T,G.WFType G T → WFEnv G → G ⊢ T ⊴  T.
20 intros 3.elim H
21   [apply SA_Refl_TVar [apply H2|assumption]
22   |apply SA_Top [assumption|apply WFT_Top]
23   |apply (SA_Arrow ? ? ? ? ? (H2 H5) (H4 H5))
24   |apply (SA_All ? ? ? ? ? (H2 H5));intros;apply (H4 ? H6)
25      [intro;apply H6;apply (fv_WFT ? ? ? (WFT_Forall ? ? ? H1 H3));
26       simplify;autobatch
27      |autobatch]]
28 qed.
29
30 (*
31  * A slightly more general variant to lemma A.2.2, where weakening isn't
32  * defined as concatenation of any two disjoint environments, but as
33  * set inclusion.
34  *)
35
36 lemma JS_weakening : ∀G,T,U.G ⊢ T ⊴ U → ∀H.WFEnv H → incl ? G H → H ⊢ T ⊴ U.
37 intros 4;elim H
38   [apply (SA_Top ? ? H4);apply (WFT_env_incl ? ? H2 ? (incl_bound_fv ? ? H5))
39   |apply (SA_Refl_TVar ? ? H4);apply (incl_bound_fv ? ? H5 ? H2)
40   |apply (SA_Trans_TVar ? ? ? ? ? (H3 ? H5 H6));apply H6;assumption
41   |apply (SA_Arrow ? ? ? ? ? (H2 ? H6 H7) (H4 ? H6 H7))
42   |apply (SA_All ? ? ? ? ? (H2 ? H6 H7));intros;apply H4
43      [unfold;intro;apply H8;apply (incl_bound_fv ? ? H7 ? H9)
44      |apply (WFE_cons ? ? ? ? H6 H8);autobatch
45      |unfold;intros;inversion H9;intros
46         [destruct H11;apply in_list_head
47         |destruct H13;apply in_list_cons;apply (H7 ? H10)]]]
48 qed.
49
50 theorem narrowing:∀X,G,G1,U,P,M,N.
51   G1 ⊢ P ⊴ U → (∀G2,T.G2@G1 ⊢ U ⊴ T → G2@G1 ⊢ P ⊴ T) → G ⊢ M ⊴ N →
52   ∀l.G=l@(mk_bound true X U::G1) → l@(mk_bound true X P::G1) ⊢ M ⊴ N.
53 intros 10.elim H2
54   [apply SA_Top
55     [rewrite > H5 in H3;
56      apply (WFE_Typ_subst ? ? ? ? ? ? ? H3 (JS_to_WFT1 ? ? ? H))
57     |rewrite > H5 in H4;apply (WFT_env_incl ? ? H4);apply incl_fv_env]
58   |apply SA_Refl_TVar
59     [rewrite > H5 in H3;apply (WFE_Typ_subst ? ? ? ? ? ? ? H3);
60      apply (JS_to_WFT1 ? ? ? H)
61     |rewrite > H5 in H4;rewrite < fv_env_extends;apply H4]
62   |elim (decidable_eq_nat X n)
63     [apply (SA_Trans_TVar ? ? ? P)
64       [rewrite < H7;elim l1;simplify
65         [constructor 1|constructor 2;assumption]
66       |rewrite > append_cons;apply H1;
67        lapply (WFE_bound_bound true n t1 U ? ? H3)
68         [apply (JS_to_WFE ? ? ? H4)
69         |rewrite < Hletin;rewrite < append_cons;apply (H5 ? H6)
70         |rewrite < H7;rewrite > H6;elim l1;simplify
71           [constructor 1|constructor 2;assumption]]]
72     |apply (SA_Trans_TVar ? ? ? t1)
73       [rewrite > H6 in H3;apply (lookup_env_extends ? ? ? ? ? ? ? ? ? ? H3);
74        unfold;intro;apply H7;symmetry;assumption
75       |apply (H5 ? H6)]]
76   |apply (SA_Arrow ? ? ? ? ? (H4 ? H7) (H6 ? H7))
77   |apply (SA_All ? ? ? ? ? (H4 ? H7));intros;
78    apply (H6 ? ? (mk_bound true X1 t2::l1))
79       [rewrite > H7;rewrite > fv_env_extends;apply H8
80       |simplify;rewrite < H7;reflexivity]]
81 qed.
82
83 lemma JSubtype_Arrow_inv:
84  ∀G:list bound.∀T1,T2,T3:Typ.
85   ∀P:list bound → Typ → Prop.
86    (∀n,t1.
87     (mk_bound true n t1) ∈ G → G ⊢ t1 ⊴ (Arrow T2 T3) → P G t1 → P G (TFree n)) →
88    (∀t,t1. G ⊢ T2 ⊴ t → G ⊢ t1 ⊴ T3 → P G (Arrow t t1)) →
89     G ⊢ T1 ⊴ (Arrow T2 T3) → P G T1.
90  intros;
91  generalize in match (refl_eq ? (Arrow T2 T3));
92  generalize in match (refl_eq ? G);
93  elim H2 in ⊢ (? ? ? % → ? ? ? % → %);
94   [1,2: destruct H6
95   |5: destruct H8
96   | lapply (H5 H6 H7); destruct; clear H5;
97     apply H;
98     assumption
99   | destruct;
100     clear H4 H6;
101     apply H1;
102     assumption
103   ]
104 qed.
105
106 lemma JSubtype_Forall_inv:
107  ∀G:list bound.∀T1,T2,T3:Typ.
108   ∀P:list bound → Typ → Prop.
109    (∀n,t1.
110     (mk_bound true n t1) ∈ G → G ⊢ t1 ⊴ (Forall T2 T3) → P G t1 → P G (TFree n)) →
111    (∀t,t1. G ⊢ T2 ⊴ t → (∀X. ¬(X ∈ fv_env G) → (mk_bound true X T2)::G ⊢ subst_type_nat t1 (TFree X) O ⊴ subst_type_nat T3 (TFree X) O)
112      → P G (Forall t t1)) →
113       G ⊢ T1 ⊴ (Forall T2 T3) → P G T1.
114  intros;
115  generalize in match (refl_eq ? (Forall T2 T3));
116  generalize in match (refl_eq ? G);
117  elim H2 in ⊢ (? ? ? % → ? ? ? % → %);
118   [1,2: destruct H6
119   |4: destruct H8
120   | lapply (H5 H6 H7); destruct; clear H5;
121     apply H;
122     assumption
123   | destruct;
124     clear H4 H6;
125     apply H1;
126     assumption
127   ]
128 qed.
129
130
131 lemma JS_trans_prova: ∀T,G1.WFType G1 T →
132 ∀G,R,U.incl ? (fv_env G1) (fv_env G) → G ⊢ R ⊴ T → G ⊢ T ⊴ U → G ⊢ R ⊴ U.
133 intros 3;elim H;clear H; try autobatch;
134   [rewrite > (JSubtype_Top ? ? H3);autobatch
135   |apply (JSubtype_Arrow_inv ? ? ? ? ? ? ? H6); intros;
136     [ autobatch
137     | inversion H7;intros; destruct; autobatch depth=4 width=4 size=9]
138   |apply (JSubtype_Forall_inv ? ? ? ? ? ? ? H6); intros;
139     [ autobatch
140     | inversion H7;intros; destruct;
141        [ apply SA_Top
142           [ assumption
143           | apply WFT_Forall;
144              [ autobatch
145              | intros;lapply (H8 ? H11);
146                autobatch]]
147        | apply SA_All
148           [ autobatch
149           | intros;apply (H4 X);
150              [intro;apply H13;apply H5;assumption
151                  |intro;apply H13;apply H5;apply (WFT_to_incl ? ? ? H3);
152                   assumption
153                  |simplify;autobatch
154                  |apply (narrowing X (mk_bound true X t::G) ? ? ? ? ? H9 ? ? [])
155                     [intros;apply H2
156                        [unfold;intros;lapply (H5 ? H15);rewrite > fv_append;
157                         autobatch
158                        |apply (JS_weakening ? ? ? H9)
159                           [autobatch
160                           |unfold;intros;autobatch]
161                        |assumption]
162                     |*:autobatch]
163                  |autobatch]]]]]
164 qed.
165
166 theorem JS_trans : ∀G,T,U,V.G ⊢ T ⊴ U → G ⊢ U ⊴ V → G ⊢ T ⊴ V.
167 intros 5;apply (JS_trans_prova ? G);autobatch;
168 qed.
169
170 theorem JS_narrow : ∀G1,G2,X,P,Q,T,U.
171                        (G2 @ (mk_bound true X Q :: G1)) ⊢ T ⊴ U → G1 ⊢ P ⊴ Q →
172                        (G2 @ (mk_bound true X P :: G1)) ⊢ T ⊴ U.
173 intros;apply (narrowing ? ? ? ? ? ? ? H1 ? H) [|autobatch]
174 intros;apply (JS_trans ? ? ? ? ? H2);apply (JS_weakening ? ? ? H1);
175      [autobatch|unfold;intros;autobatch]
176 qed.