X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2FFsub%2Fpart1a.ma;h=7fb8fed9c14c5433fc8cbd503a8f71e02aedb655;hb=92165c07e82baa982ec06ff9ab325f97659a3900;hp=795546299e842f8c2f2824a7c102db04e548d5d1;hpb=c5608d2704215a7004d23f4801e19df3939f0285;p=helm.git diff --git a/helm/software/matita/library/Fsub/part1a.ma b/helm/software/matita/library/Fsub/part1a.ma index 795546299..7fb8fed9c 100644 --- a/helm/software/matita/library/Fsub/part1a.ma +++ b/helm/software/matita/library/Fsub/part1a.ma @@ -13,65 +13,20 @@ (**************************************************************************) set "baseuri" "cic:/matita/Fsub/part1a/". -include "logic/equality.ma". -include "nat/nat.ma". -include "datatypes/bool.ma". -include "nat/compare.ma". -include "Fsub/util.ma". include "Fsub/defn.ma". (*** Lemma A.1 (Reflexivity) ***) -theorem JS_Refl : \forall T,G.(WFType G T) \to (WFEnv G) \to (JSubtype G T T). -apply Typ_len_ind;intro;elim U - [(* FIXME *) generalize in match H1;intro;inversion H1 - [intros;destruct H6 - |intros;destruct H5 - |*:intros;destruct H9] - |apply (SA_Refl_TVar ? ? H2);(*FIXME*)generalize in match H1;intro; - inversion H1 - [intros;destruct H6;rewrite > Hcut;assumption - |intros;destruct H5 - |*:intros;destruct H9] - |apply (SA_Top ? ? H2 H1) - |cut ((WFType G t) \land (WFType G t1)) - [elim Hcut;apply SA_Arrow - [apply H2 - [apply t_len_arrow1 - |*:assumption] - |apply H2 - [apply t_len_arrow2 - |*:assumption]] - |(*FIXME*)generalize in match H3;intro;inversion H3 - [intros;destruct H8 - |intros;destruct H7 - |intros;destruct H11;rewrite > Hcut;rewrite > Hcut1;split;assumption - |intros;destruct H11]] - |elim (fresh_name ((fv_type t1) @ (fv_env G))); - cut ((\lnot (in_list ? a (fv_type t1))) \land - (\lnot (in_list ? a (fv_env G)))) - [elim Hcut;cut (WFType G t) - [apply (SA_All2 ? ? ? ? ? a ? H7 H6 H6) - [apply H2 - [apply t_len_forall1 - |*:assumption] - |apply H2 - [rewrite > subst_O_nat;rewrite < eq_t_len_TFree_subst; - apply t_len_forall2 - |(*FIXME*)generalize in match H3;intro;inversion H3 - [intros;destruct H11 - |intros;destruct H10 - |intros;destruct H14 - |intros;destruct H14;rewrite < Hcut2 in H11; - rewrite < Hcut3 in H11;rewrite < H13;rewrite < H13 in H11; - apply (H11 ? H7 H6)] - |apply WFE_cons;assumption]] - |(*FIXME*)generalize in match H3;intro;inversion H3 - [intros;destruct H11 - |intros;destruct H10 - |intros;destruct H14 - |intros;destruct H14;rewrite > Hcut1;assumption]] - |split;unfold;intro;apply H5;apply natinG_or_inH_to_natinGH;auto]] +theorem JS_Refl : ∀T,G.WFType G T → WFEnv G → G ⊢ T ⊴ T. +intros 3;elim H + [apply (SA_Refl_TVar l n H2 H1); + |apply (SA_Top l Top H1 (WFT_Top l)); + |apply (SA_Arrow l t t1 t t1 (H2 H5) (H4 H5)) + |apply (SA_All ? ? ? ? ? (H2 H5));intros;apply (H4 X H6) + [intro;apply H6;apply (fv_WFT (Forall t t1) X l) + [apply (WFT_Forall ? ? ? H1 H3) + |simplify;autobatch] + |autobatch]] qed. (* @@ -95,114 +50,168 @@ intros 4;elim H [unfold;intro;apply H8;lapply (incl_bound_fv ? ? H7);apply (Hletin1 ? H9) |apply WFE_cons [1,2:assumption - |lapply (incl_bound_fv ? ? H7);apply (WFT_env_incl ? ? ? ? Hletin1); - apply (JS_to_WFT1 ? ? ? H1)] - |unfold;intros;inversion H9 - [intros;lapply (inj_head ? ? ? ? H11);rewrite > Hletin1;apply in_Base - |intros;lapply (inj_tail ? ? ? ? ? H13);rewrite < Hletin1 in H10; - apply in_Skip;apply (H7 ? H10)]]] + |apply (JS_to_WFT1 ? ? ? Hletin)] + |unfold;intros;elim (in_cons_case ? ? ? ? H9) + [rewrite > H10;apply in_Base + |elim H10;apply (in_Skip ? ? ? ? ? H11);apply (H7 ? H12)]]] qed. +lemma decidable_eq_Typ : \forall S,T:Typ.(S = T) ∨ (S ≠ T). +intro;elim S + [elim T + [elim (decidable_eq_nat n n1) + [rewrite > H;left;reflexivity + |right;intro;destruct H1;apply (H Hcut)] + |2,3:right;intro;destruct H + |*:right;intro;destruct H2] + |elim T + [2:elim (decidable_eq_nat n n1) + [rewrite > H;left;reflexivity + |right;intro;destruct H1;apply (H Hcut)] + |1,3:right;intro;destruct H + |*:right;intro;destruct H2] + |elim T + [3:left;reflexivity + |1,2:right;intro;destruct H + |*:right;intro;destruct H2] + |elim T + [1,2,3:right;intro;destruct H2 + |elim (H t2) + [rewrite > H4;elim (H1 t3) + [rewrite > H5;left;reflexivity + |right;intro;apply H5;destruct H6;assumption] + |right;intro;apply H4;destruct H5;assumption] + |right;intro;destruct H4] + |elim T + [1,2,3:right;intro;destruct H2 + |right;intro;destruct H4 + |elim (H t2) + [rewrite > H4;elim (H1 t3) + [rewrite > H5;left;reflexivity + |right;intro;apply H5;destruct H6;assumption] + |right;intro;apply H4;destruct H5;assumption]]] +qed. + +lemma decidable_eq_bound: ∀b1,b2:bound.(b1 = b2) ∨ (b1 ≠ b2). +intros;elim b1;elim b2;elim (decidable_eq_nat n n1) + [rewrite < H;elim (decidable_eq_Typ t t1) + [rewrite < H1;elim (bool_to_decidable_eq b b3) + [rewrite > H2;left;reflexivity + |right;intro;destruct H3;apply (H2 Hcut)] + |right;intro;destruct H2;apply (H1 Hcut1)] + |right;intro;destruct H1;apply (H Hcut1)] +qed. + (* Lemma A.3 (Transitivity and Narrowing) *) -lemma JS_trans_narrow : \forall n. - (\forall G,T,Q,U. - (t_len Q) \leq n \to (JSubtype G T Q) \to (JSubtype G Q U) \to +lemma JS_trans_narrow : \forall Q. + (\forall G,T,U. + (JSubtype G T Q) \to (JSubtype G Q U) \to (JSubtype G T U)) \land - (\forall G,H,X,P,Q,M,N. - (t_len Q) \leq n \to + (\forall G,H,X,P,M,N. (JSubtype (H @ ((mk_bound true X Q) :: G)) M N) \to (JSubtype G P Q) \to (JSubtype (H @ ((mk_bound true X P) :: G)) M N)). -intro;apply (nat_elim1 n);intros 2; -cut (\forall G,T,Q.(JSubtype G T Q) \to - \forall U.(t_len Q \leq m) \to (JSubtype G Q U) \to (JSubtype G T U)) - [cut (\forall G,M,N.(JSubtype G M N) \to - \forall G1,X,Q,G2,P. - (G = G2 @ ((mk_bound true X Q) :: G1)) \to (t_len Q) \leq m \to - (JSubtype G1 P Q) \to +apply Typ_len_ind;intros 2; +cut (\forall G,T,P. + (JSubtype G T U) \to + (JSubtype G U P) \to + (JSubtype G T P)) + [split + [assumption + |cut (\forall G,M,N.(JSubtype G M N) \to + \forall G1,X,G2,P. + (G = G2 @ ((mk_bound true X U) :: G1)) \to + (JSubtype G1 P U) \to (JSubtype (G2 @ ((mk_bound true X P) :: G1)) M N)) - [split - [intros;apply (Hcut ? ? ? H2 ? H1 H3) - |intros;apply (Hcut1 ? ? ? H3 ? ? ? ? ? ? H2 H4);reflexivity] - |intros 9;cut (incl ? (fv_env (G2 @ ((mk_bound true X Q)::G1))) - (fv_env (G2 @ ((mk_bound true X P)::G1)))) - [intros; -(* [rewrite > H6 in H2;lapply (JS_to_WFT1 ? ? ? H8); - apply (WFE_Typ_subst ? ? ? ? ? ? ? H2 Hletin) *) - generalize in match Hcut1;generalize in match H2; - generalize in match H1;generalize in match H4; - generalize in match G1;generalize in match G2;elim H1 - [apply SA_Top - [rewrite > H9 in H5;lapply (JS_to_WFT1 ? ? ? H7); - apply (WFE_Typ_subst ? ? ? ? ? ? ? H5 Hletin) - |rewrite > H9 in H6;apply (WFT_env_incl ? ? H6);elim l - [simplify;unfold;intros;assumption - |simplify;apply (incl_nat_cons ? ? ? H11)]] - |apply SA_Refl_TVar - [rewrite > H9 in H5;lapply (JS_to_WFT1 ? ? ? H7); - apply (WFE_Typ_subst ? ? ? ? ? ? ? H5 Hletin) - |apply H10;rewrite < H9;assumption] - |elim (decidable_eq_nat X n1) - [apply (SA_Trans_TVar ? ? ? P) - [rewrite < H12;elim l - [simplify;apply in_Base - |simplify;apply in_Skip;assumption] - |lapply (JS_to_WFE ? ? ? H9);rewrite > H10 in Hletin; - rewrite > H10 in H5; - lapply (WFE_bound_bound ? ? ? Q ? Hletin H5) - [lapply (H7 ? ? H8 H6 H10 H11);rewrite > Hletin1 in Hletin2; - apply (Hcut ? ? ? ? ? H3 Hletin2); - lapply (JS_to_WFE ? ? ? Hletin2); - apply (JS_weakening ? ? ? H8 ? Hletin3);unfold;intros; - elim l;simplify;apply in_Skip;assumption - |rewrite > H12;elim l - [simplify;apply in_Base - |simplify;apply in_Skip;assumption]]] - |rewrite > H10 in H5;apply (SA_Trans_TVar ? ? ? t1) - [apply (lookup_env_extends ? ? ? ? ? ? ? ? ? ? H5);unfold; - intro;apply H12;symmetry;assumption - |apply (H7 ? ? H8 H6 H10 H11)]] - |apply SA_Arrow - [apply (H6 ? ? H9 H5 H11 H12) - |apply (H8 ? ? H9 H7 H11 H12)] - |apply SA_All - [apply (H6 ? ? H9 H5 H11 H12) - |intros;apply (H8 ? ? (mk_bound true X1 t2::l) l1) - [unfold;intro;apply H13;rewrite > H11 in H14;apply (H12 ? H14) - |assumption - |apply H7;rewrite > H11;unfold;intro;apply H13;apply (H12 ? H14) - |simplify;rewrite < H11;reflexivity - |simplify;apply incl_nat_cons;assumption]]] - |elim G2 0 - [simplify;unfold;intros;assumption - |intro;elim t 0;simplify;intros;apply incl_nat_cons;assumption]]] - |intros 4;(*generalize in match H1;*)elim H1 + [intros;apply (Hcut1 ? ? ? H2 ? ? ? ? ? H3);reflexivity + |intros;cut (incl ? (fv_env (G2 @ ((mk_bound true X U)::G1))) + (fv_env (G2 @ ((mk_bound true X P)::G1)))) + [intros;generalize in match H2;generalize in match Hcut1; + generalize in match Hcut;generalize in match G2;elim H1 + [apply SA_Top + [rewrite > H8 in H4;lapply (JS_to_WFT1 ? ? ? H3); + apply (WFE_Typ_subst ? ? ? ? ? ? ? H4 Hletin) + |rewrite > H8 in H5;apply (WFT_env_incl ? ? H5 ? H7)] + |apply SA_Refl_TVar + [rewrite > H8 in H4;apply (WFE_Typ_subst ? ? ? ? ? ? ? H4); + apply (JS_to_WFT1 ? ? ? H3) + |rewrite > H8 in H5;apply (H7 ? H5)] + |elim (decidable_eq_nat X n) + [apply (SA_Trans_TVar ? ? ? P) + [rewrite < H10;elim l1 + [simplify;constructor 1 + |simplify;elim (decidable_eq_bound (mk_bound true X P) t2) + [rewrite < H12;apply in_Base + |apply (in_Skip ? ? ? ? ? H12);assumption]] + |apply H7 + [lapply (H6 ? H7 H8 H9);lapply (JS_to_WFE ? ? ? Hletin); + apply (JS_weakening ? ? ? H3 ? Hletin1);unfold;intros; + elim l1 + [simplify; + elim (decidable_eq_bound x (mk_bound true X P)) + [rewrite < H12;apply in_Base + |apply (in_Skip ? ? ? ? ? H12);assumption] + |simplify;elim (decidable_eq_bound x t2) + [rewrite < H13;apply in_Base + |apply (in_Skip ? ? ? ? ? H13);assumption]] + |lapply (WFE_bound_bound true n t1 U ? ? H4) + [apply (JS_to_WFE ? ? ? H5) + |rewrite < Hletin;apply (H6 ? H7 H8 H9) + |rewrite > H9;rewrite > H10;elim l1;simplify + [constructor 1 + |elim (decidable_eq_bound (mk_bound true n U) t2) + [rewrite > H12;apply in_Base + |apply (in_Skip ? ? ? ? ? H12);assumption]]]]] + |apply (SA_Trans_TVar ? ? ? t1) + [rewrite > H9 in H4; + apply (lookup_env_extends ? ? ? ? ? ? ? ? ? ? H4); + unfold;intro;apply H10;symmetry;assumption + |apply (H6 ? H7 H8 H9)]] + |apply SA_Arrow + [apply (H5 ? H8 H9 H10) + |apply (H7 ? H8 H9 H10)] + |apply SA_All + [apply (H5 ? H8 H9 H10) + |intros;apply (H7 ? ? (mk_bound true X1 t2::l1) H8) + [rewrite > H10;cut ((fv_env (l1@(mk_bound true X P::G1))) = + (fv_env (l1@(mk_bound true X U::G1)))) + [unfold;intro;apply H11;rewrite > Hcut2;assumption + |elim l1 + [simplify;reflexivity + |elim t4;simplify;rewrite > H12;reflexivity]] + |simplify;apply (incl_nat_cons ? ? ? H9) + |simplify;rewrite < H10;reflexivity]]] + |cut ((fv_env (G2@(mk_bound true X U::G1))) = + (fv_env (G2@(mk_bound true X P::G1)))) + [rewrite > Hcut1;unfold;intros;assumption + |elim G2 + [simplify;reflexivity + |elim t;simplify;rewrite > H4;reflexivity]]]]] + |intros 4;generalize in match H;elim H1 [inversion H5 [intros;rewrite < H8;apply (SA_Top ? ? H2 H3) |intros;destruct H9 |intros;destruct H10 |*:intros;destruct H11] |assumption - |apply (SA_Trans_TVar ? ? ? ? H2);apply (H4 ? H5 H6) + |apply (SA_Trans_TVar ? ? ? ? H2);apply (H4 H5 H6) |inversion H7 [intros;apply (SA_Top ? ? H8);rewrite < H10;apply WFT_Arrow [apply (JS_to_WFT2 ? ? ? H2) |apply (JS_to_WFT1 ? ? ? H4)] |intros;destruct H11 |intros;destruct H12 - |intros;destruct H13;elim (H (pred m)) - [apply SA_Arrow + |intros;destruct H13;apply SA_Arrow [rewrite > H12 in H2;rewrite < Hcut in H8; - apply (H15 ? ? ? ? ? H8 H2);lapply (t_len_arrow1 t2 t3); - unfold in Hletin;lapply (trans_le ? ? ? Hletin H6); - apply (t_len_pred ? ? Hletin1) + lapply (H6 t2) + [elim Hletin;apply (H15 ? ? ? H8 H2) + |apply (t_len_arrow1 t2 t3)] |rewrite > H12 in H4;rewrite < Hcut1 in H10; - apply (H15 ? ? ? ? ? H4 H10);lapply (t_len_arrow2 t2 t3); - unfold in Hletin;lapply (trans_le ? ? ? Hletin H6); - apply (t_len_pred ? ? Hletin1)] - |apply (pred_m_lt_m ? ? H6)] - |intros;destruct H13] + lapply (H6 t3) + [elim Hletin;apply (H15 ? ? ? H4 H10) + |apply (t_len_arrow2 t2 t3)]] + |intros;destruct H13] |inversion H7 [intros;apply (SA_Top ? ? H8);rewrite < H10;apply WFT_Forall [apply (JS_to_WFT2 ? ? ? H2) @@ -212,65 +221,32 @@ cut (\forall G,T,Q.(JSubtype G T Q) \to |intros;destruct H11 |intros;destruct H12 |intros;destruct H13 - |intros;destruct H13;elim (H (pred m)) - [elim (fresh_name ((fv_env e1) @ (fv_type t1) @ (fv_type t7))); - cut ((\lnot (in_list ? a (fv_env e1))) \land - (\lnot (in_list ? a (fv_type t1))) \land - (\lnot (in_list ? a (fv_type t7)))) - [elim Hcut2;elim H18;clear Hcut2 H18;apply (SA_All2 ? ? ? ? ? a) - [rewrite < Hcut in H8;rewrite > H12 in H2; - apply (H15 ? ? ? ? ? H8 H2);lapply (t_len_forall1 t2 t3); - unfold in Hletin;lapply (trans_le ? ? ? Hletin H6); - apply (t_len_pred ? ? Hletin1) - |5:lapply (H10 ? H20);rewrite > H12 in H5; - lapply (H5 ? H20 (subst_type_O t5 (TFree a))) - [apply (H15 ? ? ? ? ? ? Hletin) - [rewrite < Hcut1;rewrite > subst_O_nat; - rewrite < eq_t_len_TFree_subst; - lapply (t_len_forall2 t2 t3);unfold in Hletin2; - lapply (trans_le ? ? ? Hletin2 H6); - apply (t_len_pred ? ? Hletin3) - |rewrite < Hcut in H8; - apply (H16 e1 (nil ?) a t6 t2 ? ? ? Hletin1 H8); - lapply (t_len_forall1 t2 t3);unfold in Hletin2; - lapply (trans_le ? ? ? Hletin2 H6); - apply (t_len_pred ? ? Hletin3)] - |rewrite > subst_O_nat;rewrite < eq_t_len_TFree_subst; - lapply (t_len_forall2 t2 t3);unfold in Hletin1; - lapply (trans_le ? ? ? Hletin1 H6); - apply (trans_le ? ? ? ? Hletin2);constructor 2; - constructor 1 - |rewrite > Hcut1;rewrite > H12 in H4; - lapply (H4 ? H20);rewrite < Hcut1;apply JS_Refl - [apply (JS_to_WFT2 ? ? ? Hletin1) - |apply (JS_to_WFE ? ? ? Hletin1)]] - |*:assumption] - |split - [split - [unfold;intro;apply H17; - apply (natinG_or_inH_to_natinGH ? (fv_env e1));right; - assumption - |unfold;intro;apply H17; - apply (natinG_or_inH_to_natinGH - ((fv_type t1) @ (fv_type t7)));left; - apply natinG_or_inH_to_natinGH;right;assumption] - |unfold;intro;apply H17; - apply (natinG_or_inH_to_natinGH - ((fv_type t1) @ (fv_type t7)));left; - apply natinG_or_inH_to_natinGH;left;assumption]] - |apply (pred_m_lt_m ? ? H6)]]]] + |intros;destruct H13;subst;apply SA_All + [lapply (H6 t4) + [elim Hletin;apply (H12 ? ? ? H8 H2) + |apply t_len_forall1] + |intros;(*destruct H12;*)subst; + lapply (H6 (subst_type_nat t5 (TFree X) O)) + [elim Hletin;apply H13 + [lapply (H6 t4) + [elim Hletin1;apply (H16 l1 [] X t6); + [simplify;apply H4;assumption + |assumption] + |apply t_len_forall1] + |apply (H10 ? H12)] + |rewrite < eq_t_len_TFree_subst; + apply t_len_forall2]]]]] qed. theorem JS_trans : \forall G,T,U,V.(JSubtype G T U) \to (JSubtype G U V) \to (JSubtype G T V). -intros;elim (JS_trans_narrow (t_len U));apply (H2 ? ? ? ? ? H H1);constructor 1; +intros;elim JS_trans_narrow;autobatch; qed. theorem JS_narrow : \forall G1,G2,X,P,Q,T,U. (JSubtype (G2 @ (mk_bound true X Q :: G1)) T U) \to (JSubtype G1 P Q) \to (JSubtype (G2 @ (mk_bound true X P :: G1)) T U). -intros;elim (JS_trans_narrow (t_len Q));apply (H3 ? ? ? ? ? ? ? ? H H1); -constructor 1; +intros;elim JS_trans_narrow;autobatch; qed.