]> matita.cs.unibo.it Git - helm.git/blob - matita/library/Fsub/defn.ma
library/Fsub: minor fix
[helm.git] / matita / library / Fsub / defn.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/defn".
16 include "logic/equality.ma".
17 include "nat/nat.ma".
18 include "datatypes/bool.ma".
19 include "nat/compare.ma".
20 include "list/list.ma".
21
22 (*** useful definitions and lemmas not really related to Fsub ***)
23
24 lemma eqb_case : \forall x,y.(eqb x y) = true \lor (eqb x y) = false.
25 intros;elim (eqb x y);auto;
26 qed.
27        
28 lemma eq_eqb_case : \forall x,y.((x = y) \land (eqb x y) = true) \lor
29                                 ((x \neq y) \land (eqb x y) = false).
30 intros;lapply (eqb_to_Prop x y);elim (eqb_case x y)
31   [rewrite > H in Hletin;simplify in Hletin;left;auto
32   |rewrite > H in Hletin;simplify in Hletin;right;auto]
33 qed.
34
35 let rec max m n \def
36   match (leb m n) with
37      [true \Rightarrow n
38      |false \Rightarrow m]. 
39
40 (*** representation of Fsub types ***)  
41 inductive Typ : Set \def
42   | TVar : nat \to Typ            (* type var *)
43   | TFree: nat \to Typ            (* free type name *)
44   | Top : Typ                     (* maximum type *)
45   | Arrow : Typ \to Typ \to Typ   (* functions *) 
46   | Forall : Typ \to Typ \to Typ. (* universal type *)
47   
48 (*** representation of Fsub terms ***)
49 inductive Term : Set \def
50   | Var : nat \to Term            (* variable *)
51   | Free : nat \to Term          (* free name *)
52   | Abs : Typ \to Term \to Term   (* abstraction *)
53   | App : Term \to Term \to Term  (* function application *)
54   | TAbs : Typ \to Term \to Term  (* type abstraction *)
55   | TApp : Term \to Typ \to Term. (* type application *)
56   
57 (* representation of bounds *)
58
59 record bound : Set \def { 
60                           istype : bool;    (* is subtyping bound? *)
61                           name   : nat ;    (* name *)
62                           btype  : Typ      (* type to which the name is bound *)
63                         }.
64                
65 (* representation of Fsub typing environments *)
66 definition Env \def (list bound).
67 definition Empty \def (nil bound).
68 definition Cons \def \lambda G,X,T.((mk_bound false X T) :: G).
69 definition TCons \def \lambda G,X,T.((mk_bound true X T) :: G).
70
71 definition env_append : Env \to Env \to Env \def \lambda G,H.(H @ G).
72   
73 notation "hvbox(\Forall S. break T)" 
74   non associative with precedence 90
75 for @{ 'forall $S $T}.
76
77 notation "hvbox(#x)"
78   with precedence 60
79   for @{'var $x}.
80   
81 notation "hvbox(##x)"
82   with precedence 61
83   for @{'tvar $x}.
84
85 notation "hvbox(!x)"
86   with precedence 60
87   for @{'name $x}.
88   
89 notation "hvbox(!!x)"
90   with precedence 61
91   for @{'tname $x}.
92
93 notation "hvbox(s break \mapsto t)"
94   right associative with precedence 55
95   for @{ 'arrow $s $t }.
96
97 interpretation "universal type" 'forall S T = (cic:/matita/test/Typ.ind#xpointer(1/1/5) S T).
98
99 interpretation "bound var" 'var x = (cic:/matita/test/Typ.ind#xpointer(1/1/1) x).
100
101 interpretation "bound tvar" 'tvar x = (cic:/matita/test/Typ.ind#xpointer(1/1/3) x).
102
103 interpretation "bound tname" 'tname x = (cic:/matita/test/Typ.ind#xpointer(1/1/2) x).
104
105 interpretation "arrow type" 'arrow S T = (cic:/matita/test/Typ.ind#xpointer(1/1/4) S T). 
106
107 (*** Various kinds of substitution, not all will be used probably ***)
108
109 (* substitutes i-th dangling index in type T with type U *)
110 let rec subst_type_nat T U i \def
111     match T with
112     [ (TVar n) \Rightarrow match (eqb n i) with
113       [ true \Rightarrow U
114       | false \Rightarrow T]
115     | (TFree X) \Rightarrow T
116     | Top \Rightarrow T
117     | (Arrow T1 T2) \Rightarrow (Arrow (subst_type_nat T1 U i) (subst_type_nat T2 U i))
118     | (Forall T1 T2) \Rightarrow (Forall (subst_type_nat T1 U i) (subst_type_nat T2 U (S i))) ].
119
120 (* substitutes 0-th dangling index in type T with type U *)
121 let rec subst_type_O T U  \def subst_type_nat T U O. 
122
123 (* substitutes 0-th dangling index in term t with term u *)
124 let rec subst_term_O t u  \def
125   let rec aux t0 i \def
126     match t0 with
127     [ (Var n) \Rightarrow match (eqb n i) with
128       [ true \Rightarrow u
129       | false \Rightarrow t0]
130     | (Free X) \Rightarrow t0
131     | (Abs T1 t1) \Rightarrow (Abs T1 (aux t1 (S i)))
132     | (App t1 t2) \Rightarrow (App (aux t1 i) (aux t2 i))
133     | (TAbs T1 t1) \Rightarrow (TAbs T1 (aux t1 (S i)))
134     | (TApp t1 T1) \Rightarrow (TApp (aux t1 i) T1) ]
135   in aux t O.
136
137 (* substitutes 0-th dangling index in term T, which shall be a TVar,
138    with type U *)
139 let rec subst_term_tO t T  \def
140   let rec aux t0 i \def
141     match t0 with
142     [ (Var n) \Rightarrow t0 
143     | (Free X) \Rightarrow t0
144     | (Abs T1 t1) \Rightarrow (Abs (subst_type_nat T1 T i) (aux t1 (S i)))
145     | (App t1 t2) \Rightarrow (App (aux t1 i) (aux t2 i))
146     | (TAbs T1 t1) \Rightarrow (TAbs (subst_type_nat T1 T i) (aux t1 (S i)))
147     | (TApp t1 T1) \Rightarrow (TApp (aux t1 i) (subst_type_nat T1 T i)) ]
148   in aux t O.
149
150 (* substitutes (TFree X) in type T with type U *)
151 let rec subst_type_tfree_type T X U on T \def
152   match T with
153     [ (TVar n) \Rightarrow T
154     | (TFree Y) \Rightarrow match (eqb X Y) with
155       [ true \Rightarrow U
156       | false \Rightarrow T ]
157     | Top \Rightarrow T
158     | (Arrow T1 T2) \Rightarrow (Arrow (subst_type_tfree_type T1 X U) 
159                                        (subst_type_tfree_type T2 X U))
160     | (Forall T1 T2) \Rightarrow (Forall (subst_type_tfree_type T1 X U) 
161                                        (subst_type_tfree_type T2 X U)) ].
162             
163 (*** height of T's syntactic tree ***)
164
165 let rec t_len T \def
166   match T with
167      [(TVar n) \Rightarrow (S O)
168      |(TFree X) \Rightarrow (S O)
169      |Top \Rightarrow (S O)
170      |(Arrow T1 T2) \Rightarrow (S (max (t_len T1) (t_len T2)))
171      |(Forall T1 T2) \Rightarrow (S (max (t_len T1) (t_len T2)))].
172
173 (* 
174 let rec fresh_name G n \def
175   match G with
176     [ nil \Rightarrow n
177     | (cons b H) \Rightarrow match (leb (fresh_name H n) (name b)) with
178       [ true \Rightarrow (S (name b))
179       | false \Rightarrow (fresh_name H n) ]].
180
181 lemma freshname_Gn_geq_n : \forall G,n.((fresh_name G n) >= n).
182 intro;elim G
183   [simplify;unfold;constructor 1
184   |simplify;cut ((leb (fresh_name l n) (name s)) = true \lor
185                  (leb (fresh_name l n) (name s) = false))
186      [elim Hcut
187         [lapply (leb_to_Prop (fresh_name l n) (name s));rewrite > H1 in Hletin;
188          simplify in Hletin;rewrite > H1;simplify;lapply (H n);
189          unfold in Hletin1;unfold;
190          apply (trans_le ? ? ? Hletin1);
191          apply (trans_le ? ? ? Hletin);constructor 2;constructor 1
192         |rewrite > H1;simplify;apply H]
193      |elim (leb (fresh_name l n) (name s)) [left;reflexivity|right;reflexivity]]]
194 qed.
195
196 lemma freshname_consGX_gt_X : \forall G,X,T,b,n.
197           (fresh_name (cons ? (mk_bound b X T) G) n) > X.
198 intros.unfold.unfold.simplify.cut ((leb (fresh_name G n) X) = true \lor 
199                                    (leb (fresh_name G n) X) = false)
200   [elim Hcut
201      [rewrite > H;simplify;constructor 1
202      |rewrite > H;simplify;lapply (leb_to_Prop (fresh_name G n) X);
203       rewrite > H in Hletin;simplify in Hletin;
204       lapply (not_le_to_lt ? ? Hletin);unfold in Hletin1;assumption]
205   |elim (leb (fresh_name G n) X) [left;reflexivity|right;reflexivity]]
206 qed.
207
208 lemma freshname_case : \forall G,X,T,b,n.
209   (fresh_name ((mk_bound b X T) :: G) n) = (fresh_name G n) \lor
210   (fresh_name ((mk_bound b X T) :: G) n) = (S X).
211 intros.simplify.cut ((leb (fresh_name G n) X) = true \lor
212                  (leb (fresh_name G n) X) = false)
213   [elim Hcut
214      [rewrite > H;simplify;right;reflexivity
215      |rewrite > H;simplify;left;reflexivity]
216   |elim (leb (fresh_name G n) X)
217      [left;reflexivity|right;reflexivity]]
218 qed.
219
220 lemma freshname_monotone_n : \forall G,m,n.(m \leq n) \to
221                              ((fresh_name G m) \leq (fresh_name G n)).
222 intros.elim G
223   [simplify;assumption
224   |simplify;cut ((leb (fresh_name l m) (name s)) = true \lor
225                  (leb (fresh_name l m) (name s)) = false)
226      [cut ((leb (fresh_name l n) (name s)) = true \lor
227               (leb (fresh_name l n) (name s)) = false)
228         [elim Hcut
229            [rewrite > H2;simplify;elim Hcut1
230               [rewrite > H3;simplify;constructor 1 
231               |rewrite > H3;simplify;
232                lapply (leb_to_Prop (fresh_name l n) (name s));
233                rewrite > H3 in Hletin;simplify in Hletin;
234                lapply (not_le_to_lt ? ? Hletin);unfold in Hletin1;assumption]
235            |rewrite > H2;simplify;elim Hcut1
236               [rewrite > H3;simplify;
237                lapply (leb_to_Prop (fresh_name l m) (name s));
238                rewrite > H2 in Hletin;simplify in Hletin;
239                lapply (not_le_to_lt ? ? Hletin);unfold in Hletin1;
240                lapply (leb_to_Prop (fresh_name l n) (name s));
241                rewrite > H3 in Hletin2;
242                simplify in Hletin2;lapply (trans_le ? ? ? Hletin1 H1);
243                lapply (trans_le ? ? ? Hletin3 Hletin2);
244                absurd ((S (name s)) \leq (name s))
245                  [assumption|apply not_le_Sn_n]
246               |rewrite > H3;simplify;assumption]]
247         |elim (leb (fresh_name l n) (name s)) 
248            [left;reflexivity|right;reflexivity]]
249      |elim (leb (fresh_name l m) (name s)) [left;reflexivity|right;reflexivity]]]
250 qed.
251
252 lemma freshname_monotone_G : \forall G,X,T,b,n.
253                    (fresh_name G n) \leq (fresh_name ((mk_bound b X T) :: G) n).
254 intros.simplify.cut ((leb (fresh_name G n) X) = true \lor
255                      (leb (fresh_name G n) X) = false)
256   [elim Hcut
257      [rewrite > H;simplify;lapply (leb_to_Prop (fresh_name G n) X);
258       rewrite > H in Hletin;simplify in Hletin;constructor 2;assumption
259      |rewrite > H;simplify;constructor 1]
260   |elim (leb (fresh_name G n) X)
261      [left;reflexivity|right;reflexivity]]
262 qed.*)
263
264 lemma subst_O_nat : \forall T,U.((subst_type_O T U) = (subst_type_nat T U O)).
265 intros;elim T;simplify;reflexivity;
266 qed.
267
268 (* FIXME: these definitions shouldn't be part of the poplmark challenge
269    - use destruct instead, when hopefully it will get fixed... *) 
270
271 definition head \def
272   \lambda G:(list bound).match G with
273     [ nil \Rightarrow (mk_bound false O Top)
274     | (cons b H) \Rightarrow b].
275
276 definition head_nat \def
277   \lambda G:(list nat).match G with
278     [ nil \Rightarrow O
279     | (cons n H) \Rightarrow n].
280
281 lemma inj_head : \forall h1,h2:bound.\forall t1,t2:Env.
282                  ((h1::t1) = (h2::t2)) \to (h1 = h2).
283 intros.lapply (eq_f ? ? head ? ? H).simplify in Hletin.assumption.
284 qed.
285
286 lemma inj_head_nat : \forall h1,h2:nat.\forall t1,t2:(list nat).
287                  ((h1::t1) = (h2::t2)) \to (h1 = h2).
288 intros.lapply (eq_f ? ? head_nat ? ? H).simplify in Hletin.assumption.
289 qed.
290
291 lemma inj_tail : \forall A.\forall h1,h2:A.\forall t1,t2:(list A).
292                  ((h1::t1) = (h2::t2)) \to (t1 = t2).
293 intros.lapply (eq_f ? ? (tail ?) ? ? H).simplify in Hletin.assumption.
294 qed.
295
296 (* end of fixme *) 
297
298 (*** definitions and theorems about lists ***)
299
300 inductive in_list (A : Set) : A \to (list A) \to Prop \def
301   | in_Base : \forall x:A.\forall l:(list A).
302               (in_list A x (x :: l))
303   | in_Skip : \forall x,y:A.\forall l:(list A).
304               (in_list A x l) \to (in_list A x (y :: l)).
305
306 (* var binding is in env judgement *)                
307 definition var_bind_in_env : bound \to Env \to Prop \def
308   \lambda b,G.(in_list bound b G).
309
310 (* FIXME: use the map in library/list (when there will be one) *)
311 definition map : \forall A,B,f.((list A) \to (list B)) \def
312   \lambda A,B,f.let rec map (l : (list A)) : (list B) \def
313     match l in list return \lambda l0:(list A).(list B) with
314       [nil \Rightarrow (nil B)
315       |(cons (a:A) (t:(list A))) \Rightarrow 
316         (cons B (f a) (map t))] in map.
317
318 definition fv_env : (list bound) \to (list nat) \def
319   \lambda G.(map ? ? (\lambda b.match b with
320       [(mk_bound B X T) \Rightarrow X]) G).
321
322 (* variable is in env judgement *)              
323 definition var_in_env : nat \to Env \to Prop \def
324   \lambda x,G.(in_list nat x (fv_env G)).
325   
326 definition var_type_in_env : nat \to Env \to Prop \def
327   \lambda x,G.\exists T.(var_bind_in_env (mk_bound true x T) G).
328
329 definition incl : \forall A.(list A) \to (list A) \to Prop \def
330   \lambda A,l,m.\forall x.(in_list A x l) \to (in_list A x m).               
331               
332 let rec fv_type T \def
333   match T with
334     [(TVar n) \Rightarrow []
335     |(TFree x) \Rightarrow [x]
336     |Top \Rightarrow []
337     |(Arrow U V) \Rightarrow ((fv_type U) @ (fv_type V))
338     |(Forall U V) \Rightarrow ((fv_type U) @ (fv_type V))].
339
340 lemma var_notinbG_notinG : \forall G,x,b.
341                            (\lnot (var_in_env x (b::G))) 
342                            \to \lnot (var_in_env x G).
343 intros 3.elim b.unfold.intro.elim H.unfold.simplify.constructor 2.exact H1.
344 qed.
345
346 lemma in_list_nil : \forall A,x.\lnot (in_list A x []).
347 intros.unfold.intro.inversion H
348   [intros;lapply (sym_eq ? ? ? H2);absurd (a::l = [])
349      [assumption|apply nil_cons]
350   |intros;lapply (sym_eq ? ? ? H4);absurd (a1::l = [])
351      [assumption|apply nil_cons]]
352 qed.
353
354 lemma notin_cons : \forall A,x,y,l.\lnot (in_list A x (y::l)) \to
355                       (y \neq x) \land \lnot (in_list A x l).
356 intros.split
357   [unfold;intro;apply H;rewrite > H1;constructor 1
358   |unfold;intro;apply H;constructor 2;assumption]
359 qed.
360
361 lemma boundinenv_natinfv : \forall x,G.
362                               (\exists B,T.(in_list ? (mk_bound B x T) G)) \to
363                               (in_list ? x (fv_env G)).
364 intros 2;elim G
365   [elim H;elim H1;lapply (in_list_nil ? ? H2);elim Hletin
366   |elim H1;elim H2;inversion H3
367      [intros;rewrite < H4;simplify;apply in_Base
368      |intros;elim a3;simplify;apply in_Skip;
369       lapply (inj_tail ? ? ? ? ? H7);rewrite > Hletin in H;apply H;
370       apply ex_intro
371         [apply a
372         |apply ex_intro
373            [apply a1
374            |rewrite > H6;assumption]]]]
375 qed.
376
377 lemma nat_in_list_case : \forall G,H,n.(in_list nat n (H @ G)) \to 
378                                (in_list nat n G) \lor (in_list nat n H).
379 intros 3.elim H
380   [simplify in H1;left;assumption
381   |simplify in H2;inversion H2
382     [intros;lapply (inj_head_nat ? ? ? ? H4);rewrite > Hletin;
383      right;apply in_Base
384     |intros;lapply (inj_tail ? ? ? ? ? H6);rewrite < Hletin in H3;
385      rewrite > H5 in H1;lapply (H1 H3);elim Hletin1
386        [left;assumption|right;apply in_Skip;assumption]]]
387 qed.
388
389 lemma natinG_or_inH_to_natinGH : \forall G,H,n.
390                       (in_list nat n G) \lor (in_list nat n H) \to
391                       (in_list nat n (H @ G)).
392 intros.elim H1
393   [elim H
394      [simplify;assumption
395      |simplify;apply in_Skip;assumption]
396   |generalize in match H2;elim H2
397      [simplify;apply in_Base
398      |lapply (H4 H3);simplify;apply in_Skip;assumption]]
399 qed.
400
401 lemma natinfv_boundinenv : \forall x,G.(in_list ? x (fv_env G)) \to
402                               \exists B,T.(in_list ? (mk_bound B x T) G).
403 intros 2;elim G 0
404   [simplify;intro;lapply (in_list_nil ? ? H);elim Hletin
405   |intros 3;elim s;simplify in H1;inversion H1
406      [intros;rewrite < H2;simplify;apply ex_intro
407         [apply b
408         |apply ex_intro
409            [apply t
410            |lapply (inj_head_nat ? ? ? ? H3);rewrite > H2;rewrite < Hletin;
411             apply in_Base]]
412      |intros;lapply (inj_tail ? ? ? ? ? H5);rewrite < Hletin in H2;
413       rewrite < H4 in H2;lapply (H H2);elim Hletin1;elim H6;apply ex_intro
414         [apply a2
415         |apply ex_intro
416            [apply a3
417            |apply in_Skip;rewrite < H4;assumption]]]]
418 qed.
419            
420 lemma incl_bound_fv : \forall l1,l2.(incl ? l1 l2) \to 
421                          (incl ? (fv_env l1) (fv_env l2)).
422 intros.unfold in H.unfold.intros.apply boundinenv_natinfv.
423 lapply (natinfv_boundinenv ? ? H1).elim Hletin.elim H2.apply ex_intro
424   [apply a
425   |apply ex_intro
426      [apply a1
427      |apply (H ? H3)]]
428 qed.
429
430 (* lemma incl_cons : \forall x,l1,l2.
431                   (incl bound l1 l2) \to (incl bound (x :: l1) (x :: l2)).
432 intros.unfold in H.unfold.intros.inversion H1
433   [intros;lapply (inj_head ? ? ? ? H3);rewrite > Hletin;apply in_Base
434   |intros;apply in_Skip;apply H;lapply (inj_tail ? ? ? ? ? H5);rewrite > Hletin;
435    assumption]
436 qed. *)
437
438 lemma incl_nat_cons : \forall x,l1,l2.
439                   (incl nat l1 l2) \to (incl nat (x :: l1) (x :: l2)).
440 intros.unfold in H.unfold.intros.inversion H1
441   [intros;lapply (inj_head_nat ? ? ? ? H3);rewrite > Hletin;apply in_Base
442   |intros;apply in_Skip;apply H;lapply (inj_tail ? ? ? ? ? H5);rewrite > Hletin;
443    assumption]
444 qed.
445
446 lemma boundin_envappend_case : \forall G,H,b.(var_bind_in_env b (H @ G)) \to 
447                                (var_bind_in_env b G) \lor (var_bind_in_env b H).
448 intros 3.elim H
449   [simplify in H1;left;assumption
450   |unfold in H2;inversion H2
451     [intros;simplify in H4;lapply (inj_head ? ? ? ? H4);rewrite > Hletin;
452      right;apply in_Base
453     |intros;simplify in H6;lapply (inj_tail ? ? ? ? ? H6);rewrite < Hletin in H3;
454      rewrite > H5 in H1;lapply (H1 H3);elim Hletin1
455        [left;assumption|right;apply in_Skip;assumption]]]
456 qed.
457
458 lemma varin_envappend_case: \forall G,H,x.(var_in_env x (H @ G)) \to
459                             (var_in_env x G) \lor (var_in_env x H).
460 intros 3.elim H 0
461   [simplify;intro;left;assumption
462   |intros 2;elim s;simplify in H2;inversion H2
463      [intros;lapply (inj_head_nat ? ? ? ? H4);rewrite > Hletin;right;
464       simplify;constructor 1
465      |intros;lapply (inj_tail ? ? ? ? ? H6);
466       lapply H1
467         [rewrite < H5;elim Hletin1
468            [left;assumption|right;simplify;constructor 2;assumption]
469         |unfold var_in_env;unfold fv_env;rewrite > Hletin;rewrite > H5;
470          assumption]]]
471 qed.
472
473 lemma boundinG_or_boundinH_to_boundinGH : \forall G,H,b.
474                       (var_bind_in_env b G) \lor (var_bind_in_env b H) \to
475                       (var_bind_in_env b (H @ G)).
476 intros.elim H1
477   [elim H
478      [simplify;assumption
479      |simplify;apply in_Skip;assumption]
480   |generalize in match H2;elim H2
481      [simplify;apply in_Base
482      |lapply (H4 H3);simplify;apply in_Skip;assumption]]
483 qed. 
484
485
486 lemma varinG_or_varinH_to_varinGH : \forall G,H,x.
487                           (var_in_env x G) \lor (var_in_env x H) \to
488                           (var_in_env x (H @ G)).
489 intros.elim H1 0
490   [elim H
491      [simplify;assumption
492      |elim s;simplify;constructor 2;apply (H2 H3)]
493   |elim H 0
494      [simplify;intro;lapply (in_list_nil nat x H2);elim Hletin
495      |intros 2;elim s;simplify in H3;inversion H3
496         [intros;lapply (inj_head_nat ? ? ? ? H5);rewrite > Hletin;simplify;
497          constructor 1
498         |intros;simplify;constructor 2;rewrite < H6;apply H2;
499          lapply (inj_tail ? ? ? ? ? H7);rewrite > H6;unfold;unfold fv_env;
500          rewrite > Hletin;assumption]]]
501 qed.
502
503 lemma varbind_to_append : \forall G,b.(var_bind_in_env b G) \to
504                           \exists G1,G2.(G = (G2 @ (b :: G1))).
505 intros.generalize in match H.elim H
506   [apply ex_intro [apply l|apply ex_intro [apply Empty|reflexivity]]
507   |lapply (H2 H1);elim Hletin;elim H4;rewrite > H5;
508    apply ex_intro 
509      [apply a2|apply ex_intro [apply (a1 :: a3)|simplify;reflexivity]]]
510 qed.
511   
512 (*** Type Well-Formedness judgement ***)
513
514 inductive WFType : Env \to Typ \to Prop \def
515   | WFT_TFree : \forall X,G.(in_list ? X (fv_env G)) 
516                 \to (WFType G (TFree X))
517   | WFT_Top : \forall G.(WFType G Top)
518   | WFT_Arrow : \forall G,T,U.(WFType G T) \to (WFType G U) \to 
519                 (WFType G (Arrow T U))
520   | WFT_Forall : \forall G,T,U.(WFType G T) \to
521                  (\forall X:nat.
522                     (\lnot (in_list ? X (fv_env G))) \to
523                     (\lnot (in_list ? X (fv_type U))) \to
524                     (WFType ((mk_bound true X T) :: G) 
525                        (subst_type_O U (TFree X)))) \to 
526                  (WFType G (Forall T U)).
527
528 (*** Environment Well-Formedness judgement ***)
529
530 inductive WFEnv : Env \to Prop \def
531   | WFE_Empty : (WFEnv Empty)
532   | WFE_cons : \forall B,X,T,G.(WFEnv G) \to 
533                \lnot (in_list ? X (fv_env G)) \to
534                   (WFType G T) \to (WFEnv ((mk_bound B X T) :: G)).
535             
536 (*** Subtyping judgement ***)              
537 inductive JSubtype : Env \to Typ \to Typ \to Prop \def
538   | SA_Top : \forall G:Env.\forall T:Typ.(WFEnv G) \to
539              (WFType G T) \to (JSubtype G T Top)
540   | SA_Refl_TVar : \forall G:Env.\forall X:nat.(WFEnv G) \to (var_in_env X G) 
541                    \to (JSubtype G (TFree X) (TFree X))
542   | SA_Trans_TVar : \forall G:Env.\forall X:nat.\forall T:Typ.
543                     \forall U:Typ.
544                     (var_bind_in_env (mk_bound true X U) G) \to
545                     (JSubtype G U T) \to (JSubtype G (TFree X) T)
546   | SA_Arrow : \forall G:Env.\forall S1,S2,T1,T2:Typ.
547                (JSubtype G T1 S1) \to (JSubtype G S2 T2) \to
548                (JSubtype G (Arrow S1 S2) (Arrow T1 T2))
549   | SA_All : \forall G:Env.\forall S1,S2,T1,T2:Typ.
550              (JSubtype G T1 S1) \to
551              (\forall X:nat.\lnot (var_in_env X G) \to
552                 (JSubtype ((mk_bound true X T1) :: G) 
553                    (subst_type_O S2 (TFree X)) (subst_type_O T2 (TFree X)))) \to
554              (JSubtype G (Forall S1 S2) (Forall T1 T2)).
555
556 (*** Typing judgement ***)
557 inductive JType : Env \to Term \to Typ \to Prop \def
558   | T_Var : \forall G:Env.\forall x:nat.\forall T:Typ.
559             (WFEnv G) \to (var_bind_in_env (mk_bound false x T) G) \to
560             (JType G (Free x) T)
561   | T_Abs : \forall G.\forall T1,T2:Typ.\forall t2:Term.
562             \forall x:nat.
563             (JType ((mk_bound false x T1)::G) (subst_term_O t2 (Free x)) T2) \to
564             (JType G (Abs T1 t2) (Arrow T1 T2))
565   | T_App : \forall G.\forall t1,t2:Term.\forall T2:Typ.
566             \forall T1:Typ.(JType G t1 (Arrow T1 T2)) \to (JType G t2 T1) \to
567             (JType G (App t1 t2) T2)
568   | T_TAbs : \forall G:Env.\forall T1,T2:Typ.\forall t2:Term.
569              \forall X:nat.
570              (JType ((mk_bound true X T1)::G) 
571              (subst_term_tO t2 (TFree X)) (subst_type_O T2 (TFree X)))
572              \to (JType G (TAbs T1 t2) (Forall T1 T2))
573   | T_TApp : \forall G:Env.\forall t1:Term.\forall T2,T12:Typ.
574              \forall X:nat.\forall T11:Typ.
575              (JType G t1 (Forall T11 (subst_type_tfree_type T12 X (TVar O)))) \to 
576              (JSubtype G T2 T11)
577              \to (JType G (TApp t1 T2) (subst_type_tfree_type T12 X T2))
578   | T_Sub : \forall G:Env.\forall t:Term.\forall T:Typ.
579             \forall S:Typ.(JType G t S) \to (JSubtype G S T) \to (JType G t T).
580
581
582 lemma WFT_env_incl : \forall G,T.(WFType G T) \to
583                      \forall H.(incl ? (fv_env G) (fv_env H)) \to (WFType H T).
584 intros 4.generalize in match H1.elim H
585   [apply WFT_TFree;unfold in H3;apply (H3 ? H2)
586   |apply WFT_Top
587   |apply WFT_Arrow [apply (H3 ? H6)|apply (H5 ? H6)]
588   |apply WFT_Forall 
589      [apply (H3 ? H6)
590      |intros;apply H5
591         [unfold;intro;unfold in H7;apply H7;unfold in H6;apply(H6 ? H9)
592         |assumption
593         |simplify;apply (incl_nat_cons ? ? ? H6)]]]
594 qed.
595
596 (*** definitions and theorems about swaps ***)
597
598 definition swap : nat \to nat \to nat \to nat \def
599   \lambda u,v,x.match (eqb x u) with
600     [true \Rightarrow v
601     |false \Rightarrow match (eqb x v) with
602        [true \Rightarrow u
603        |false \Rightarrow x]].
604        
605 lemma swap_left : \forall x,y.(swap x y x) = y.
606 intros;unfold swap;rewrite > eqb_n_n;simplify;reflexivity;
607 qed.
608
609 lemma swap_right : \forall x,y.(swap x y y) = x.
610 intros;unfold swap;elim (eq_eqb_case y x)
611   [elim H;rewrite > H2;simplify;rewrite > H1;reflexivity
612   |elim H;rewrite > H2;simplify;rewrite > eqb_n_n;simplify;reflexivity]
613 qed.
614
615 lemma swap_other : \forall x,y,z.(z \neq x) \to (z \neq y) \to (swap x y z) = z.
616 intros;unfold swap;elim (eq_eqb_case z x)
617   [elim H2;lapply (H H3);elim Hletin
618   |elim H2;rewrite > H4;simplify;elim (eq_eqb_case z y)
619      [elim H5;lapply (H1 H6);elim Hletin
620      |elim H5;rewrite > H7;simplify;reflexivity]]
621 qed. 
622
623 lemma swap_inv : \forall u,v,x.(swap u v (swap u v x)) = x.
624 intros;unfold in match (swap u v x);elim (eq_eqb_case x u)
625   [elim H;rewrite > H2;simplify;rewrite > H1;apply swap_right
626   |elim H;rewrite > H2;simplify;elim (eq_eqb_case x v)
627      [elim H3;rewrite > H5;simplify;rewrite > H4;apply swap_left
628      |elim H3;rewrite > H5;simplify;apply (swap_other ? ? ? H1 H4)]]
629 qed.
630
631 lemma swap_inj : \forall u,v,x,y.(swap u v x) = (swap u v y) \to x = y.
632 intros;unfold swap in H;elim (eq_eqb_case x u)
633   [elim H1;elim (eq_eqb_case y u)
634      [elim H4;rewrite > H5;assumption
635      |elim H4;rewrite > H3 in H;rewrite > H6 in H;simplify in H;
636       elim (eq_eqb_case y v)
637         [elim H7;rewrite > H9 in H;simplify in H;rewrite > H in H8;
638          lapply (H5 H8);elim Hletin
639         |elim H7;rewrite > H9 in H;simplify in H;elim H8;symmetry;assumption]]
640   |elim H1;elim (eq_eqb_case y u)
641      [elim H4;rewrite > H3 in H;rewrite > H6 in H;simplify in H;
642       elim (eq_eqb_case x v)
643         [elim H7;rewrite > H9 in H;simplify in H;rewrite < H in H8;
644          elim H2;assumption
645         |elim H7;rewrite > H9 in H;simplify in H;elim H8;assumption]
646      |elim H4;rewrite > H3 in H;rewrite > H6 in H;simplify in H;
647       elim (eq_eqb_case x v)
648         [elim H7;rewrite > H9 in H;elim (eq_eqb_case y v)
649            [elim H10;rewrite > H11;assumption
650            |elim H10;rewrite > H12 in H;simplify in H;elim H5;symmetry;
651             assumption]
652         |elim H7;rewrite > H9 in H;elim (eq_eqb_case y v)
653            [elim H10;rewrite > H12 in H;simplify in H;elim H2;assumption
654            |elim H10;rewrite > H12 in H;simplify in H;assumption]]]]
655 qed. 
656
657 lemma fv_subst_type_nat : \forall x,T,y,n.(in_list ? x (fv_type T)) \to
658                          (in_list ? x (fv_type (subst_type_nat T (TFree y) n))).
659 intros 3;elim T 0
660   [intros;simplify in H;elim (in_list_nil ? ? H)
661   |simplify;intros;assumption
662   |simplify;intros;assumption
663   |intros;simplify in H2;elim (nat_in_list_case ? ? ? H2)
664      [simplify;apply natinG_or_inH_to_natinGH;left;apply (H1 ? H3)
665      |simplify;apply natinG_or_inH_to_natinGH;right;apply (H ? H3)]
666   |intros;simplify in H2;elim (nat_in_list_case ? ? ? H2)
667      [simplify;apply natinG_or_inH_to_natinGH;left;apply (H1 ? H3)
668      |simplify;apply natinG_or_inH_to_natinGH;right;apply (H ? H3)]]
669 qed.
670
671 lemma fv_subst_type_O : \forall x,T,y.(in_list ? x (fv_type T)) \to
672                          (in_list ? x (fv_type (subst_type_O T (TFree y)))).
673 intros;rewrite > subst_O_nat;apply (fv_subst_type_nat ? ? ? ? H);
674 qed.
675
676 let rec swap_Typ u v T on T \def
677   match T with
678      [(TVar n) \Rightarrow (TVar n)
679      |(TFree X) \Rightarrow (TFree (swap u v X))
680      |Top \Rightarrow Top
681      |(Arrow T1 T2) \Rightarrow (Arrow (swap_Typ u v T1) (swap_Typ u v T2))
682      |(Forall T1 T2) \Rightarrow (Forall (swap_Typ u v T1) (swap_Typ u v T2))].
683      
684 lemma swap_Typ_inv : \forall u,v,T.(swap_Typ u v (swap_Typ u v T)) = T.
685 intros;elim T
686   [simplify;reflexivity
687   |simplify;rewrite > swap_inv;reflexivity
688   |simplify;reflexivity
689   |simplify;rewrite > H;rewrite > H1;reflexivity
690   |simplify;rewrite > H;rewrite > H1;reflexivity]
691 qed.
692
693 lemma swap_Typ_not_free : \forall u,v,T.\lnot (in_list ? u (fv_type T)) \to
694                       \lnot (in_list ? v (fv_type T)) \to (swap_Typ u v T) = T.
695 intros 3;elim T 0
696   [intros;simplify;reflexivity
697   |simplify;intros;cut (n \neq u \land n \neq v)
698      [elim Hcut;rewrite > (swap_other ? ? ? H2 H3);reflexivity
699      |split
700         [unfold;intro;apply H;rewrite > H2;apply in_Base
701         |unfold;intro;apply H1;rewrite > H2;apply in_Base]]
702   |simplify;intros;reflexivity
703   |simplify;intros;cut ((\lnot (in_list ? u (fv_type t)) \land
704                          \lnot (in_list ? u (fv_type t1))) \land
705                         (\lnot (in_list ? v (fv_type t)) \land
706                          \lnot (in_list ? v (fv_type t1))))
707      [elim Hcut;elim H4;elim H5;clear Hcut H4 H5;rewrite > (H H6 H8);
708       rewrite > (H1 H7 H9);reflexivity
709      |split
710         [split;unfold;intro;apply H2;apply natinG_or_inH_to_natinGH;auto
711         |split;unfold;intro;apply H3;apply natinG_or_inH_to_natinGH;auto]]
712   |simplify;intros;cut ((\lnot (in_list ? u (fv_type t)) \land
713                          \lnot (in_list ? u (fv_type t1))) \land
714                         (\lnot (in_list ? v (fv_type t)) \land
715                          \lnot (in_list ? v (fv_type t1))))
716      [elim Hcut;elim H4;elim H5;clear Hcut H4 H5;rewrite > (H H6 H8);
717       rewrite > (H1 H7 H9);reflexivity
718      |split
719         [split;unfold;intro;apply H2;apply natinG_or_inH_to_natinGH;auto
720         |split;unfold;intro;apply H3;apply natinG_or_inH_to_natinGH;auto]]]
721 qed.
722         
723 lemma subst_type_nat_swap : \forall u,v,T,X,m.
724          (swap_Typ u v (subst_type_nat T (TFree X) m)) =
725          (subst_type_nat (swap_Typ u v T) (TFree (swap u v X)) m).
726 intros 4;elim T
727   [simplify;elim (eqb_case n m);rewrite > H;simplify;reflexivity
728   |simplify;reflexivity
729   |simplify;reflexivity
730   |simplify;rewrite > H;rewrite > H1;reflexivity
731   |simplify;rewrite > H;rewrite > H1;reflexivity]
732 qed.
733
734 lemma subst_type_O_swap : \forall u,v,T,X.
735          (swap_Typ u v (subst_type_O T (TFree X))) =
736          (subst_type_O (swap_Typ u v T) (TFree (swap u v X))).
737 intros 4;rewrite > (subst_O_nat (swap_Typ u v T));rewrite > (subst_O_nat T);
738 apply subst_type_nat_swap;
739 qed.
740
741 lemma in_fv_type_swap : \forall u,v,x,T.((in_list ? x (fv_type T)) \to
742               (in_list ? (swap u v x) (fv_type (swap_Typ u v T)))) \land
743              ((in_list ? (swap u v x) (fv_type (swap_Typ u v T))) \to
744               (in_list ? x (fv_type T))).
745 intros;split
746   [elim T 0
747      [simplify;intros;elim (in_list_nil ? ? H)
748      |simplify;intros;cut (x = n)
749         [rewrite > Hcut;apply in_Base
750         |inversion H
751            [intros;lapply (inj_head_nat ? ? ? ? H2);rewrite > Hletin;
752             reflexivity
753            |intros;lapply (inj_tail ? ? ? ? ? H4);rewrite < Hletin in H1;
754             elim (in_list_nil ? ? H1)]]
755      |simplify;intro;elim (in_list_nil ? ? H)
756      |simplify;intros;elim (nat_in_list_case ? ? ? H2)
757         [apply natinG_or_inH_to_natinGH;left;apply (H1 H3)
758         |apply natinG_or_inH_to_natinGH;right;apply (H H3)]
759      |simplify;intros;elim (nat_in_list_case ? ? ? H2)
760         [apply natinG_or_inH_to_natinGH;left;apply (H1 H3)
761         |apply natinG_or_inH_to_natinGH;right;apply (H H3)]]
762   |elim T 0
763      [simplify;intros;elim (in_list_nil ? ? H)
764      |simplify;intros;cut ((swap u v x) = (swap u v n))
765         [lapply (swap_inj ? ? ? ? Hcut);rewrite > Hletin;apply in_Base
766         |inversion H
767            [intros;lapply (inj_head_nat ? ? ? ? H2);rewrite > Hletin;
768             reflexivity
769            |intros;lapply (inj_tail ? ? ? ? ? H4);rewrite < Hletin in H1;
770             elim (in_list_nil ? ? H1)]]
771      |simplify;intro;elim (in_list_nil ? ? H)
772      |simplify;intros;elim (nat_in_list_case ? ? ? H2)
773         [apply natinG_or_inH_to_natinGH;left;apply (H1 H3)
774         |apply natinG_or_inH_to_natinGH;right;apply (H H3)]
775      |simplify;intros;elim (nat_in_list_case ? ? ? H2)
776         [apply natinG_or_inH_to_natinGH;left;apply (H1 H3)
777         |apply natinG_or_inH_to_natinGH;right;apply (H H3)]]]
778 qed.
779         
780 definition swap_bound : nat \to nat \to bound \to bound \def
781   \lambda u,v,b.match b with
782      [(mk_bound B X T) \Rightarrow (mk_bound B (swap u v X) (swap_Typ u v T))].
783
784 definition swap_Env : nat \to nat \to Env \to Env \def
785   \lambda u,v,G.(map ? ? (\lambda b.(swap_bound u v b)) G). 
786
787 lemma lookup_swap : \forall x,u,v,T,B,G.(in_list ? (mk_bound B x T) G) \to
788     (in_list ? (mk_bound B (swap u v x) (swap_Typ u v T)) (swap_Env u v G)).
789 intros 6;elim G 0
790   [intros;elim (in_list_nil ? ? H)
791   |intro;elim s;simplify;inversion H1
792      [intros;lapply (inj_head ? ? ? ? H3);rewrite < H2 in Hletin;
793       destruct Hletin;rewrite > Hcut;rewrite > Hcut1;rewrite > Hcut2;
794       apply in_Base
795      |intros;lapply (inj_tail ? ? ? ? ? H5);rewrite < Hletin in H2;
796       rewrite < H4 in H2;apply in_Skip;apply (H H2)]]
797 qed.
798
799 lemma in_FV_subst : \forall x,T,U,n.(in_list ? x (fv_type T)) \to
800                                 (in_list ? x (fv_type (subst_type_nat T U n))).
801 intros 3;elim T
802   [simplify in H;inversion H
803      [intros;lapply (sym_eq ? ? ? H2);absurd (a::l = [])
804         [assumption|apply nil_cons]
805      |intros;lapply (sym_eq ? ? ? H4);absurd (a1::l = [])
806         [assumption|apply nil_cons]]
807   |simplify;simplify in H;assumption
808   |simplify in H;simplify;assumption
809   |simplify in H2;simplify;apply natinG_or_inH_to_natinGH;
810    lapply (nat_in_list_case ? ? ? H2);elim Hletin
811      [left;apply (H1 ? H3)
812      |right;apply (H ? H3)]
813   |simplify in H2;simplify;apply natinG_or_inH_to_natinGH;
814    lapply (nat_in_list_case ? ? ? H2);elim Hletin
815      [left;apply (H1 ? H3)
816      |right;apply (H ? H3)]]
817 qed.
818
819 lemma in_dom_swap : \forall u,v,x,G.
820                        ((in_list ? x (fv_env G)) \to 
821                        (in_list ? (swap u v x) (fv_env (swap_Env u v G)))) \land
822                        ((in_list ? (swap u v x) (fv_env (swap_Env u v G))) \to
823                        (in_list ? x (fv_env G))).
824 intros;split
825   [elim G 0
826      [simplify;intro;elim (in_list_nil ? ? H)
827      |intro;elim s 0;simplify;intros;inversion H1
828         [intros;lapply (inj_head_nat ? ? ? ? H3);rewrite > Hletin;apply in_Base
829         |intros;lapply (inj_tail ? ? ? ? ? H5);rewrite < Hletin in H2;
830          rewrite > H4 in H;apply in_Skip;apply (H H2)]]
831   |elim G 0
832      [simplify;intro;elim (in_list_nil ? ? H)
833      |intro;elim s 0;simplify;intros;inversion H1
834         [intros;lapply (inj_head_nat ? ? ? ? H3);rewrite < H2 in Hletin;
835          lapply (swap_inj ? ? ? ? Hletin);rewrite > Hletin1;apply in_Base
836         |intros;lapply (inj_tail ? ? ? ? ? H5);rewrite < Hletin in H2;
837          rewrite > H4 in H;apply in_Skip;apply (H H2)]]]
838 qed.
839
840 (*** lemma on fresh names ***)
841
842 lemma fresh_name : \forall l:(list nat).\exists n.\lnot (in_list ? n l).
843 cut (\forall l:(list nat).\exists n.\forall m.
844         (n \leq m) \to \lnot (in_list ? m l))
845   [intros;lapply (Hcut l);elim Hletin;apply ex_intro
846      [apply a
847      |apply H;constructor 1]
848   |intros;elim l
849     [apply ex_intro 
850        [apply O
851        |intros;unfold;intro;inversion H1
852           [intros;lapply (sym_eq ? ? ? H3);absurd (a::l1 = [])
853              [assumption|apply nil_cons]
854           |intros;lapply (sym_eq ? ? ? H5);absurd (a1::l1 = [])
855              [assumption|apply nil_cons]]]
856     |elim H;lapply (decidable_eq_nat a s);elim Hletin
857        [apply ex_intro
858           [apply (S a)
859           |intros;unfold;intro;inversion H4
860              [intros;lapply (inj_head_nat ? ? ? ? H6);rewrite < Hletin1 in H5;
861               rewrite < H2 in H5;rewrite > H5 in H3;
862               apply (not_le_Sn_n ? H3)
863              |intros;lapply (inj_tail ? ? ? ? ? H8);rewrite < Hletin1 in H5;
864               rewrite < H7 in H5;
865               apply (H1 m ? H5);lapply (le_S ? ? H3);
866               apply (le_S_S_to_le ? ? Hletin2)]]
867        |cut ((leb a s) = true \lor (leb a s) = false)
868           [elim Hcut
869              [apply ex_intro
870                 [apply (S s)
871                 |intros;unfold;intro;inversion H5
872                    [intros;lapply (inj_head_nat ? ? ? ? H7);rewrite > H6 in H4;
873                     rewrite < Hletin1 in H4;apply (not_le_Sn_n ? H4)
874                    |intros;lapply (inj_tail ? ? ? ? ? H9);
875                     rewrite < Hletin1 in H6;lapply (H1 a1)
876                       [apply (Hletin2 H6)
877                       |lapply (leb_to_Prop a s);rewrite > H3 in Hletin2;
878                        simplify in Hletin2;rewrite < H8;
879                        apply (trans_le ? ? ? Hletin2);
880                        apply (trans_le ? ? ? ? H4);constructor 2;constructor 1]]]
881              |apply ex_intro
882                 [apply a
883                 |intros;lapply (leb_to_Prop a s);rewrite > H3 in Hletin1;
884                  simplify in Hletin1;lapply (not_le_to_lt ? ? Hletin1);
885                  unfold in Hletin2;unfold;intro;inversion H5
886                    [intros;lapply (inj_head_nat ? ? ? ? H7);
887                     rewrite < Hletin3 in H6;rewrite > H6 in H4;
888                     apply (Hletin1 H4)
889                    |intros;lapply (inj_tail ? ? ? ? ? H9);
890                     rewrite < Hletin3 in H6;rewrite < H8 in H6;
891                     apply (H1 ? H4 H6)]]]
892           |elim (leb a s);auto]]]]
893 qed.
894
895 (*** lemmas on well-formedness ***)
896
897 lemma fv_WFT : \forall T,x,G.(WFType G T) \to (in_list ? x (fv_type T)) \to
898                   (in_list ? x (fv_env G)).
899 intros 4.elim H
900   [simplify in H2;inversion H2
901      [intros;lapply (inj_head_nat ? ? ? ? H4);rewrite < Hletin;assumption
902      |intros;lapply (inj_tail ? ? ? ? ? H6);rewrite < Hletin in H3;
903       inversion H3
904         [intros;lapply (sym_eq ? ? ? H8);absurd (a2 :: l2 = []) 
905            [assumption|apply nil_cons]
906         |intros;lapply (sym_eq ? ? ? H10);
907             absurd (a3 :: l2 = []) [assumption|apply nil_cons]]]
908   |simplify in H1;lapply (in_list_nil ? x H1);elim Hletin
909   |simplify in H5;lapply (nat_in_list_case ? ? ? H5);elim Hletin
910      [apply (H4 H6)
911      |apply (H2 H6)]
912   |simplify in H5;lapply (nat_in_list_case ? ? ? H5);elim Hletin
913      [lapply (fresh_name ((fv_type t1) @ (fv_env e)));elim Hletin1;
914       cut ((\lnot (in_list ? a (fv_type t1))) \land
915            (\lnot (in_list ? a (fv_env e))))
916         [elim Hcut;lapply (H4 ? H9 H8)
917            [cut (x \neq a)
918               [simplify in Hletin2;
919                (* FIXME trick *);generalize in match Hletin2;intro;
920                inversion Hletin2
921                  [intros;lapply (inj_head_nat ? ? ? ? H12);
922                   rewrite < Hletin3 in H11;lapply (Hcut1 H11);elim Hletin4
923                  |intros;lapply (inj_tail ? ? ? ? ? H14);rewrite > Hletin3;
924                   assumption]
925               |unfold;intro;apply H8;rewrite < H10;assumption]
926            |rewrite > subst_O_nat;apply in_FV_subst;assumption]
927         |split
928            [unfold;intro;apply H7;apply natinG_or_inH_to_natinGH;right;
929             assumption
930            |unfold;intro;apply H7;apply natinG_or_inH_to_natinGH;left;
931             assumption]]
932      |apply (H2 H6)]]
933 qed.
934            
935 lemma WFE_consG_to_WFT : \forall G.\forall b,X,T.
936                          (WFEnv ((mk_bound b X T)::G)) \to (WFType G T).
937 intros.
938 inversion H
939   [intro;reduce in H1;destruct H1
940   |intros;lapply (inj_head ? ? ? ? H5);lapply (inj_tail ? ? ? ? ? H5);
941    destruct Hletin;rewrite > Hletin1;rewrite > Hcut2;assumption]
942 qed.
943          
944 lemma WFE_consG_WFE_G : \forall G.\forall b.
945                          (WFEnv (b::G)) \to (WFEnv G).
946 intros.
947 inversion H
948   [intro;reduce in H1;destruct H1
949   |intros;lapply (inj_tail ? ? ? ? ? H5);rewrite > Hletin;assumption]
950 qed.
951
952 lemma WFT_swap : \forall u,v,G,T.(WFType G T) \to
953                     (WFType (swap_Env u v G) (swap_Typ u v T)).
954 intros.elim H
955   [simplify;apply WFT_TFree;lapply (natinfv_boundinenv ? ? H1);elim Hletin;
956    elim H2;apply boundinenv_natinfv;apply ex_intro
957      [apply a
958      |apply ex_intro 
959         [apply (swap_Typ u v a1)
960         |apply lookup_swap;assumption]]
961   |simplify;apply WFT_Top
962   |simplify;apply WFT_Arrow
963      [assumption|assumption]
964   |simplify;apply WFT_Forall
965      [assumption
966      |intros;rewrite < (swap_inv u v);
967       cut (\lnot (in_list ? (swap u v X) (fv_type t1)))
968         [cut (\lnot (in_list ? (swap u v X) (fv_env e)))
969            [generalize in match (H4 ? Hcut1 Hcut);simplify;
970             rewrite > subst_type_O_swap;intro;assumption
971            |lapply (in_dom_swap u v (swap u v X) e);elim Hletin;unfold;
972             intros;lapply (H7 H9);rewrite > (swap_inv u v) in Hletin1;
973             apply (H5 Hletin1)] 
974         |generalize in match (in_fv_type_swap u v (swap u v X) t1);intros;
975          elim H7;unfold;intro;lapply (H8 H10);
976          rewrite > (swap_inv u v) in Hletin;apply (H6 Hletin)]]]
977 qed.
978
979 lemma WFE_swap : \forall u,v,G.(WFEnv G) \to (WFEnv (swap_Env u v G)).
980 intros 3.elim G 0
981   [intro;simplify;assumption
982   |intros 2;elim s;simplify;constructor 2
983      [apply H;apply (WFE_consG_WFE_G ? ? H1)
984      |unfold;intro;lapply (in_dom_swap u v n l);elim Hletin;lapply (H4 H2);
985       (* FIXME trick *)generalize in match H1;intro;inversion H1
986         [intros;absurd ((mk_bound b n t)::l = [])
987            [assumption|apply nil_cons]
988         |intros;lapply (inj_head ? ? ? ? H10);lapply (inj_tail ? ? ? ? ? H10);
989          destruct Hletin2;rewrite < Hcut1 in H8;rewrite < Hletin3 in H8;
990          apply (H8 Hletin1)]
991      |apply (WFT_swap u v l t);inversion H1
992         [intro;absurd ((mk_bound b n t)::l = [])
993            [assumption|apply nil_cons]
994         |intros;lapply (inj_head ? ? ? ? H6);lapply (inj_tail ? ? ? ? ? H6);
995          destruct Hletin;rewrite > Hletin1;rewrite > Hcut2;assumption]]]
996 qed.
997
998 (*** some exotic inductions and related lemmas ***) 
999
1000 (* TODO : relocate the following 3 lemmas *)
1001
1002 lemma max_case : \forall m,n.(max m n) = match (leb m n) with
1003       [ false \Rightarrow n
1004       | true \Rightarrow m ].
1005 intros;elim m;simplify;reflexivity;
1006 qed.
1007
1008 lemma not_t_len_lt_SO : \forall T.\lnot (t_len T) < (S O).
1009 intros;elim T
1010   [simplify;unfold;intro;unfold in H;elim (not_le_Sn_n ? H)
1011   |simplify;unfold;intro;unfold in H;elim (not_le_Sn_n ? H)
1012   |simplify;unfold;intro;unfold in H;elim (not_le_Sn_n ? H)
1013   |simplify;unfold;rewrite > max_case;elim (leb (t_len t) (t_len t1))
1014      [simplify in H2;apply H1;apply (trans_lt ? ? ? ? H2);unfold;constructor 1
1015      |simplify in H2;apply H;apply (trans_lt ? ? ? ? H2);unfold;constructor 1]
1016   |simplify;unfold;rewrite > max_case;elim (leb (t_len t) (t_len t1))
1017      [simplify in H2;apply H1;apply (trans_lt ? ? ? ? H2);unfold;constructor 1
1018      |simplify in H2;apply H;apply (trans_lt ? ? ? ? H2);unfold;constructor 1]]
1019 qed.
1020
1021 lemma t_len_gt_O : \forall T.(t_len T) > O.
1022 intro;elim T
1023   [simplify;unfold;unfold;constructor 1
1024   |simplify;unfold;unfold;constructor 1
1025   |simplify;unfold;unfold;constructor 1
1026   |simplify;lapply (max_case (t_len t) (t_len t1));rewrite > Hletin;
1027    elim (leb (t_len t) (t_len t1))
1028      [simplify;unfold;unfold;constructor 2;unfold in H1;unfold in H1;assumption
1029      |simplify;unfold;unfold;constructor 2;unfold in H;unfold in H;assumption]
1030   |simplify;lapply (max_case (t_len t) (t_len t1));rewrite > Hletin;
1031    elim (leb (t_len t) (t_len t1))
1032      [simplify;unfold;unfold;constructor 2;unfold in H1;unfold in H1;assumption
1033      |simplify;unfold;unfold;constructor 2;unfold in H;unfold in H;assumption]]
1034 qed.
1035
1036 lemma Typ_len_ind : \forall P:Typ \to Prop.
1037                        (\forall U.(\forall V.((t_len V) < (t_len U)) \to (P V))
1038                            \to (P U))
1039                        \to \forall T.(P T).
1040 cut (\forall P:Typ \to Prop.
1041         (\forall U.(\forall V.((t_len V) < (t_len U)) \to (P V))
1042             \to (P U))
1043         \to \forall T,n.(n = (t_len T)) \to (P T))                      
1044   [intros;apply (Hcut ? H ? (t_len T));reflexivity
1045   |intros 4;generalize in match T;apply (nat_elim1 n);intros;
1046    generalize in match H2;elim t 
1047      [apply H;intros;simplify in H4;elim (not_t_len_lt_SO ? H4)
1048      |apply H;intros;simplify in H4;elim (not_t_len_lt_SO ? H4)
1049      |apply H;intros;simplify in H4;elim (not_t_len_lt_SO ? H4)
1050      |apply H;intros;apply (H1 (t_len V))
1051         [rewrite > H5;assumption
1052         |reflexivity]
1053      |apply H;intros;apply (H1 (t_len V))
1054         [rewrite > H5;assumption
1055         |reflexivity]]]
1056 qed.
1057
1058 lemma t_len_arrow1 : \forall T1,T2.(t_len T1) < (t_len (Arrow T1 T2)).
1059 intros.simplify.
1060 (* FIXME!!! BUG?!?! *)
1061 cut ((max (t_len T1) (t_len T2)) = match (leb (t_len T1) (t_len T2)) with
1062       [ false \Rightarrow (t_len T2)
1063       | true \Rightarrow (t_len T1) ])
1064   [rewrite > Hcut;cut ((leb (t_len T1) (t_len T2)) = false \lor
1065                        (leb (t_len T1) (t_len T2)) = true)
1066      [lapply (leb_to_Prop (t_len T1) (t_len T2));elim Hcut1
1067         [rewrite > H;rewrite > H in Hletin;simplify;constructor 1
1068         |rewrite > H;rewrite > H in Hletin;simplify;simplify in Hletin;
1069          unfold;apply le_S_S;assumption]
1070      |elim (leb (t_len T1) (t_len T2));auto]
1071   |elim T1;simplify;reflexivity]
1072 qed.
1073
1074 lemma t_len_arrow2 : \forall T1,T2.(t_len T2) < (t_len (Arrow T1 T2)).
1075 intros.simplify.
1076 (* FIXME!!! BUG?!?! *)
1077 cut ((max (t_len T1) (t_len T2)) = match (leb (t_len T1) (t_len T2)) with
1078       [ false \Rightarrow (t_len T2)
1079       | true \Rightarrow (t_len T1) ])
1080   [rewrite > Hcut;cut ((leb (t_len T1) (t_len T2)) = false \lor
1081                        (leb (t_len T1) (t_len T2)) = true)
1082      [lapply (leb_to_Prop (t_len T1) (t_len T2));elim Hcut1
1083         [rewrite > H;rewrite > H in Hletin;simplify;simplify in Hletin;
1084          lapply (not_le_to_lt ? ? Hletin);unfold in Hletin1;unfold;
1085          constructor 2;assumption
1086         |rewrite > H;simplify;unfold;constructor 1]
1087      |elim (leb (t_len T1) (t_len T2));auto]
1088   |elim T1;simplify;reflexivity]
1089 qed.
1090
1091 lemma t_len_forall1 : \forall T1,T2.(t_len T1) < (t_len (Forall T1 T2)).
1092 intros.simplify.
1093 (* FIXME!!! BUG?!?! *)
1094 cut ((max (t_len T1) (t_len T2)) = match (leb (t_len T1) (t_len T2)) with
1095       [ false \Rightarrow (t_len T2)
1096       | true \Rightarrow (t_len T1) ])
1097   [rewrite > Hcut;cut ((leb (t_len T1) (t_len T2)) = false \lor
1098                        (leb (t_len T1) (t_len T2)) = true)
1099      [lapply (leb_to_Prop (t_len T1) (t_len T2));elim Hcut1
1100         [rewrite > H;rewrite > H in Hletin;simplify;constructor 1
1101         |rewrite > H;rewrite > H in Hletin;simplify;simplify in Hletin;
1102          unfold;apply le_S_S;assumption]
1103      |elim (leb (t_len T1) (t_len T2));auto]
1104   |elim T1;simplify;reflexivity]
1105 qed.
1106
1107 lemma t_len_forall2 : \forall T1,T2.(t_len T2) < (t_len (Forall T1 T2)).
1108 intros.simplify.
1109 (* FIXME!!! BUG?!?! *)
1110 cut ((max (t_len T1) (t_len T2)) = match (leb (t_len T1) (t_len T2)) with
1111       [ false \Rightarrow (t_len T2)
1112       | true \Rightarrow (t_len T1) ])
1113   [rewrite > Hcut;cut ((leb (t_len T1) (t_len T2)) = false \lor
1114                        (leb (t_len T1) (t_len T2)) = true)
1115      [lapply (leb_to_Prop (t_len T1) (t_len T2));elim Hcut1
1116         [rewrite > H;rewrite > H in Hletin;simplify;simplify in Hletin;
1117          lapply (not_le_to_lt ? ? Hletin);unfold in Hletin1;unfold;
1118          constructor 2;assumption
1119         |rewrite > H;simplify;unfold;constructor 1]
1120      |elim (leb (t_len T1) (t_len T2));auto]
1121   |elim T1;simplify;reflexivity]
1122 qed.
1123
1124 lemma eq_t_len_TFree_subst : \forall T,n,X.(t_len T) = 
1125                                          (t_len (subst_type_nat T (TFree X) n)).
1126 intro.elim T
1127   [simplify;elim (eqb n n1)
1128      [simplify;reflexivity
1129      |simplify;reflexivity]
1130   |simplify;reflexivity
1131   |simplify;reflexivity
1132   |simplify;lapply (H n X);lapply (H1 n X);rewrite < Hletin;rewrite < Hletin1;
1133    reflexivity
1134   |simplify;lapply (H n X);lapply (H1 (S n) X);rewrite < Hletin;
1135    rewrite < Hletin1;reflexivity]
1136 qed.
1137
1138 lemma swap_env_not_free : \forall u,v,G.(WFEnv G) \to 
1139                                         \lnot (in_list ? u (fv_env G)) \to
1140                                         \lnot (in_list ? v (fv_env G)) \to
1141                                         (swap_Env u v G) = G.
1142 intros 3.elim G 0
1143   [simplify;intros;reflexivity
1144   |intros 2;elim s 0;simplify;intros;lapply (notin_cons ? ? ? ? H2);
1145    lapply (notin_cons ? ? ? ? H3);elim Hletin;elim Hletin1;
1146    lapply (swap_other ? ? ? H4 H6);lapply (WFE_consG_to_WFT ? ? ? ? H1);
1147    cut (\lnot (in_list ? u (fv_type t)))
1148      [cut (\lnot (in_list ? v (fv_type t)))
1149         [lapply (swap_Typ_not_free ? ? ? Hcut Hcut1);
1150          lapply (WFE_consG_WFE_G ? ? H1);
1151          lapply (H Hletin5 H5 H7);
1152          rewrite > Hletin2;rewrite > Hletin4;rewrite > Hletin6;reflexivity
1153         |unfold;intro;apply H7;
1154          apply (fv_WFT ? ? ? Hletin3 H8)] 
1155      |unfold;intro;apply H5;apply (fv_WFT ? ? ? Hletin3 H8)]]
1156 qed.
1157
1158 (*** alternative "constructor" for universal types' well-formedness ***)
1159
1160 lemma WFT_Forall2 : \forall G,X,T,T1,T2.
1161                        (WFEnv G) \to
1162                        (WFType G T1) \to
1163                        \lnot (in_list ? X (fv_type T2)) \to
1164                        \lnot (in_list ? X (fv_env G)) \to
1165                        (WFType ((mk_bound true X T)::G) 
1166                           (subst_type_O T2 (TFree X))) \to
1167                     (WFType G (Forall T1 T2)).
1168 intros.apply WFT_Forall
1169   [assumption
1170   |intros;generalize in match (WFT_swap X X1 ? ? H4);simplify;
1171    rewrite > swap_left;
1172    rewrite > (swap_env_not_free X X1 G H H3 H5);
1173    rewrite > subst_type_O_swap;rewrite > swap_left;
1174    rewrite > (swap_Typ_not_free ? ? T2 H2 H6);
1175    intro;apply (WFT_env_incl ? ? H7);unfold;simplify;intros;assumption]
1176 qed.
1177
1178 (*** lemmas relating subtyping and well-formedness ***)
1179
1180 lemma JS_to_WFE : \forall G,T,U.(JSubtype G T U) \to (WFEnv G).
1181 intros;elim H;assumption.
1182 qed.
1183
1184 lemma JS_to_WFT : \forall G,T,U.(JSubtype G T U) \to ((WFType G T) \land 
1185                                                       (WFType G U)).
1186 intros;elim H
1187   [split [assumption|apply WFT_Top]
1188   |split;apply WFT_TFree;assumption
1189   |split 
1190      [apply WFT_TFree;apply boundinenv_natinfv;apply ex_intro
1191         [apply true | apply ex_intro [apply t1 |assumption]]
1192      |elim H3;assumption]
1193   |elim H2;elim H4;split;apply WFT_Arrow;assumption
1194   |elim H2;split
1195      [lapply (fresh_name ((fv_env e) @ (fv_type t1)));
1196       elim Hletin;cut ((\lnot (in_list ? a (fv_env e))) \land
1197                        (\lnot (in_list ? a (fv_type t1))))
1198         [elim Hcut;apply (WFT_Forall2 ? a t2 ? ? (JS_to_WFE ? ? ? H1) H6 H9 H8);
1199          lapply (H4 ? H8);elim Hletin1;assumption
1200         |split;unfold;intro;apply H7;apply natinG_or_inH_to_natinGH
1201            [right;assumption
1202            |left;assumption]]
1203      |lapply (fresh_name ((fv_env e) @ (fv_type t3)));
1204       elim Hletin;cut ((\lnot (in_list ? a (fv_env e))) \land
1205                        (\lnot (in_list ? a (fv_type t3))))
1206         [elim Hcut;apply (WFT_Forall2 ? a t2 ? ? (JS_to_WFE ? ? ? H1) H5 H9 H8);
1207          lapply (H4 ? H8);elim Hletin1;assumption
1208         |split;unfold;intro;apply H7;apply natinG_or_inH_to_natinGH
1209            [right;assumption
1210            |left;assumption]]]]
1211 qed.
1212
1213 lemma JS_to_WFT1 : \forall G,T,U.(JSubtype G T U) \to (WFType G T).
1214 intros;lapply (JS_to_WFT ? ? ? H);elim Hletin;assumption.
1215 qed.
1216
1217 lemma JS_to_WFT2 : \forall G,T,U.(JSubtype G T U) \to (WFType G U).
1218 intros;lapply (JS_to_WFT ? ? ? H);elim Hletin;assumption.
1219 qed.
1220
1221 (*** lemma relating subtyping and swaps ***)
1222
1223 lemma JS_swap : \forall u,v,G,T,U.(JSubtype G T U) \to
1224                    (JSubtype (swap_Env u v G) (swap_Typ u v T) (swap_Typ u v U)).
1225 intros 6.elim H
1226   [simplify;apply SA_Top
1227      [apply WFE_swap;assumption
1228      |apply WFT_swap;assumption]
1229   |simplify;apply SA_Refl_TVar
1230      [apply WFE_swap;assumption
1231      |unfold in H2;unfold;lapply (in_dom_swap u v n e);elim Hletin;
1232       apply (H3 H2)]
1233   |simplify;apply SA_Trans_TVar
1234      [apply (swap_Typ u v t1)
1235      |apply lookup_swap;assumption
1236      |assumption]
1237   |simplify;apply SA_Arrow;assumption
1238   |simplify;apply SA_All
1239      [assumption
1240      |intros;lapply (H4 (swap u v X))
1241         [simplify in Hletin;rewrite > subst_type_O_swap in Hletin;
1242          rewrite > subst_type_O_swap in Hletin;rewrite > swap_inv in Hletin;
1243          assumption
1244         |unfold;intro;apply H5;unfold;
1245          lapply (in_dom_swap u v (swap u v X) e);
1246          elim Hletin;rewrite > swap_inv in H7;apply H7;assumption]]]
1247 qed.
1248
1249 lemma fresh_WFT : \forall x,G,T.(WFType G T) \to \lnot (in_list ? x (fv_env G))
1250                      \to \lnot (in_list ? x (fv_type T)).
1251 intros;unfold;intro;apply H1;apply (fv_WFT ? ? ? H H2);
1252 qed.
1253
1254 lemma fresh_subst_type_O : \forall x,T1,B,G,T,y.
1255                   (WFType ((mk_bound B x T1)::G) (subst_type_O T (TFree x))) \to
1256                   \lnot (in_list ? y (fv_env G)) \to (x \neq y) \to
1257                   \lnot (in_list ? y (fv_type T)).
1258 intros;unfold;intro;
1259 cut (in_list ? y (fv_env ((mk_bound B x T1) :: G)))
1260   [simplify in Hcut;inversion Hcut
1261      [intros;apply H2;lapply (inj_head_nat ? ? ? ? H5);rewrite < H4 in Hletin;
1262       assumption
1263      |intros;apply H1;rewrite > H6;lapply (inj_tail ? ? ? ? ? H7);
1264       rewrite > Hletin;assumption]
1265   |apply (fv_WFT (subst_type_O T (TFree x)) ? ? H);
1266    apply fv_subst_type_O;assumption]
1267 qed.
1268
1269 (*** alternative "constructor" for subtyping between universal types ***)
1270
1271 lemma SA_All2 : \forall G,S1,S2,T1,T2,X.(JSubtype G T1 S1) \to
1272                    \lnot (in_list ? X (fv_env G)) \to
1273                    \lnot (in_list ? X (fv_type S2)) \to
1274                    \lnot (in_list ? X (fv_type T2)) \to
1275                    (JSubtype ((mk_bound true X T1) :: G)
1276                       (subst_type_O S2 (TFree X))
1277                       (subst_type_O T2 (TFree X))) \to
1278                    (JSubtype G (Forall S1 S2) (Forall T1 T2)).
1279 intros;apply (SA_All ? ? ? ? ? H);intros;
1280 lapply (decidable_eq_nat X X1);elim Hletin
1281   [rewrite < H6;assumption
1282   |elim (JS_to_WFT ? ? ? H);elim (JS_to_WFT ? ? ? H4);
1283    cut (\lnot (in_list ? X1 (fv_type S2)))
1284      [cut (\lnot (in_list ? X1 (fv_type T2)))
1285         [cut (((mk_bound true X1 T1)::G) =
1286               (swap_Env X X1 ((mk_bound true X T1)::G)))
1287            [rewrite > Hcut2;
1288             cut (((subst_type_O S2 (TFree X1)) =
1289                    (swap_Typ X X1 (subst_type_O S2 (TFree X)))) \land
1290                  ((subst_type_O T2 (TFree X1)) =
1291                    (swap_Typ X X1 (subst_type_O T2 (TFree X)))))
1292               [elim Hcut3;rewrite > H11;rewrite > H12;apply JS_swap;
1293                assumption
1294               |split
1295                  [rewrite > (subst_type_O_swap X X1 S2 X);
1296                   rewrite > (swap_Typ_not_free X X1 S2 H2 Hcut); 
1297                   rewrite > swap_left;reflexivity
1298                  |rewrite > (subst_type_O_swap X X1 T2 X);
1299                   rewrite > (swap_Typ_not_free X X1 T2 H3 Hcut1); 
1300                   rewrite > swap_left;reflexivity]]
1301            |simplify;lapply (JS_to_WFE ? ? ? H);
1302             rewrite > (swap_env_not_free X X1 G Hletin1 H1 H5);
1303             cut ((\lnot (in_list ? X (fv_type T1))) \land
1304                  (\lnot (in_list ? X1 (fv_type T1))))
1305               [elim Hcut2;rewrite > (swap_Typ_not_free X X1 T1 H11 H12);
1306                rewrite > swap_left;reflexivity
1307               |split
1308                  [unfold;intro;apply H1;apply (fv_WFT T1 X G H7 H11)
1309                  |unfold;intro;apply H5;apply (fv_WFT T1 X1 G H7 H11)]]]
1310         |unfold;intro;apply H5;lapply (fv_WFT ? X1 ? H10)
1311            [inversion Hletin1
1312               [intros;simplify in H13;lapply (inj_head_nat ? ? ? ? H13);
1313                rewrite < H12 in Hletin2;lapply (H6 Hletin2);elim Hletin3
1314               |intros;simplify in H15;lapply (inj_tail ? ? ? ? ? H15);
1315                rewrite < Hletin2 in H12;rewrite < H14 in H12;lapply (H5 H12);
1316                elim Hletin3]
1317            |rewrite > subst_O_nat;apply in_FV_subst;assumption]]
1318      |unfold;intro;apply H5;lapply (fv_WFT ? X1 ? H9)
1319         [inversion Hletin1
1320            [intros;simplify in H13;lapply (inj_head_nat ? ? ? ? H13);
1321             rewrite < H12 in Hletin2;lapply (H6 Hletin2);elim Hletin3
1322            |intros;simplify in H15;lapply (inj_tail ? ? ? ? ? H15);
1323             rewrite < Hletin2 in H12;rewrite < H14 in H12;lapply (H5 H12);
1324             elim Hletin3]
1325         |rewrite > subst_O_nat;apply in_FV_subst;assumption]]]
1326 qed.