]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/POPLmark/Fsub/defn.ma
New version.
[helm.git] / helm / software / matita / contribs / POPLmark / Fsub / defn.ma
index 3bd8d2bdcdf75f2d83be67bd2077fc2b47c4948d..184166ed9346616940e739b222dd408a327e623b 100644 (file)
@@ -12,7 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/Fsub/defn".
 include "Fsub/util.ma".
 
 (*** representation of Fsub types ***)  
@@ -44,16 +43,6 @@ let rec subst_type_nat T U i \def
     | (Arrow T1 T2) \Rightarrow (Arrow (subst_type_nat T1 U i) (subst_type_nat T2 U i))
     | (Forall T1 T2) \Rightarrow (Forall (subst_type_nat T1 U i) (subst_type_nat T2 U (S i))) ].
 
-(*** height of T's syntactic tree ***)
-
-let rec t_len T \def
-  match T with
-     [(TVar n) \Rightarrow (S O)
-     |(TFree X) \Rightarrow (S O)
-     |Top \Rightarrow (S O)
-     |(Arrow T1 T2) \Rightarrow (S (max (t_len T1) (t_len T2)))
-     |(Forall T1 T2) \Rightarrow (S (max (t_len T1) (t_len T2)))].
-
 (*** definitions about lists ***)
 
 definition fv_env : (list bound) \to (list nat) \def
@@ -147,11 +136,6 @@ notation "hvbox(S [# n ↦ T])"
 interpretation "subst bound var" 'substvar S T n =
   (cic:/matita/Fsub/defn/subst_type_nat.con S T n).  
 
-notation "hvbox(|T|)"
-  non associative with precedence 30 for @{ 'tlen $T }.
-interpretation "type length" 'tlen T =
-  (cic:/matita/Fsub/defn/t_len.con T).  
-
 notation "hvbox(!X ⊴ T)"
   non associative with precedence 60 for @{ 'subtypebound $X $T }.
 interpretation "subtyping bound" 'subtypebound X T =
@@ -168,18 +152,19 @@ intros 2;elim G
   [elim H;elim H1;lapply (not_in_list_nil ? ? H2);elim Hletin
   |elim H1;elim H2;elim (in_list_cons_case ? ? ? ? H3)
      [rewrite < H4;simplify;apply in_list_head
-     |simplify;apply in_list_cons;apply H;apply (ex_intro ? ? a);
-      apply (ex_intro ? ? a1);assumption]]
+     |simplify;apply in_list_cons;apply H;apply (ex_intro ? ? a1);
+      apply (ex_intro ? ? a2);assumption]]
 qed.
 
 lemma natinfv_boundinenv : \forall x,G.(in_list ? x (fv_env G)) \to
                               \exists B,T.(in_list ? (mk_bound B x T) G).
 intros 2;elim G 0
   [simplify;intro;lapply (not_in_list_nil ? ? H);elim Hletin
-  |intros 3;elim t;simplify in H1;elim (in_list_cons_case ? ? ? ? H1)
-     [rewrite < H2;apply (ex_intro ? ? b);apply (ex_intro ? ? t1);apply in_list_head
-     |elim (H H2);elim H3;apply (ex_intro ? ? a);
-      apply (ex_intro ? ? a1);apply in_list_cons;assumption]]
+  |intros 3;
+   elim a;simplify in H1;elim (in_list_cons_case ? ? ? ? H1)
+     [rewrite < H2;apply (ex_intro ? ? b);apply (ex_intro ? ? t);apply in_list_head
+     |elim (H H2);elim H3;apply (ex_intro ? ? a1);
+      apply (ex_intro ? ? a2);apply in_list_cons;assumption]]
 qed.
 
 lemma incl_bound_fv : \forall l1,l2.(incl ? l1 l2) \to 
@@ -215,7 +200,7 @@ lemma fv_env_extends : \forall H,x,B,C,T,U,G.
                           (fv_env (H @ ((mk_bound B x T) :: G))) = 
                           (fv_env (H @ ((mk_bound C x U) :: G))).
 intros;elim H
-  [simplify;reflexivity|elim t;simplify;rewrite > H1;reflexivity]
+  [simplify;reflexivity|elim a;simplify;rewrite > H1;reflexivity]
 qed.
 
 lemma lookup_env_extends : \forall G,H,B,C,D,T,U,V,x,y.
@@ -252,12 +237,12 @@ cut (\forall l:(list nat).\exists n.\forall m.
   |intros;elim l
     [apply (ex_intro ? ? O);intros;unfold;intro;elim (not_in_list_nil ? ? H1)
     |elim H;
-     apply (ex_intro ? ? (S (max a t))).
+     apply (ex_intro ? ? (S (max a1 a))).
      intros.unfold. intro.
      elim (in_list_cons_case ? ? ? ? H3)
       [rewrite > H4 in H2.autobatch
       |elim H4
-         [apply (H1 m ? H4).apply (trans_le ? (max a t));autobatch
+         [apply (H1 m ? H4).apply (trans_le ? (max a1 a));autobatch
          |assumption]]]]
 qed.
 
@@ -324,7 +309,7 @@ intros 7;elim H 0
   [simplify;intros;(*FIXME*)generalize in match H1;intro;inversion H1;intros
      [lapply (nil_cons ? G (mk_bound B x T));elim (Hletin H4)
      |destruct H8;apply (WFE_cons ? ? ? ? H4 H6 H2)]
-  |intros;simplify;generalize in match H2;elim t;simplify in H4;
+  |intros;simplify;generalize in match H2;elim a;simplify in H4;
    inversion H4;intros
      [destruct H5
      |destruct H9;apply WFE_cons
@@ -376,6 +361,7 @@ intros.inversion H;intros
   |destruct H5|*:destruct H6]
 qed.
 
+(*
 (* elim vs inversion *)
 lemma JS_trans_TFree: ∀G,T,X.G ⊢ T ⊴ (TFree X) →
   ∀U.G ⊢ (TFree X) ⊴ U → G ⊢ T ⊴ U.
@@ -387,7 +373,8 @@ intros 4.cut (∀Y.TFree Y = TFree X → ∀U.G ⊢ (TFree Y) ⊴ U → G ⊢ T
     |apply (SA_Trans_TVar ? ? ? ? H1);apply (H3 Y);assumption
     |*:destruct H5]]
 qed.
+*)
 
 lemma fv_append : ∀G,H.fv_env (G @ H) = (fv_env G @ fv_env H).
 intro;elim G;simplify;autobatch paramodulation;
-qed.
\ No newline at end of file
+qed.