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=a5a7eb39b9bad97d52d836ad1401329cff5b58a3;hp=0d997a5af617cf67b7947bc8fc5fe91cfb695ca8;hpb=13088dbb8e54833dfbe2d6c38b08d78fc36452a8;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 0d997a5af..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,63 +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. -(* FIXME -alias symbol "eq" (instance 2) = "setoid1 eq". -alias symbol "eq" (instance 1) = "setoid2 eq". theorem BTop_to_OBTop_faithful: faithful2 ?? BTop_to_OBTop. - intros 5; - whd in e; unfold BTop_to_OBTop in e; simplify in e; - change in match (oA ?) in e with (A o1); - whd in f g; - alias symbol "OR_f_minus_star" (instance 1) = "relation f⎻*". - change in match ((o_continuous_relation_of_continuous_relation_morphism o1 o2 f)⎻* ) in e with ((foo ?? f)⎻* ); - change in match ((o_continuous_relation_of_continuous_relation_morphism o1 o2 g)⎻* ) in e with ((foo ?? g)⎻* ); - whd; whd in o1 o2; - intro b; - alias symbol "OR_f_minus" (instance 1) = "relation f⎻". - letin fb ≝ ((ext ?? f) b); - lapply (e fb); - whd in Hletin:(? ? ? % %); - cases (Hletin); simplify in s s1; - split; - [2: intro; simplify; - lapply depth=0 (s b); intro; apply (Hletin1 ? a ?) - [ 2: whd in f1; - change in Hletin with ((foo ?? f)⎻* - - alias symbol "OR_f_minus_star" (instance 4) = "relation f⎻*". - alias symbol "OR_f_minus_star" (instance 4) = "relation f⎻*". -change in e with - (comp2 SET1 (Ω^o1) (Ω^o1) (Ω^o2) (A o1) (foo o1 o2 f)⎻* =_1 - comp2 SET1 (Ω^o1) (Ω^o1) (Ω^o2) (A o1) (foo o1 o2 g)⎻* - ); - - - - change in e with (comp1 SET (Ω^o1) ?? (A o1) (foo o1 o2 f)⎻* = ((foo o1 o2 g)⎻* ∘ A o1)); - unfold o_continuous_relation_of_continuous_relation_morphism in e; - unfold o_continuous_relation_of_continuous_relation in e; - simplify in e; + 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. *)