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/basic_pairs.ma".
17 record concrete_space : Type ≝
19 converges: ∀a: concr bp.∀U,V: form bp. a ⊩ U → a ⊩ V → a ⊩ (U ↓ V);
20 all_covered: ∀x: concr bp. x ⊩ form bp
23 definition bp': concrete_space → basic_pair ≝ λc.bp c.
27 record convergent_relation_pair (CS1,CS2: concrete_space) : Type ≝
28 { rp:> arrows1 ? CS1 CS2;
31 extS ?? rp \sub\c (BPextS CS2 (b ↓ c)) =
32 BPextS CS1 ((extS ?? rp \sub\f b) ↓ (extS ?? rp \sub\f c));
34 extS ?? rp\sub\c (BPextS CS2 (form CS2)) = BPextS CS1 (form CS1)
37 definition rp' : ∀CS1,CS2. convergent_relation_pair CS1 CS2 → relation_pair CS1 CS2 ≝
38 λCS1,CS2,c. rp CS1 CS2 c.
42 definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid1.
45 [ apply (convergent_relation_pair c c1)
48 apply (relation_pair_equality c c1 c2 c3);
49 | intros 1; apply refl1;
50 | intros 2; apply sym1;
51 | intros 3; apply trans1]]
54 definition rp'': ∀CS1,CS2.convergent_relation_space_setoid CS1 CS2 → arrows1 BP CS1 CS2 ≝
59 definition convergent_relation_space_composition:
60 ∀o1,o2,o3: concrete_space.
62 (convergent_relation_space_setoid o1 o2)
63 (convergent_relation_space_setoid o2 o3)
64 (convergent_relation_space_setoid o1 o3).
65 intros; constructor 1;
66 [ intros; whd in c c1 ⊢ %;
68 [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption]
70 change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
71 change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? (? ? ? (? ? ? %) ?) ?)))
72 with (c1 \sub \f ∘ c \sub \f);
73 change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? ? (? ? ? (? ? ? %) ?))))
74 with (c1 \sub \f ∘ c \sub \f);
75 apply (.= (extS_com ??????));
76 apply (.= (†(respects_converges ?????)));
77 apply (.= (respects_converges ?????));
78 apply (.= (†(((extS_com ??????) \sup -1)‡(extS_com ??????)\sup -1)));
80 | change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
81 apply (.= (extS_com ??????));
82 apply (.= (†(respects_all_covered ???)));
83 apply (.= respects_all_covered ???);
86 change with (b ∘ a = b' ∘ a');
87 change in H with (rp'' ?? a = rp'' ?? a');
88 change in H1 with (rp'' ?? b = rp ?? b');
93 definition CSPA: category1.
95 [ apply concrete_space
96 | apply convergent_relation_space_setoid
97 | intro; constructor 1;
101 apply (.= (equalset_extS_id_X_X ??));
102 apply (.= (†((equalset_extS_id_X_X ??)\sup -1‡
103 (equalset_extS_id_X_X ??)\sup -1)));
105 | apply (.= (equalset_extS_id_X_X ??));
107 | apply convergent_relation_space_composition
109 change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
113 change with (a ∘ id1 ? o1 = a);
114 apply (.= id_neutral_right1 ????);
117 change with (id1 ? o2 ∘ a = a);
118 apply (.= id_neutral_left1 ????);