]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/sets/sets.ma
3f5a4feeeceb634c199b7b7c8727c3289f729025
[helm.git] / helm / software / matita / nlibrary / sets / sets.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 "logic/equality.ma".
16
17 nrecord powerset (A: Type) : Type[1] ≝ { mem: A → CProp }.
18
19 interpretation "powerset" 'powerset A = (powerset A).
20
21 interpretation "subset construction" 'subset \eta.x = (mk_powerset ? x).
22
23 interpretation "mem" 'mem a S = (mem ? S a).
24
25 ndefinition subseteq ≝ λA.λU,V.∀a:A. a ∈ U → a ∈ V.
26
27 interpretation "subseteq" 'subseteq U V = (subseteq ? U V).
28
29 ntheorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
30  #A; #S; #x; #H; nassumption;
31 nqed.
32
33 ntheorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
34  #A; #S1; #S2; #S3; #H12; #H23; #x; #H;
35  napply (H23 …); napply (H12 …); nassumption;
36 nqed.
37
38 ndefinition overlaps ≝ λA.λU,V:Ω \sup A.∃x:A.x ∈ U ∧ x ∈ V.
39
40 interpretation "overlaps" 'overlaps U V = (overlaps ? U V).
41
42 ndefinition intersects ≝ λA.λU,V:Ω \sup A. {x | x ∈ U ∧ x ∈ V }.
43
44 interpretation "intersects" 'intersects U V = (intersects ? U V).
45
46 ndefinition union ≝ λA.λU,V:Ω \sup A. {x | x ∈ U ∨ x ∈ V }.
47
48 interpretation "union" 'union U V = (union ? U V).
49
50 ndefinition singleton ≝ λA.λa:A.{b | a=b}.
51
52 interpretation "singleton" 'singl a = (singleton ? a).