]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/subsets.ma
WARNING: partial commit to try to understand something.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / subsets.ma
diff --git a/helm/software/matita/contribs/formal_topology/overlap/subsets.ma b/helm/software/matita/contribs/formal_topology/overlap/subsets.ma
new file mode 100644 (file)
index 0000000..17128ae
--- /dev/null
@@ -0,0 +1,139 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "logic/cprop_connectives.ma".
+include "categories.ma".
+
+record powerset_carrier (A: SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }.
+
+definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp2 ≝
+ λA:SET.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a.
+
+theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
+ intros 6; intros 2;
+ apply s1; apply s;
+ assumption.
+qed.
+
+definition powerset_setoid1: SET → SET1.
+ 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.
+
+interpretation "powerset" 'powerset A = (powerset_setoid1 A).
+
+interpretation "subset construction" 'subset \eta.x =
+ (mk_powerset_carrier _ (mk_unary_morphism1 _ CPROP x _)).
+
+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 ????? (prop11 ?? u ?? e));
+    cases e1; whd in s s1;
+    split; intro;
+     [ apply s; assumption
+     | apply s1; assumption]]
+qed.
+
+interpretation "mem" 'mem a S = (fun21 ___ (mem _) a S).
+
+definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
+ intros;
+ constructor 1;
+  [ apply (λU,V. subseteq_operator ? U V)
+  | intros;
+    cases e; cases e1;
+    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 "subseteq" 'subseteq U V = (fun21 ___ (subseteq _) U V).
+
+theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
+ intros 4; assumption.
+qed.
+
+theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
+ intros; apply transitive_subseteq_operator; [apply S2] assumption.
+qed.
+
+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 H; exists; [1,4: apply w]
+     [ apply (. #‡e); assumption
+     | apply (. #‡e1); assumption
+     | apply (. #‡(e \sup -1)); assumption;
+     | apply (. #‡(e1 \sup -1)); assumption]]
+qed.
+
+interpretation "overlaps" 'overlaps U V = (fun21 ___ (overlaps _) U V).
+
+definition intersects:
+ ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
+ intros;
+ constructor 1;
+  [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
+    intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1;
+  | intros;
+    split; intros 2; simplify in f ⊢ %;
+    [ apply (. (#‡e)‡(#‡e1)); assumption
+    | apply (. (#‡(e \sup -1))‡(#‡(e1 \sup -1))); assumption]]
+qed.
+
+interpretation "intersects" 'intersects U V = (fun21 ___ (intersects _) U V).
+
+definition union:
+ ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
+ intros;
+ constructor 1;
+  [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
+    intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1
+  | intros;
+    split; intros 2; simplify in f ⊢ %;
+    [ apply (. (#‡e)‡(#‡e1)); assumption
+    | apply (. (#‡(e \sup -1))‡(#‡(e1 \sup -1))); assumption]]
+qed.
+
+interpretation "union" 'union U V = (fun21 ___ (union _) U V).
+
+definition singleton: ∀A:setoid. unary_morphism1 A (Ω \sup A).
+ intros; constructor 1;
+  [ apply (λa:A.{b | eq ? a b}); unfold setoid1_of_setoid; simplify;
+    intros; simplify;
+    split; intro;
+    apply (.= e1);
+     [ apply e | apply (e \sup -1) ]
+  | unfold setoid1_of_setoid; simplify;
+    intros; split; intros 2; simplify in f ⊢ %; apply trans;
+     [ apply a |4: apply a'] try assumption; apply sym; assumption]
+qed.
+
+interpretation "singleton" 'singl a = (fun11 __ (singleton _) a).
\ No newline at end of file