X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2FPOPLmark%2FFsub%2Fpart1a_inversion.ma;h=c1e9090974efb131486b0017b49c1ac1e377743f;hb=233826389b4c0c4192c1eb1cacc8cfa99b2750f4;hp=f96e07679c4b4be9a406cc8a0ec936d4f3f5c6cf;hpb=ef3e622c49ce8a0478c3ef1326d4f179aff3d1ed;p=helm.git diff --git a/helm/software/matita/contribs/POPLmark/Fsub/part1a_inversion.ma b/helm/software/matita/contribs/POPLmark/Fsub/part1a_inversion.ma index f96e07679..c1e909097 100644 --- a/helm/software/matita/contribs/POPLmark/Fsub/part1a_inversion.ma +++ b/helm/software/matita/contribs/POPLmark/Fsub/part1a_inversion.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/Fsub/part1a_inversion/". include "Fsub/defn.ma". (*** Lemma A.1 (Reflexivity) ***) @@ -23,7 +22,8 @@ intros 3.elim H |apply (SA_Arrow ? ? ? ? ? (H2 H5) (H4 H5)) |apply (SA_All ? ? ? ? ? (H2 H5));intros;apply (H4 ? H6) [intro;apply H6;apply (fv_WFT ? ? ? (WFT_Forall ? ? ? H1 H3)); - simplify;autobatch + simplify; + autobatch; |autobatch]] qed. @@ -43,8 +43,8 @@ intros 4;elim H [unfold;intro;apply H8;apply (incl_bound_fv ? ? H7 ? H9) |apply (WFE_cons ? ? ? ? H6 H8);autobatch |unfold;intros;inversion H9;intros - [destruct H11;apply in_Base - |destruct H13;apply in_Skip;apply (H7 ? H10)]]] + [destruct H11;apply in_list_head + |destruct H13;apply in_list_cons;apply (H7 ? H10)]]] qed. theorem narrowing:∀X,G,G1,U,P,M,N. @@ -89,9 +89,7 @@ lemma JSubtype_Arrow_inv: G ⊢ T1 ⊴ (Arrow T2 T3) → P G T1. intros; generalize in match (refl_eq ? (Arrow T2 T3)); - change in ⊢ (% → ?) with (Arrow T2 T3 = Arrow T2 T3); generalize in match (refl_eq ? G); - change in ⊢ (% → ?) with (G = G); elim H2 in ⊢ (? ? ? % → ? ? ? % → %); [1,2: destruct H6 |5: destruct H8 @@ -105,37 +103,59 @@ lemma JSubtype_Arrow_inv: ] qed. +lemma JSubtype_Forall_inv: + ∀G:list bound.∀T1,T2,T3:Typ. + ∀P:list bound → Typ → Prop. + (∀n,t1. + (mk_bound true n t1) ∈ G → G ⊢ t1 ⊴ (Forall T2 T3) → P G t1 → P G (TFree n)) → + (∀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) + → P G (Forall t t1)) → + G ⊢ T1 ⊴ (Forall T2 T3) → P G T1. + intros; + generalize in match (refl_eq ? (Forall T2 T3)); + generalize in match (refl_eq ? G); + elim H2 in ⊢ (? ? ? % → ? ? ? % → %); + [1,2: destruct H6 + |4: destruct H8 + | lapply (H5 H6 H7); destruct; clear H5; + apply H; + assumption + | destruct; + clear H4 H6; + apply H1; + assumption + ] +qed. + + lemma JS_trans_prova: ∀T,G1.WFType G1 T → ∀G,R,U.incl ? (fv_env G1) (fv_env G) → G ⊢ R ⊴ T → G ⊢ T ⊴ U → G ⊢ R ⊴ U. intros 3;elim H;clear H; try autobatch; [rewrite > (JSubtype_Top ? ? H3);autobatch |apply (JSubtype_Arrow_inv ? ? ? ? ? ? ? H6); intros; [ autobatch - | inversion H7;intros; destruct; autobatch depth=4 width=4 size=9 - ] - |generalize in match H7;generalize in match H4;generalize in match H2; - generalize in match H5;clear H7 H4 H2 H5; - generalize in match (refl_eq ? (Forall t t1));elim H6 in ⊢ (? ? ? %→%);destruct; - [apply (SA_Trans_TVar ? ? ? ? H);apply (H4 ? H7 H8 H9 H10);reflexivity - |inversion H11;intros;destruct; - [apply SA_Top - [autobatch - |apply WFT_Forall - [autobatch - |intros;lapply (H4 ? H13);autobatch]] - |apply SA_All - [autobatch paramodulation - |intros;apply (H10 X) - [intro;apply H15;apply H8;assumption - |intro;apply H15;apply H8;apply (WFT_to_incl ? ? ? H3); + | inversion H7;intros; destruct; autobatch depth=4 width=4 size=9] + |apply (JSubtype_Forall_inv ? ? ? ? ? ? ? H6); intros; + [ autobatch + | inversion H7;intros; destruct; + [ apply SA_Top + [ assumption + | apply WFT_Forall; + [ autobatch + | intros;lapply (H8 ? H11); + autobatch]] + | apply SA_All + [ autobatch + | intros;apply (H4 X); + [intro;apply H13;apply H5;assumption + |intro;apply H13;apply H5;apply (WFT_to_incl ? ? ? H3); assumption |simplify;autobatch - |apply (narrowing X (mk_bound true X t::l1) - ? ? ? ? ? H7 ? ? []) - [intros;apply H9 - [unfold;intros;lapply (H8 ? H17);rewrite > fv_append; + |apply (narrowing X (mk_bound true X t::G) ? ? ? ? ? H9 ? ? []) + [intros;apply H2 + [unfold;intros;lapply (H5 ? H15);rewrite > fv_append; autobatch - |apply (JS_weakening ? ? ? H7) + |apply (JS_weakening ? ? ? H9) [autobatch |unfold;intros;autobatch] |assumption]