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 "categories.ma".
17 record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }.
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.
22 theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
28 definition powerset_setoid1: SET → SET1.
31 [ apply (powerset_carrier T)
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 ]
41 interpretation "powerset" 'powerset A = (powerset_setoid1 A).
43 interpretation "subset construction" 'subset \eta.x =
44 (mk_powerset_carrier _ (mk_unary_morphism1 _ CPROP x _)).
46 definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP.
49 [ apply (λx,S. mem_operator ? S x)
51 cases b; clear b; cases b'; clear b'; simplify; intros;
52 apply (trans1 ????? (prop11 ?? u ?? e));
53 cases e1; whd in s s1;
56 | apply s1; assumption]]
59 interpretation "mem" 'mem a S = (fun21 ___ (mem _) a S).
61 definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
64 [ apply (λU,V. subseteq_operator ? U V)
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) ]]
74 interpretation "subseteq" 'subseteq U V = (fun21 ___ (subseteq _) U V).
78 theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
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.
86 definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
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))
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]]
99 interpretation "overlaps" 'overlaps U V = (fun21 ___ (overlaps _) U V).
101 definition intersects:
102 ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
105 [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
106 intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1;
108 split; intros 2; simplify in f ⊢ %;
109 [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
110 | apply (. (#‡e)‡(#‡e1)); assumption]]
113 interpretation "intersects" 'intersects U V = (fun21 ___ (intersects _) U V).
116 ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
119 [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
120 intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1
122 split; intros 2; simplify in f ⊢ %;
123 [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
124 | apply (. (#‡e)‡(#‡e1)); assumption]]
127 interpretation "union" 'union U V = (fun21 ___ (union _) U V).
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 | eq ? a b}); unfold setoid1_of_setoid; simplify;
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]
142 interpretation "singleton" 'singl a = (fun11 __ (singleton _) a).
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‡#); ]
151 | intros; split; intros 2; simplify in f ⊢ %; intro;
152 [ apply (. (#‡(e i)^-1)); apply f;
153 | apply (. (#‡e i)); apply f]]
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‡#)); ]
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]]