]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/formal_topology/subsets.ma
decentralizing core notation
[helm.git] / matita / matita / lib / formal_topology / subsets.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/categories.ma".
16 include "basics/core_notation/singl_1.ma".
17 include "basics/core_notation/subset_1.ma".
18 (*
19 record powerset_carrier (A: objs1 SET) : Type[1] ≝ { mem_operator: A ⇒_1 CPROP }.
20 interpretation "powerset low" 'powerset A = (powerset_carrier A).
21 notation "hvbox(a break ∈. b)" non associative with precedence 45 for @{ 'mem_low $a $b }.
22 interpretation "memlow" 'mem_low a S = (fun11 ?? (mem_operator ? S) a).
23
24 definition subseteq_operator: ∀A: objs1 SET. Ω^A → Ω^A → CProp[0] ≝
25  λA:objs1 SET.λU,V.∀a:A. a ∈. U → a ∈. V.
26
27 theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
28  intros 6; intros 2;
29  apply s1; apply s;
30  assumption.
31 qed.
32
33 definition powerset_setoid1: SET → SET1.
34  intros (T);
35  constructor 1;
36   [ apply (powerset_carrier T)
37   | constructor 1;
38      [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U)
39      | simplify; intros; split; intros 2; assumption
40      | simplify; intros (x y H); cases H; split; assumption
41      | simplify; intros (x y z H H1); cases H; cases H1; split;
42        apply transitive_subseteq_operator; [1,4: apply y ]
43        assumption ]]
44 qed.
45
46 interpretation "powerset" 'powerset A = (powerset_setoid1 A).
47
48 interpretation "subset construction" 'subset \eta.x =
49  (mk_powerset_carrier ? (mk_unary_morphism1 ? CPROP x ?)).
50
51 definition mem: ∀A. A × Ω^A ⇒_1 CPROP.
52  intros;
53  constructor 1;
54   [ apply (λx,S. mem_operator ? S x)
55   | intros 5;
56     cases b; clear b; cases b'; clear b'; simplify; intros;
57     apply (trans1 ????? (prop11 ?? u ?? e));
58     cases e1; whd in s s1;
59     split; intro;
60      [ apply s; assumption
61      | apply s1; assumption]]
62 qed.
63
64 interpretation "mem" 'mem a S = (fun21 ??? (mem ?) a S).
65
66 definition subseteq: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
67  intros;
68  constructor 1;
69   [ apply (λU,V. subseteq_operator ? U V)
70   | intros;
71     cases e; cases e1;
72     split; intros 1;
73     [ apply (transitive_subseteq_operator ????? s2);
74       apply (transitive_subseteq_operator ???? s1 s4)
75     | apply (transitive_subseteq_operator ????? s3);
76       apply (transitive_subseteq_operator ???? s s4) ]]
77 qed.
78
79 interpretation "subseteq" 'subseteq U V = (fun21 ??? (subseteq ?) U V).
80
81 theorem subseteq_refl: ∀A.∀S:Ω^A.S ⊆ S.
82  intros 4; assumption.
83 qed.
84
85 theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω^A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
86  intros; apply transitive_subseteq_operator; [apply S2] assumption.
87 qed.
88
89 definition overlaps: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
90  intros;
91  constructor 1;
92   [ apply (λA:objs1 SET.λU,V:Ω^A.(exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V) : CProp[0]))
93   | intros;
94     constructor 1; intro; cases e2; exists; [1,4: apply w]
95      [ apply (. #‡e^-1); assumption
96      | apply (. #‡e1^-1); assumption
97      | apply (. #‡e); assumption;
98      | apply (. #‡e1); assumption]]
99 qed.
100
101 interpretation "overlaps" 'overlaps U V = (fun21 ??? (overlaps ?) U V).
102
103 definition intersects: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
104  intros;
105  constructor 1;
106   [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
107     intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1;
108   | intros;
109     split; intros 2; simplify in f ⊢ %;
110     [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
111     | apply (. (#‡e)‡(#‡e1)); assumption]]
112 qed.
113
114 interpretation "intersects" 'intersects U V = (fun21 ??? (intersects ?) U V).
115
116 definition union: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
117  intros;
118  constructor 1;
119   [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
120     intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1
121   | intros;
122     split; intros 2; simplify in f ⊢ %;
123     [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
124     | apply (. (#‡e)‡(#‡e1)); assumption]]
125 qed.
126
127 interpretation "union" 'union U V = (fun21 ??? (union ?) U V).
128
129 (* qua non riesco a mettere set *)
130 definition singleton: ∀A:setoid. A ⇒_1 Ω^A.
131  intros; constructor 1;
132   [ apply (λa:A.{b | a =_0 b}); unfold setoid1_of_setoid; simplify;
133     intros; simplify;
134     split; intro;
135     apply (.= e1);
136      [ apply e | apply (e^-1) ]
137   | unfold setoid1_of_setoid; simplify;
138     intros; split; intros 2; simplify in f ⊢ %; apply trans;
139      [ apply a |4: apply a'] try assumption; apply sym; assumption]
140 qed.
141
142 interpretation "singleton" 'singl a = (fun11 ?? (singleton ?) a).
143 notation > "{ term 19 a : S }" with precedence 90 for @{fun11 ?? (singleton $S) $a}.
144
145 definition big_intersects: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^A.
146  intros; constructor 1;
147   [ intro; whd; whd in I;
148     apply ({x | ∀i:I. x ∈ c i});
149     simplify; intros; split; intros; [ apply (. (e^-1‡#)); | apply (. e‡#); ]
150     apply f;
151   | intros; split; intros 2; simplify in f ⊢ %; intro;
152      [ apply (. (#‡(e i)^-1)); apply f;
153      | apply (. (#‡e i)); apply f]]
154 qed.
155
156 definition big_union: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^A.
157  intros; constructor 1;
158   [ intro; whd; whd in A; whd in I;
159     apply ({x | ∃i:I. x ∈ c i });
160     simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w]
161     [ apply (. (e^-1‡#)); | apply (. (e‡#)); ]
162     apply x;
163   | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w]
164      [ apply (. (#‡(e w)^-1)); apply x;
165      | apply (. (#‡e w)); apply x]]
166 qed.
167
168 notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (\emsp) term 90 p)" 
169 non associative with precedence 55 for @{ 'bigcup $p }.
170 notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (ident i ∈  I) break term 90 p)" 
171 non associative with precedence 55 for @{ 'bigcup_mk (λ${ident i}:$I.$p) }.
172 notation > "hovbox(⋃ f)" non associative with precedence 65 for @{ 'bigcup $f }.
173
174 notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (\emsp) term 90 p)" 
175 non associative with precedence 55 for @{ 'bigcap $p }.
176 notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (ident i ∈  I) break term 90 p)" 
177 non associative with precedence 55 for @{ 'bigcap_mk (λ${ident i}:$I.$p) }.
178 notation > "hovbox(⋂ f)" non associative with precedence 65 for @{ 'bigcap $f }.
179
180 interpretation "bigcup" 'bigcup f = (fun12 ?? (big_union ??) f).
181 interpretation "bigcap" 'bigcap f = (fun12 ?? (big_intersects ??) f).
182 interpretation "bigcup mk" 'bigcup_mk f = (fun12 ?? (big_union ??) (mk_unary_morphism2 ?? f ?)).
183 interpretation "bigcap mk" 'bigcap_mk f = (fun12 ?? (big_intersects ??) (mk_unary_morphism2 ?? f ?)).
184 *)