]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
Sambin's result holds trivially since most of the fields of the objects of
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.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 "o-basic_pairs.ma".
16 include "o-saturations.ma".
17
18 definition A : ∀b:BP. unary_morphism1 (form b) (form b).
19 intros; constructor 1;
20  [ apply (λx.□_b (Ext⎽b x));
21  | intros; apply  (†(†e));]
22 qed.
23
24 lemma down_p : ∀S:SET1.∀I:SET.∀u:S⇒S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
25 intros; apply (†(†e));
26 qed.
27
28 record concrete_space : Type2 ≝
29  { bp:> BP;
30    (*distr : is_distributive (form bp);*)
31    downarrow: unary_morphism1 (form bp) (form bp);
32    downarrow_is_sat: is_o_saturation ? downarrow;
33    converges: ∀q1,q2.
34      (Ext⎽bp q1 ∧ (Ext⎽bp q2)) = (Ext⎽bp ((downarrow q1) ∧ (downarrow q2)));
35    all_covered: Ext⎽bp (oa_one (form bp)) = oa_one (concr bp);
36    il2: ∀I:SET.∀p:arrows2 SET1 I (form bp).
37      downarrow (∨ { x ∈ I | downarrow (p x) | down_p ???? }) =
38      ∨ { x ∈ I | downarrow (p x) | down_p ???? };
39    il1: ∀q.downarrow (A ? q) = A ? q
40  }.
41
42 interpretation "o-concrete space downarrow" 'downarrow x = 
43   (fun11 __ (downarrow _) x).
44
45 definition binary_downarrow : 
46   ∀C:concrete_space.binary_morphism1 (form C) (form C) (form C).
47 intros; constructor 1;
48 [ intros; apply (↓ c ∧ ↓ c1);
49 | intros;
50   alias symbol "prop2" = "prop21".
51   alias symbol "prop1" = "prop11".
52   apply ((†e)‡(†e1));]
53 qed.
54
55 interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 _ _ _ (binary_downarrow _) a b).
56
57 record convergent_relation_pair (CS1,CS2: concrete_space) : Type2 ≝
58  { rp:> arrows2 ? CS1 CS2;
59    respects_converges:
60     ∀b,c. eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (rp\sub\f⎻ b ↓ rp\sub\f⎻ c));
61    respects_all_covered:
62      eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (oa_one (form CS2))))
63            (Ext⎽CS1 (oa_one (form CS1)))
64  }.
65
66 definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid2.
67  intros;
68  constructor 1;
69   [ apply (convergent_relation_pair c c1)
70   | constructor 1;
71      [ intros;
72        apply (relation_pair_equality c c1 c2 c3);
73      | intros 1; apply refl2;
74      | intros 2; apply sym2; 
75      | intros 3; apply trans2]]
76 qed.
77
78 definition convergent_relation_space_of_convergent_relation_space_setoid: 
79   ∀CS1,CS2.carr2 (convergent_relation_space_setoid CS1 CS2) → 
80     convergent_relation_pair CS1 CS2  ≝ λP,Q,c.c.
81 coercion convergent_relation_space_of_convergent_relation_space_setoid.
82
83 definition convergent_relation_space_composition:
84  ∀o1,o2,o3: concrete_space.
85   binary_morphism2
86    (convergent_relation_space_setoid o1 o2)
87    (convergent_relation_space_setoid o2 o3)
88    (convergent_relation_space_setoid o1 o3).
89  intros; constructor 1;
90      [ intros; whd in t t1 ⊢ %;
91        constructor 1;
92         [ apply (c1 ∘ c);
93         | intros;
94           change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
95           alias symbol "trans" = "trans1".
96           apply (.= († (respects_converges : ?)));
97           apply (respects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2));
98         | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (form o3)))));
99           apply (.= (†(respects_all_covered :?)));
100           apply rule (respects_all_covered ?? c);]
101      | intros;
102        change with (b ∘ a = b' ∘ a'); 
103        change in e with (rp ?? a = rp ?? a');
104        change in e1 with (rp ?? b = rp ?? b');
105        apply (e‡e1);]
106 qed.
107
108 definition CSPA: category2.
109  constructor 1;
110   [ apply concrete_space
111   | apply convergent_relation_space_setoid
112   | intro; constructor 1;
113      [ apply id2
114      | intros; apply refl1;
115      | apply refl1]
116   | apply convergent_relation_space_composition
117   | intros; simplify; whd in a12 a23 a34;
118     change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
119     apply rule ASSOC;
120   | intros; simplify;
121     change with (a ∘ id2 BP o1 = a);
122     apply (id_neutral_right2 : ?);
123   | intros; simplify;
124     change with (id2 ? o2 ∘ a = a);
125     apply (id_neutral_left2 : ?);]
126 qed.
127
128 definition concrete_space_of_CSPA : objs2 CSPA → concrete_space ≝ λx.x.
129 coercion concrete_space_of_CSPA.
130
131 definition convergent_relation_space_setoid_of_arrows2_CSPA :
132  ∀P,Q. arrows2 CSPA P Q → convergent_relation_space_setoid P Q ≝ λP,Q,x.x.
133 coercion convergent_relation_space_setoid_of_arrows2_CSPA.
134