]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/datatypes/subsets.ma
7c9a13195f2ba6f07635f3067c5e4e942a07d095
[helm.git] / helm / software / matita / library / datatypes / 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 "logic/cprop_connectives.ma".
16 include "datatypes/categories.ma".
17
18 record powerset_carrier (A: setoid) : Type ≝ { mem_operator: A ⇒ CPROP }.
19
20 definition subseteq_operator ≝
21  λA:setoid.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a.
22
23 theorem transitive_subseteq_operator: ∀A. transitive ? (subseteq_operator A).
24  intros 6; intros 2;
25  apply s1; apply s;
26  assumption.
27 qed.
28
29 definition powerset_setoid: setoid → setoid1.
30  intros (T);
31  constructor 1;
32   [ apply (powerset_carrier T)
33   | constructor 1;
34      [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U)
35      | simplify; intros; split; intros 2; assumption
36      | simplify; intros (x y H); cases H; split; assumption
37      | simplify; intros (x y z H H1); cases H; cases H1; split;
38        apply transitive_subseteq_operator; [1,4: apply y ]
39        assumption ]]
40 qed.
41
42 interpretation "powerset" 'powerset A = (powerset_setoid A).
43
44 interpretation "subset construction" 'subset \eta.x =
45  (mk_powerset_carrier _ (mk_unary_morphism _ CPROP x _)).
46
47 definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP.
48  intros;
49  constructor 1;
50   [ apply (λx,S. mem_operator ? S x)
51   | intros 5;
52     cases b; clear b; cases b'; clear b'; simplify; intros;
53     apply (trans1 ????? (prop_1 ?? u ?? H));
54     cases H1; whd in s s1;
55     split; intro;
56      [ apply s; assumption
57      | apply s1; assumption]]
58 qed.     
59
60 interpretation "mem" 'mem a S = (fun1 ___ (mem _) a S).
61
62 definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
63  intros;
64  constructor 1;
65   [ apply (λU,V. subseteq_operator ? U V)
66   | intros;
67     cases H; cases H1;
68     split; intros 1;
69     [ apply (transitive_subseteq_operator ????? s2);
70       apply (transitive_subseteq_operator ???? s1 s4)
71     | apply (transitive_subseteq_operator ????? s3);
72       apply (transitive_subseteq_operator ???? s s4) ]]
73 qed.
74
75 interpretation "subseteq" 'subseteq U V = (fun1 ___ (subseteq _) U V).
76
77 theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
78  intros 4; assumption.
79 qed.
80
81 theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
82  intros; apply transitive_subseteq_operator; [apply S2] assumption.
83 qed.
84
85 definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
86  intros;
87  constructor 1;
88   [ apply (λA.λU,V:Ω \sup A.exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V))
89   | intros;
90     constructor 1; intro; cases H2; exists; [1,4: apply w]
91      [ apply (. #‡H); assumption
92      | apply (. #‡H1); assumption
93      | apply (. #‡(H \sup -1)); assumption;
94      | apply (. #‡(H1 \sup -1)); assumption]]
95 qed.
96
97 interpretation "overlaps" 'overlaps U V = (fun1 ___ (overlaps _) U V).
98
99 definition intersects:
100  ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
101  intros;
102  constructor 1;
103   [ apply (λU,V. {x | x ∈ U ∧ x ∈ V });
104     intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1;
105   | intros;
106     split; intros 2; simplify in f ⊢ %;
107     [ apply (. (#‡H)‡(#‡H1)); assumption
108     | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
109 qed.
110
111 interpretation "intersects" 'intersects U V = (fun1 ___ (intersects _) U V).
112
113 definition union:
114  ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
115  intros;
116  constructor 1;
117   [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
118     intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1
119   | intros;
120     split; intros 2; simplify in f ⊢ %;
121     [ apply (. (#‡H)‡(#‡H1)); assumption
122     | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
123 qed.
124
125 interpretation "union" 'union U V = (fun1 ___ (union _) U V).
126
127 definition singleton: ∀A:setoid. unary_morphism A (Ω \sup A).
128  intros; constructor 1;
129   [ apply (λA:setoid.λa:A.{b | a=b});
130     intros; simplify;
131     split; intro;
132     apply (.= H1);
133      [ apply H | apply (H \sup -1) ]
134   | intros; split; intros 2; simplify in f ⊢ %; apply trans;
135      [ apply a |4: apply a'] try assumption; apply sym; assumption]
136 qed.
137
138 interpretation "singleton" 'singl a = (fun_1 __ (singleton _) a).