]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/datatypes/subsets.ma
Definition of formal_topologies.
[helm.git] / helm / software / matita / library / datatypes / subsets.ma
index 96ec347c75f3e8fba8076ea589a2d6a8e7cc4d52..5483dfa37f1fd648a38eca5727cd9cb9d05cc134 100644 (file)
 (**************************************************************************)
 
 include "logic/cprop_connectives.ma".
+include "datatypes/categories.ma".
 
-record powerset (A: Type) : Type ≝ { char: A → CProp }.
+record powerset_carrier (A: setoid) : Type ≝ { mem_operator: A ⇒ CPROP }.
 
-interpretation "powerset" 'powerset A = (powerset A).
+definition subseteq_operator ≝
+ λA:setoid.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a.
 
-interpretation "subset construction" 'subset \eta.x = (mk_powerset _ x).
-
-definition mem ≝ λA.λS:Ω \sup A.λx:A. match S with [mk_powerset c ⇒ c x].
-
-interpretation "mem" 'mem a S = (mem _ S a).
-
-definition overlaps ≝ λA:Type.λU,V:Ω \sup A.exT2 ? (λa:A. a ∈ U) (λa.a ∈ V).
-
-interpretation "overlaps" 'overlaps U V = (overlaps _ U V).
-
-definition subseteq ≝ λA:Type.λU,V:Ω \sup A.∀a:A. a ∈ U → a ∈ V.
-
-interpretation "subseteq" 'subseteq U V = (subseteq _ U V).
-
-definition intersects ≝ λA:Type.λU,V:Ω \sup A.{a | a ∈ U ∧ a ∈ V}.
-
-interpretation "intersects" 'intersects U V = (intersects _ U V).
-
-definition union ≝ λA:Type.λU,V:Ω \sup A.{a | a ∈ U ∨ a ∈ V}.
-
-interpretation "union" 'union U V = (union _ U V).
-
-record ssigma (A:Type) (S: powerset A) : Type ≝
- { witness:> A;
-   proof:> witness ∈ S
- }.
+theorem transitive_subseteq_operator: ∀A. transitive ? (subseteq_operator A).
+ intros 6; intros 2;
+ apply s1; apply s;
+ assumption.
+qed.
 
-coercion ssigma.
+definition powerset_setoid: setoid → setoid1.
+ intros (T);
+ constructor 1;
+  [ apply (powerset_carrier T)
+  | constructor 1;
+     [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U)
+     | simplify; intros; split; intros 2; assumption
+     | simplify; intros (x y H); cases H; split; assumption
+     | simplify; intros (x y z H H1); cases H; cases H1; split;
+       apply transitive_subseteq_operator; [1,4: apply y ]
+       assumption ]]
+qed.
 
-record binary_relation (A,B: Type) (U: Ω \sup A) (V: Ω \sup B) : Type ≝
- { satisfy:2> U → V → CProp }.
+interpretation "powerset" 'powerset A = (powerset_setoid A).
 
-(*notation < "hvbox (x  (\circ term 19 r \frac \nbsp \circ) y)"  with precedence 45 for @{'satisfy $r $x $y}.*)
-notation < "hvbox (x \nbsp \natur term 90 r \nbsp y)"  with precedence 45 for @{'satisfy $r $x $y}.
-notation > "hvbox (x \natur term 90 r y)"  with precedence 45 for @{'satisfy $r $x $y}.
-interpretation "relation applied" 'satisfy r x y = (satisfy ____ r x y).
+interpretation "subset construction" 'subset \eta.x =
+ (mk_powerset_carrier _ (mk_unary_morphism _ CPROP x _)).
 
-definition composition:
- ∀A,B,C.∀U1: Ω \sup A.∀U2: Ω \sup B.∀U3: Ω \sup C.
-  binary_relation ?? U1 U2 → binary_relation ?? U2 U3 →
-   binary_relation ?? U1 U3.
- intros (A B C U1 U2 U3 R12 R23);
+definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP.
+ intros;
+ constructor 1;
+  [ apply (λx,S. mem_operator ? S x)
+  | intros 5;
+    cases b; clear b; cases b'; clear b'; simplify; intros;
+    apply (trans1 ????? (prop_1 ?? u ?? H));
+    cases H1; whd in s s1;
+    split; intro;
+     [ apply s; assumption
+     | apply s1; assumption]]
+qed.     
+
+interpretation "mem" 'mem a S = (fun1 ___ (mem _) a S).
+
+definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
+ intros;
  constructor 1;
- intros (s1 s3);
- apply (∃s2. s1 ♮R12 s2 ∧ s2 ♮R23 s3);
+  [ apply (λU,V. subseteq_operator ? U V)
+  | intros;
+    cases H; cases H1;
+    split; intros 1;
+    [ apply (transitive_subseteq_operator ????? s2);
+      apply (transitive_subseteq_operator ???? s1 s4)
+    | apply (transitive_subseteq_operator ????? s3);
+      apply (transitive_subseteq_operator ???? s s4) ]]
 qed.
 
-interpretation "binary relation composition" 'compose x y = (composition ______ x y).
+interpretation "subseteq" 'subseteq U V = (fun1 ___ (subseteq _) U V).
 
-definition equal_relations ≝
- λA,B,U,V.λr,r': binary_relation A B U V.
-  ∀x,y. r x y ↔ r' x y.
+definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
+ intros;
+ constructor 1;
+  [ apply (λA.λU,V:Ω \sup A.exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V))
+  | intros;
+    constructor 1; intro; cases H2; exists; [1,4: apply w]
+     [ apply (. #‡H); assumption
+     | apply (. #‡H1); assumption
+     | apply (. #‡(H \sup -1)); assumption;
+     | apply (. #‡(H1 \sup -1)); assumption]]
+qed.
 
-interpretation "equal relation" 'eq x y = (equal_relations ____ x y).
+interpretation "overlaps" 'overlaps U V = (fun1 ___ (overlaps _) U V).
 
-lemma refl_equal_relations: ∀A,B,U,V. reflexive ? (equal_relations A B U V).
- intros 5; intros 2; split; intro; assumption.
+definition intersects:
+ ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
+ intros;
+ constructor 1;
+  [ apply (λU,V. {x | x ∈ U ∧ x ∈ V });
+    intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1;
+  | intros;
+    split; intros 2; simplify in f ⊢ %;
+    [ apply (. (#‡H)‡(#‡H1)); assumption
+    | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
 qed.
 
-lemma sym_equal_relations: ∀A,B,U,V. symmetric ? (equal_relations A B U V).
- intros 7; intros 2; split; intro;
-  [ apply (fi ?? (H ??)) | apply (if ?? (H ??))] assumption.
-qed.
+interpretation "intersects" 'intersects U V = (fun1 ___ (intersects _) U V).
 
-lemma trans_equal_relations: ∀A,B,U,V. transitive ? (equal_relations A B U V).
- intros 9; intros 2; split; intro;
-  [ apply (if ?? (H1 ??)) | apply (fi ?? (H ??)) ]
-  [ apply (if ?? (H ??)) | apply (fi ?? (H1 ??)) ]
-  assumption.
+definition union:
+ ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
+ intros;
+ constructor 1;
+  [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
+    intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1
+  | intros;
+    split; intros 2; simplify in f ⊢ %;
+    [ apply (. (#‡H)‡(#‡H1)); assumption
+    | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
 qed.
 
-lemma equal_morphism:
- ∀A,B,U,V.∀r1,r1',r2,r2':binary_relation A B U V.
-  r1' = r1 → r2 = r2' → r1 = r2 → r1' = r2'.
- intros 13;
- split; intro;
-  [ apply (if ?? (H1 ??));
-    apply (if ?? (H2 ??));
-    apply (if ?? (H ??));
-    assumption
-  | apply (fi ?? (H ??));
-    apply (fi ?? (H2 ??));
-    apply (fi ?? (H1 ??));
-    assumption
-  ]
-qed.
+interpretation "union" 'union U V = (fun1 ___ (union _) U V).
 
-lemma associative_composition:
- ∀A,B,C,D.∀U1,U2,U3,U4.
-  ∀r1:binary_relation A B U1 U2.
-   ∀r2:binary_relation B C U2 U3.
-    ∀r3:binary_relation C D U3 U4.
-     (r1 ∘ r2) ∘ r3 = r1 ∘ (r2 ∘ r3).
- intros 13;
+definition singleton: ∀A:setoid. A → Ω \sup A.
+ apply (λA:setoid.λa:A.{b | a=b});
+ intros; simplify;
  split; intro;
- cases H; clear H; cases H1; clear H1;
- [cases H; clear H | cases H2; clear H2]
- cases H1; clear H1;
- exists; try assumption;
- split; try assumption;
- exists; try assumption;
- split; assumption.
-qed.
-
-lemma composition_morphism:
- ∀A,B,C.∀U1,U2,U3.
-  ∀r1,r1':binary_relation A B U1 U2.
-   ∀r2,r2':binary_relation B C U2 U3.
-    r1 = r1' → r2 = r2' → r1 ∘ r2 = r1' ∘ r2'.
- intros 14; split; intro;
- cases H2; clear H2; cases H3; clear H3;
- [ lapply (if ?? (H x w) H2) | lapply (fi ?? (H x w) H2) ]
- [ lapply (if ?? (H1 w y) H4)| lapply (fi ?? (H1 w y) H4) ]
- exists; try assumption;
- split; assumption.
+ apply (.= H1);
+  [ apply H | apply (H \sup -1) ]
 qed.
 
-include "logic/equality.ma".
-
-definition singleton ≝ λA:Type.λa:A.{b | a=b}.
-
 interpretation "singleton" 'singl a = (singleton _ a).