1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "o-basic_topologies.ma".
18 definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
19 intros; constructor 1;
20 [ apply (λU:Ω \sup S.{a | ∃b:carrbt S. b ∈ U ∧ a ∈ A ? (singleton ? b)});
21 intros; simplify; split; intro; cases H1; cases x; exists [1,3: apply w]
22 split; try assumption; [ apply (. H‡#) | apply (. H \sup -1‡#) ] assumption
23 | intros; split; intros 2; cases f; exists; [1,3: apply w] cases x; split;
24 try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
27 interpretation "downarrow" 'downarrow a = (fun_1 __ (downarrow _) a).
29 definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
30 intros; constructor 1;
31 [ apply (λU,V. ↓U ∩ ↓V);
32 | intros; apply (.= (†H)‡(†H1)); apply refl1]
35 interpretation "ffintersects" 'fintersects U V = (fun1 ___ (ffintersects _) U V).
38 record formal_topology: Type ≝
40 converges: ∀U,V: bt. A ? (U ↓ V) = (A ? U ∧ A ? V)
45 definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
46 intros; constructor 1;
47 [ apply (λb,c:S. (singleton S b) ↓ (singleton S c));
48 | intros; apply (.= (†H)‡(†H1)); apply refl1]
51 interpretation "ffintersects'" 'fintersects U V = (fun1 ___ (ffintersects' _) U V).
53 record formal_map (S,T: formal_topology) : Type ≝
54 { cr:> continuous_relation_setoid S T;
55 C1: ∀b,c. extS ?? cr (b ↓ c) = ext ?? cr b ↓ ext ?? cr c;
59 definition formal_map_setoid: formal_topology → formal_topology → setoid1.
60 intros (S T); constructor 1;
61 [ apply (formal_map S T);
63 [ apply (λf,f1: formal_map S T.f=f1);
64 | simplify; intros 1; apply refl1
65 | simplify; intros 2; apply sym1
66 | simplify; intros 3; apply trans1]]
70 ∀S,T: formal_topology.∀f:formal_map_setoid S T.∀U,V: Ω \sup T.
71 extS ?? f (U ↓ V) = extS ?? f U ↓ extS ?? f V.
73 definition formal_map_composition:
74 ∀o1,o2,o3: formal_topology.
76 (formal_map_setoid o1 o2)
77 (formal_map_setoid o2 o3)
78 (formal_map_setoid o1 o3).
79 intros; constructor 1;
80 [ intros; whd in c c1; constructor 1;
81 [ apply (comp1 BTop ??? c c1);
83 apply (.= (extS_com ??? c c1 ?));
84 apply (.= †(C1 ?????));
85 apply (.= (C1' ?????));
86 apply (.= ((†((extS_singleton ????) \sup -1))‡(†((extS_singleton ????) \sup -1))));
87 apply (.= (extS_com ??????)\sup -1 ‡ (extS_com ??????) \sup -1);
88 apply (.= (extS_singleton ????)‡(extS_singleton ????));
90 | apply (.= (extS_com ??? c c1 ?));
91 apply (.= (†(C2 ???)));
95 change with (comp1 BTop ??? a b = comp1 BTop ??? a' b');
96 apply prop1; assumption]