X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fformal_topology%2Fbasic_topologies_to_o-basic_topologies.ma;h=58b75fb680ececa60755fcb2e77b7105b9245254;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=b1592f6eadf1e4d39ee775408ae75de1620a3509;hpb=5fee26d2afb3a67370c92481bfbfdbd9ebed741e;p=helm.git diff --git a/helm/software/matita/library/formal_topology/basic_topologies_to_o-basic_topologies.ma b/helm/software/matita/library/formal_topology/basic_topologies_to_o-basic_topologies.ma index b1592f6ea..58b75fb68 100644 --- a/helm/software/matita/library/formal_topology/basic_topologies_to_o-basic_topologies.ma +++ b/helm/software/matita/library/formal_topology/basic_topologies_to_o-basic_topologies.ma @@ -49,39 +49,43 @@ definition BTop_to_OBTop: carr3 ((category2_of_category1 BTop) ⇒_\c3 OBTop). | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_comp*);] qed. -(* -alias symbol "eq" (instance 2) = "setoid1 eq". -alias symbol "eq" (instance 1) = "setoid2 eq". -theorem BTop_to_OBTop_faithful: - ∀S,T.∀f,g:arrows2 (category2_of_category1 BTop) S T. - map_arrows2 ?? BTop_to_OBTop ?? f = map_arrows2 ?? BTop_to_OBTop ?? g → f=g. - intros; change with (∀b.A ? (ext ?? f b) = A ? (ext ?? g b)); - apply (POW_faithful); - apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) f \sub \c (⊩ \sub T)); - apply sym2; - apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) g \sub \c (⊩ \sub T)); - apply sym2; - apply e; +theorem BTop_to_OBTop_faithful: faithful2 ?? BTop_to_OBTop. + intros 5; apply (continuous_relation_eq_inv' o1 o2 f g); apply e; qed. include "formal_topology/notation.ma". -theorem BTop_to_OBTop_full: - ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BTop_to_OBTop S T g = f). - intros; +theorem BTop_to_OBTop_full: full2 ?? BTop_to_OBTop. + intros 3 (S T); cases (POW_full (carrbt S) (carrbt T) (Ocont_rel ?? f)) (g Hg); - exists[ + (* cases Hg; *) + exists [ constructor 1; [ apply g - | apply hide; intros; lapply (Oreduced ?? f ? e); + | unfold image_coercion; cases daemon (*apply hide; intros; lapply (Oreduced ?? f ? e); unfold image_coercion; cases Hg; lapply (e3 U) as K; apply (.= K); - apply (.= Hletin); apply rule (†(K^-1)); - | apply hide; intros; lapply (Osaturated ?? f ? e); + apply (.= Hletin); apply rule (†(K^-1)); *) + | cases daemon (* apply hide; intros; lapply (Osaturated ?? f ? e); cases Hg; lapply (e1 U) as K; apply (.= K); - apply (.= Hletin); apply rule (†(K^-1)); + apply (.= Hletin); apply rule (†(K^-1)); *) ] | simplify; unfold BTop_to_OBTop; simplify; + cases Hg; unfold o_continuous_relation_of_continuous_relation_morphism; + simplify; + change with ((orelation_of_relation ?? g)⎻* ∘ oA (o_basic_topology_of_basic_topology S) = + f⎻* ∘ oA (o_basic_topology_of_basic_topology S)); + + + change with (g⎻* ∘ oA (o_basic_topology_of_basic_topology S) = + f⎻* ∘ oA (o_basic_topology_of_basic_topology S)); + apply sym2; whd in T; + intro; + apply trans2; [2: apply sym2; [2: apply Hg; + + whd in ⊢ (?(??%%)???); + apply (.= Hg^-1); unfold o_continuous_relation_of_continuous_relation_morphism; simplify; - cases Hg; whd; simplify; intro; + intro; simplify; + unfold image_coercion; cases Hg; whd; simplify; intro; simplify; qed. *)