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 "formal_topology/o-basic_pairs.ma".
16 include "formal_topology/o-saturations.ma".
18 definition A : ∀b:OBP. unary_morphism1 (Oform b) (Oform b).
19 intros; constructor 1;
20 [ apply (λx.□⎽b (Ext⎽b x));
21 | intros; apply (†(†e));]
24 lemma down_p : ∀S:SET1.∀I:SET.∀u:S ⇒_1 S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a =_1 a'→u (c a) =_1 u (c a').
25 intros; apply (†(†e));
28 record Oconcrete_space : Type[2] ≝
30 (*distr : is_distributive (form bp);*)
31 Odownarrow: unary_morphism1 (Oform Obp) (Oform Obp);
32 Odownarrow_is_sat: is_o_saturation ? Odownarrow;
34 (Ext⎽Obp q1 ∧ (Ext⎽Obp q2)) = (Ext⎽Obp ((Odownarrow q1) ∧ (Odownarrow q2)));
35 Oall_covered: Ext⎽Obp (oa_one (Oform Obp)) = oa_one (Oconcr Obp);
36 Oil2: ∀I:SET.∀p:arrows2 SET1 I (Oform Obp).
37 Odownarrow (∨ { x ∈ I | Odownarrow (p x) | down_p ???? }) =
38 ∨ { x ∈ I | Odownarrow (p x) | down_p ???? };
39 Oil1: ∀q.Odownarrow (A ? q) = A ? q
42 interpretation "o-concrete space downarrow" 'downarrow x =
43 (fun11 ?? (Odownarrow ?) x).
45 definition Obinary_downarrow :
46 ∀C:Oconcrete_space.binary_morphism1 (Oform C) (Oform C) (Oform C).
47 intros; constructor 1;
48 [ intros; apply (↓ c ∧ ↓ c1);
50 alias symbol "prop2" = "prop21".
51 alias symbol "prop1" = "prop11".
55 interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 ? ? ? (Obinary_downarrow ?) a b).
57 record Oconvergent_relation_pair (CS1,CS2: Oconcrete_space) : Type[2] ≝
58 { Orp:> arrows2 ? CS1 CS2;
60 ∀b,c. eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (Orp\sub\f⎻ b ↓ Orp\sub\f⎻ c));
61 Orespects_all_covered:
62 eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (oa_one (Oform CS2))))
63 (Ext⎽CS1 (oa_one (Oform CS1)))
66 definition Oconvergent_relation_space_setoid: Oconcrete_space → Oconcrete_space → setoid2.
69 [ apply (Oconvergent_relation_pair c c1)
72 apply (Orelation_pair_equality c c1 c2 c3);
73 | intros 1; apply refl2;
74 | intros 2; apply sym2;
75 | intros 3; apply trans2]]
78 definition Oconvergent_relation_space_of_Oconvergent_relation_space_setoid:
79 ∀CS1,CS2.carr2 (Oconvergent_relation_space_setoid CS1 CS2) →
80 Oconvergent_relation_pair CS1 CS2 ≝ λP,Q,c.c.
81 coercion Oconvergent_relation_space_of_Oconvergent_relation_space_setoid.
83 definition Oconvergent_relation_space_composition:
84 ∀o1,o2,o3: Oconcrete_space.
86 (Oconvergent_relation_space_setoid o1 o2)
87 (Oconvergent_relation_space_setoid o2 o3)
88 (Oconvergent_relation_space_setoid o1 o3).
89 intros; constructor 1;
90 [ intros; whd in t t1 ⊢ %;
94 change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
95 alias symbol "trans" = "trans1".
96 apply (.= († (Orespects_converges : ?)));
97 apply (Orespects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2));
98 | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (Oform o3)))));
99 apply (.= (†(Orespects_all_covered :?)));
100 apply rule (Orespects_all_covered ?? c);]
102 change with (b ∘ a = b' ∘ a');
103 change in e with (Orp ?? a = Orp ?? a');
104 change in e1 with (Orp ?? b = Orp ?? b');
108 definition OCSPA: category2.
110 [ apply Oconcrete_space
111 | apply Oconvergent_relation_space_setoid
112 | intro; constructor 1;
114 | intros; apply refl1;
116 | apply Oconvergent_relation_space_composition
117 | intros; simplify; whd in a12 a23 a34;
118 change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
121 change with (a ∘ id2 OBP o1 = a);
122 apply (id_neutral_right2 : ?);
124 change with (id2 ? o2 ∘ a = a);
125 apply (id_neutral_left2 : ?);]
128 definition Oconcrete_space_of_OCSPA : objs2 OCSPA → Oconcrete_space ≝ λx.x.
129 coercion Oconcrete_space_of_OCSPA.
131 definition Oconvergent_relation_space_setoid_of_arrows2_OCSPA :
132 ∀P,Q. arrows2 OCSPA P Q → Oconvergent_relation_space_setoid P Q ≝ λP,Q,x.x.
133 coercion Oconvergent_relation_space_setoid_of_arrows2_OCSPA.