]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/subsets.ma
virtuals for () removed and bound to 'o'
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / subsets.ma
index 4bbe7fdbebd35ce17de26a3bbc5ef5f672503477..3aaf2d3d6b34b6da73f52a678e6cfb7d95c5f73b 100644 (file)
@@ -16,7 +16,7 @@ include "categories.ma".
 
 record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }.
 
-definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp1 ≝
+definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp0 ≝
  λA:objs1 SET.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a.
 
 theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
@@ -147,58 +147,21 @@ definition big_intersects:
   [ intro; whd; whd in I;
     apply ({x | ∀i:I. x ∈ t i});
     simplify; intros; split; intros; [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ]
-    apply H;
+    apply f;
   | intros; split; intros 2; simplify in f ⊢ %; intro;
      [ apply (. (#‡(e i))); apply f;
      | apply (. (#‡(e i)\sup -1)); apply f]]
 qed.
 
-axiom (*definition*) big_union:
+definition big_union:
  ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)).
-(* intros; constructor 1;
+ intros; constructor 1;
   [ intro; whd; whd in A; whd in I;
-    apply ({x | ∃i:I. x ∈ t i});
-    simplify; intros; split; intros; cases H; clear H; exists; [1,3:apply w]
+    apply ({x | ∃i:carr I. x ∈ t i });
+    simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w]
     [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ]
     apply x;
   | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w]
      [ apply (. (#‡(e w))); apply x;
      | apply (. (#‡(e w)\sup -1)); apply x]]
-qed.*)
-
-(* incluso prima non funziona piu' nulla *)
-include "o-algebra.ma".
-
-
-axiom daemon: False.
-definition SUBSETS: objs1 SET → OAlgebra.
- intro A; constructor 1;
-  [ apply (Ω \sup A);
-  | apply subseteq;
-  | apply overlaps;
-  | apply big_intersects;
-  | apply big_union;
-  | apply ({x | True});
-    simplify; intros; apply (refl1 ? (eq1 CPROP));
-  | apply ({x | False});
-    simplify; intros; apply (refl1 ? (eq1 CPROP));
-  | intros; whd; intros; assumption
-  | intros; whd; split; assumption
-  | intros; apply transitive_subseteq_operator; [2: apply f; | skip | assumption]
-  | intros; cases f; exists [apply w] assumption
-  | intros; intros 2; apply (f ? f1 i);
-  |(* intros; intros 2; apply f; exists; [apply i] assumption;*)
-  | intros 3; cases f;
-  | intros 3; constructor 1;
-  | intros; cases f; exists; [apply w]
-     [ assumption
-     | whd; intros; cases i; simplify; assumption]
-  |(* intros; split; intro;
-     [ cases f; cases x1; exists [apply w1] exists [apply w] assumption;
-     | cases H; cases x; exists; [apply w1] [assumption | exists; [apply w] assumption]]*)
-  | intros; intros 2; cases (H (singleton ? a) ?);
-     [ exists; [apply a] [assumption | change with (a = a); apply refl1;]
-     | change in x1 with (a = w); change with (mem A a q); apply (. (x1 \sup -1‡#));
-       assumption]]
-cases daemon;
-qed.
\ No newline at end of file
+qed.