]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/sets.ma
naive sets (A-> Prop)
[helm.git] / matita / matita / lib / basics / sets.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic        
3     ||A||  Library of Mathematics, developed at the Computer Science     
4     ||T||  Department of the University of Bologna, Italy.                     
5     ||I||                                                                 
6     ||T||  
7     ||A||  This file is distributed under the terms of the 
8     \   /  GNU General Public License Version 2        
9      \ /      
10       V_______________________________________________________________ *)
11
12 include "basics/logic.ma".
13
14 (**** a subset of A is just an object of type A→Prop ****)
15
16 definition empty_set ≝ λA:Type[0].λa:A.False.
17 notation "\emptyv" non associative with precedence 90 for @{'empty_set}.
18 interpretation "empty set" 'empty_set = (empty_set ?).
19
20 definition singleton ≝ λA.λx,a:A.x=a.
21 (* notation "{x}" non associative with precedence 90 for @{'sing_lang $x}. *)
22 interpretation "singleton" 'singl x = (singleton ? x).
23
24 definition union : ∀A:Type[0].∀P,Q.A → Prop ≝ λA,P,Q,a.P a ∨ Q a.
25 interpretation "union" 'union a b = (union ? a b).
26
27 definition intersection : ∀A:Type[0].∀P,Q.A→Prop ≝ λA,P,Q,a.P a ∧ Q a.
28 interpretation "intersection" 'intersects a b = (intersection ? a b).
29
30 definition subset: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ λA,P,Q.∀a:A.(P a → Q a).
31 interpretation "subset" 'subseteq a b = (intersection ? a b).
32
33 (* extensional equality *)
34 definition eqP ≝ λA:Type[0].λP,Q:A → Prop.∀a:A.P a ↔ Q a.
35 notation "A =1 B" non associative with precedence 45 for @{'eqP $A $B}.
36 interpretation "extensional equality" 'eqP a b = (eqP ? a b).
37
38 lemma eqP_sym: ∀U.∀A,B:U →Prop. 
39   A =1 B → B =1 A.
40 #U #A #B #eqAB #a @iff_sym @eqAB qed.
41  
42 lemma eqP_trans: ∀U.∀A,B,C:U →Prop. 
43   A =1 B → B =1 C → A =1 C.
44 #U #A #B #C #eqAB #eqBC #a @iff_trans // qed.
45
46 lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. 
47   A =1 C  → A ∪ B =1 C ∪ B.
48 #U #A #B #C #eqAB #a @iff_or_r @eqAB qed.
49   
50 lemma eqP_union_l: ∀U.∀A,B,C:U →Prop. 
51   B =1 C  → A ∪ B =1 A ∪ C.
52 #U #A #B #C #eqBC #a @iff_or_l @eqBC qed.
53   
54 (* set equalities *)
55 lemma union_comm : ∀U.∀A,B:U →Prop. 
56   A ∪ B =1 B ∪ A.
57 #U #A #B #a % * /2/ qed. 
58
59 lemma union_assoc: ∀U.∀A,B,C:U → Prop. 
60   A ∪ B ∪ C =1 A ∪ (B ∪ C).
61 #S #A #B #C #w % [* [* /3/ | /3/] | * [/3/ | * /3/]
62 qed.   
63
64 lemma cap_comm : ∀U.∀A,B:U →Prop. 
65   A ∩ B =1 B ∩ A.
66 #U #A #B #a % * /2/ qed. 
67
68 lemma union_idemp: ∀U.∀A:U →Prop. 
69   A ∪ A =1 A.
70 #U #A #a % [* // | /2/] qed. 
71
72 lemma cap_idemp: ∀U.∀A:U →Prop. 
73   A ∩ A =1 A.
74 #U #A #a % [* // | /2/] qed. 
75
76