]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/sets/sets.ma
81f0f07c7a20c2e306a74e68e9612d3b2bde8973
[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/cprop.ma".
16
17 nrecord powerset (A: setoid) : Type[1] ≝ { mem_op: unary_morphism1 A CPROP }.
18
19 interpretation "powerset" 'powerset A = (powerset A).
20
21 interpretation "subset construction" 'subset \eta.x =
22  (mk_powerset ? (mk_unary_morphism1 ? CPROP x ?)).
23
24 interpretation "mem" 'mem a S = (mem_op ? S a).
25
26 ndefinition subseteq ≝ λA:setoid.λU,V.∀a:A. a ∈ U → a ∈ V.
27
28 interpretation "subseteq" 'subseteq U V = (subseteq ? U V).
29
30 ntheorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
31  #A; #S; #x; #H; nassumption;
32 nqed.
33
34 ntheorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
35  #A; #S1; #S2; #S3; #H12; #H23; #x; #H;
36  napply H23; napply H12; nassumption;
37 nqed.
38
39 ndefinition powerset_setoid1: setoid → setoid1.
40  #S; napply mk_setoid1
41   [ napply (Ω \sup S)
42   | napply mk_equivalence_relation1
43      [ #A; #B; napply (∀x. iff (x ∈ A) (x ∈ B))
44      | nwhd; #x; #x0; napply mk_iff; #H; nassumption
45      | nwhd; #x; #y; #H; #A; napply mk_iff; #K
46         [ napply (fi ?? (H ?)) | napply (if ?? (H ?)) ]
47         nassumption
48      | nwhd; #A; #B; #C; #H1; #H2; #H3; napply mk_iff; #H4
49         [ napply (if ?? (H2 ?)); napply (if ?? (H1 ?)); nassumption
50         | napply (fi ?? (H1 ?)); napply (fi ?? (H2 ?)); nassumption]##]
51 nqed.
52
53 unification hint 0 (∀A.(λx,y.True) (Ω \sup A) (carr1 (powerset_setoid1 A))).
54
55 ndefinition mem: ∀A:setoid. binary_morphism1 A (powerset_setoid1 A) CPROP.
56  #A; napply mk_binary_morphism1
57   [ napply (λa.λA.a ∈ A)
58   | #a; #a'; #B; #B'; #Ha; #HB; napply mk_iff; #H
59      [ napply (. (†Ha^-1)); (* CSC: notation for ∈ not working *)
60        napply (if ?? (HB ?)); nassumption
61      | napply (. (†Ha)); napply (fi ?? (HB ?)); nassumption]##]
62 nqed.
63
64 unification hint 0 (∀A,x,S. (λx,y.True) (mem_op A x S) (fun21 ??? (mem A) S x)).
65
66 ndefinition overlaps ≝ λA.λU,V:Ω \sup A.∃x:A.x ∈ U ∧ x ∈ V.
67
68 interpretation "overlaps" 'overlaps U V = (overlaps ? U V).
69
70 ndefinition intersects ≝ λA.λU,V:Ω \sup A. {x | x ∈ U ∧ x ∈ V }.
71  #a; #a'; #H; napply mk_iff; *; #H1; #H2
72   [ napply (. ((H^-1‡#)‡(H^-1‡#))); nwhd; napply conj; nassumption
73   | napply (. ((H‡#)‡(H‡#))); nwhd; napply conj; nassumption]
74 nqed.
75   
76 (*interpretation "intersects" 'intersects U V = (intersects ? U V).*)
77
78 (*
79 ndefinition union ≝ λA.λU,V:Ω \sup A. {x | x ∈ U ∨ x ∈ V }.
80
81 interpretation "union" 'union U V = (union ? U V).
82
83 ndefinition singleton ≝ λA:setoid.λa:A.{b | a=b}.
84
85 interpretation "singleton" 'singl a = (singleton ? a).*)