]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/formal_topology/basic_topologies_to_o-basic_topologies.ma
0d997a5af617cf67b7947bc8fc5fe91cfb695ca8
[helm.git] / helm / software / matita / library / formal_topology / basic_topologies_to_o-basic_topologies.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "formal_topology/basic_topologies.ma".
16 include "formal_topology/o-basic_topologies.ma".
17 include "formal_topology/relations_to_o-algebra.ma".
18
19 definition o_basic_topology_of_basic_topology: basic_topology → Obasic_topology.
20  intros (b); constructor 1;
21   [ apply (POW' b) | apply (A b) | apply (J b);
22   | apply (A_is_saturation b) | apply (J_is_reduction b) | apply (compatibility b) ]
23 qed.
24
25 definition o_continuous_relation_of_continuous_relation:
26  ∀BT1,BT2.continuous_relation BT1 BT2 →
27   Ocontinuous_relation (o_basic_topology_of_basic_topology BT1) (o_basic_topology_of_basic_topology BT2).
28  intros (BT1 BT2 c); constructor 1;
29   [ apply (orelation_of_relation ?? c) | apply (reduced ?? c) | apply (saturated ?? c) ]
30 qed.
31
32 axiom daemon: False.
33
34 lemma o_continuous_relation_of_continuous_relation_morphism :
35   ∀S,T:category2_of_category1 BTop.
36   unary_morphism2 (arrows2 (category2_of_category1 BTop) S T)
37    (arrows2 OBTop (o_basic_topology_of_basic_topology S) (o_basic_topology_of_basic_topology T)).
38 intros (S T);
39    constructor 1;
40      [ apply (o_continuous_relation_of_continuous_relation S T);
41      | cases daemon (*apply (o_relation_pair_of_relation_pair_is_morphism S T)*)]
42 qed.
43
44 definition BTop_to_OBTop: carr3 ((category2_of_category1 BTop) ⇒_\c3 OBTop).
45  constructor 1;
46   [ apply o_basic_topology_of_basic_topology;
47   | intros; apply o_continuous_relation_of_continuous_relation_morphism;
48   | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_id*);
49   | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_comp*);]
50 qed.
51
52 (* FIXME
53 alias symbol "eq" (instance 2) = "setoid1 eq".
54 alias symbol "eq" (instance 1) = "setoid2 eq".
55 theorem BTop_to_OBTop_faithful: faithful2 ?? BTop_to_OBTop.
56  intros 5;
57  whd in e; unfold BTop_to_OBTop in e; simplify in e;
58  change in match (oA ?) in e with (A o1);
59  whd in f g;
60  alias symbol "OR_f_minus_star" (instance 1) = "relation f⎻*".
61  change in match ((o_continuous_relation_of_continuous_relation_morphism o1 o2 f)⎻* ) in e with ((foo ?? f)⎻* );
62  change in match ((o_continuous_relation_of_continuous_relation_morphism o1 o2 g)⎻* ) in e with ((foo ?? g)⎻* );
63  whd; whd in o1 o2;
64  intro b;
65  alias symbol "OR_f_minus" (instance 1) = "relation f⎻".
66  letin fb ≝ ((ext ?? f) b);
67  lapply (e fb);
68  whd in Hletin:(? ? ? % %);
69  cases (Hletin); simplify in s s1;
70  split;
71   [2: intro; simplify;
72     lapply depth=0 (s b); intro; apply (Hletin1 ? a ?)
73      [ 2: whd in f1;
74  change in Hletin with ((foo ?? f)⎻*
75  
76  alias symbol "OR_f_minus_star" (instance 4) = "relation f⎻*".
77  alias symbol "OR_f_minus_star" (instance 4) = "relation f⎻*".
78 change in e with
79  (comp2 SET1 (Ω^o1) (Ω^o1) (Ω^o2) (A o1) (foo o1 o2 f)⎻* =_1
80   comp2 SET1 (Ω^o1) (Ω^o1) (Ω^o2) (A o1) (foo o1 o2 g)⎻*
81  );
82  
83  
84  
85  change in e with (comp1 SET (Ω^o1) ?? (A o1) (foo o1 o2 f)⎻*  = ((foo o1 o2 g)⎻* ∘ A o1));
86  unfold o_continuous_relation_of_continuous_relation_morphism in e;
87  unfold o_continuous_relation_of_continuous_relation in e;
88  simplify in e;
89 qed.
90
91 include "formal_topology/notation.ma".
92
93 theorem BTop_to_OBTop_full: 
94    ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BTop_to_OBTop S T g = f).
95  intros;
96  cases (POW_full (carrbt S) (carrbt T) (Ocont_rel ?? f)) (g Hg);
97  exists[
98    constructor 1;
99     [ apply g
100     | apply hide; intros; lapply (Oreduced ?? f ? e);
101       cases Hg; lapply (e3 U) as K; apply (.= K);
102       apply (.= Hletin); apply rule (†(K^-1));
103     | apply hide; intros; lapply (Osaturated ?? f ? e);
104       cases Hg; lapply (e1 U) as K; apply (.= K);
105       apply (.= Hletin); apply rule (†(K^-1));
106     ]
107  | simplify; unfold BTop_to_OBTop; simplify;
108    unfold o_continuous_relation_of_continuous_relation_morphism; simplify;
109    cases Hg; whd; simplify; intro; 
110 qed.
111 *)