]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/datatypes/subsets.ma
5483dfa37f1fd648a38eca5727cd9cb9d05cc134
[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 definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
78  intros;
79  constructor 1;
80   [ apply (λA.λU,V:Ω \sup A.exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V))
81   | intros;
82     constructor 1; intro; cases H2; exists; [1,4: apply w]
83      [ apply (. #‡H); assumption
84      | apply (. #‡H1); assumption
85      | apply (. #‡(H \sup -1)); assumption;
86      | apply (. #‡(H1 \sup -1)); assumption]]
87 qed.
88
89 interpretation "overlaps" 'overlaps U V = (fun1 ___ (overlaps _) U V).
90
91 definition intersects:
92  ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
93  intros;
94  constructor 1;
95   [ apply (λU,V. {x | x ∈ U ∧ x ∈ V });
96     intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1;
97   | intros;
98     split; intros 2; simplify in f ⊢ %;
99     [ apply (. (#‡H)‡(#‡H1)); assumption
100     | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
101 qed.
102
103 interpretation "intersects" 'intersects U V = (fun1 ___ (intersects _) U V).
104
105 definition union:
106  ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
107  intros;
108  constructor 1;
109   [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
110     intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1
111   | intros;
112     split; intros 2; simplify in f ⊢ %;
113     [ apply (. (#‡H)‡(#‡H1)); assumption
114     | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
115 qed.
116
117 interpretation "union" 'union U V = (fun1 ___ (union _) U V).
118
119 definition singleton: ∀A:setoid. A → Ω \sup A.
120  apply (λA:setoid.λa:A.{b | a=b});
121  intros; simplify;
122  split; intro;
123  apply (.= H1);
124   [ apply H | apply (H \sup -1) ]
125 qed.
126
127 interpretation "singleton" 'singl a = (singleton _ a).