]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/subsets.ma
Additional contribs.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / 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 "categories.ma".
16
17 record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }.
18
19 definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp0 ≝
20  λA:objs1 SET.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a.
21
22 theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
23  intros 6; intros 2;
24  apply s1; apply s;
25  assumption.
26 qed.
27
28 definition powerset_setoid1: SET → SET1.
29  intros (T);
30  constructor 1;
31   [ apply (powerset_carrier T)
32   | constructor 1;
33      [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U)
34      | simplify; intros; split; intros 2; assumption
35      | simplify; intros (x y H); cases H; split; assumption
36      | simplify; intros (x y z H H1); cases H; cases H1; split;
37        apply transitive_subseteq_operator; [1,4: apply y ]
38        assumption ]]
39 qed.
40
41 interpretation "powerset" 'powerset A = (powerset_setoid1 A).
42
43 interpretation "subset construction" 'subset \eta.x =
44  (mk_powerset_carrier ? (mk_unary_morphism1 ? CPROP x ?)).
45
46 definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP.
47  intros;
48  constructor 1;
49   [ apply (λx,S. mem_operator ? S x)
50   | intros 5;
51     cases b; clear b; cases b'; clear b'; simplify; intros;
52     apply (trans1 ????? (prop11 ?? u ?? e));
53     cases e1; whd in s s1;
54     split; intro;
55      [ apply s; assumption
56      | apply s1; assumption]]
57 qed.
58
59 interpretation "mem" 'mem a S = (fun21 ??? (mem ?) a S).
60
61 definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
62  intros;
63  constructor 1;
64   [ apply (λU,V. subseteq_operator ? U V)
65   | intros;
66     cases e; cases e1;
67     split; intros 1;
68     [ apply (transitive_subseteq_operator ????? s2);
69       apply (transitive_subseteq_operator ???? s1 s4)
70     | apply (transitive_subseteq_operator ????? s3);
71       apply (transitive_subseteq_operator ???? s s4) ]]
72 qed.
73
74 interpretation "subseteq" 'subseteq U V = (fun21 ??? (subseteq ?) U V).
75
76
77
78 theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
79  intros 4; assumption.
80 qed.
81
82 theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
83  intros; apply transitive_subseteq_operator; [apply S2] assumption.
84 qed.
85
86 definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
87  intros;
88  constructor 1;
89   (* Se metto x al posto di ? ottengo una universe inconsistency *)
90   [ apply (λA:objs1 SET.λU,V:Ω \sup A.(exT2 ? (λx:A.?(*x*) ∈ U) (λx:A.?(*x*) ∈ V) : CProp0))
91   | intros;
92     constructor 1; intro; cases e2; exists; [1,4: apply w]
93      [ apply (. #‡e^-1); assumption
94      | apply (. #‡e1^-1); assumption
95      | apply (. #‡e); assumption;
96      | apply (. #‡e1); assumption]]
97 qed.
98
99 interpretation "overlaps" 'overlaps U V = (fun21 ??? (overlaps ?) U V).
100
101 definition intersects:
102  ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
103  intros;
104  constructor 1;
105   [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
106     intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1;
107   | intros;
108     split; intros 2; simplify in f ⊢ %;
109     [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
110     | apply (. (#‡e)‡(#‡e1)); assumption]]
111 qed.
112
113 interpretation "intersects" 'intersects U V = (fun21 ??? (intersects ?) U V).
114
115 definition union:
116  ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup 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. unary_morphism1 A (Ω \sup 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 \sup -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
144 definition big_intersects:
145  ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup 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:
157  ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)).
158  intros; constructor 1;
159   [ intro; whd; whd in A; whd in I;
160     apply ({x | ∃i:I. x ∈ c i });
161     simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w]
162     [ apply (. (e^-1‡#)); | apply (. (e‡#)); ]
163     apply x;
164   | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w]
165      [ apply (. (#‡(e w)^-1)); apply x;
166      | apply (. (#‡e w)); apply x]]
167 qed.