X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fapply_functor.ma;h=789988194f6ade4c616334b96bf5157b77a71912;hb=3e5c359c75874748cfed8a9046031b62396e0e6d;hp=7b10b4b9fc98f05bd0736344545809630c9f89c4;hpb=427bbecb65191dc43037c8d62dbdebeccf7c1c29;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma b/helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma index 7b10b4b9f..789988194 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma @@ -23,11 +23,11 @@ record Fo (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) : Type2 ≝ { notation "ℱ\sub 1 x" non associative with precedence 60 for @{'F1 $x}. notation > "ℱ_1" non associative with precedence 90 for @{F1 ???}. -interpretation "F1" 'F1 x = (F1 ___ x). +interpretation "F1" 'F1 x = (F1 ??? x). notation "ℱ\sub 2 x" non associative with precedence 60 for @{'F2 $x}. notation > "ℱ_2" non associative with precedence 90 for @{F2 ???}. -interpretation "F2" 'F2 x = (F2 ___ x). +interpretation "F2" 'F2 x = (F2 ??? x). lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F. arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) → @@ -44,11 +44,11 @@ record Fm_c (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) (X,Y:Fo ?? F) : Type2 ≝ { notation "ℳ\sub 1 x" non associative with precedence 60 for @{'Fm1 $x}. notation > "ℳ_1" non associative with precedence 90 for @{Fm1 ?????}. -interpretation "Fm1" 'Fm1 x = (Fm1 _____ x). +interpretation "Fm1" 'Fm1 x = (Fm1 ????? x). notation "ℳ\sub 2 x" non associative with precedence 60 for @{'Fm2 $x}. notation > "ℳ_2" non associative with precedence 90 for @{Fm2 ?????}. -interpretation "Fm2" 'Fm2 x = (Fm2 _____ x). +interpretation "Fm2" 'Fm2 x = (Fm2 ????? x). definition Fm : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2. @@ -96,3 +96,27 @@ constructor 1; | intros; apply (id_neutral_right2 C2 ?? (ℳ_2 a)); | intros; apply (id_neutral_left2 C2 ?? (ℳ_2 a));] qed. + +definition faithful ≝ + λC1,C2.λF:arrows3 CAT2 C1 C2.∀S,T.∀f,g:arrows2 C1 S T. + map_arrows2 ?? F ?? f = map_arrows2 ?? F ?? g → f=g. + +definition Ylppa : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2. + faithful ?? F → let rC2 ≝ Apply ?? F in arrows3 CAT2 rC2 C1. +intros; constructor 1; +[ intro; apply (ℱ_1 o); +| intros; constructor 1; + [ intros; apply (ℳ_1 c); + | apply hide; intros; apply f; change in e with (ℳ_2 a = ℳ_2 a'); + lapply (FmP ????? a) as H1; lapply (FmP ????? a') as H2; + cut (REW ????? (map_arrows2 ?? F ?? (ℳ_1 a)) = + REW ????? (map_arrows2 ?? F ?? (ℳ_1 a')));[2: + apply (.= H1); apply (.= e); apply H2^-1;] + clear H1 H2 e; cases S in a a' Hcut; cases T; + cases H; cases H1; simplify; intros; assumption;] +| intro; apply rule #; +| intros; simplify; apply rule #;] +qed. + + +