]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/subsets.ma
1. CProp_n fixed to be equal to Type_n to better understand what is happening
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / subsets.ma
index 9847d4c55d07c2f3d5f2f83c9179446c9a2793d3..73872bfed5f8bceeb5605d28fd614abb6fe83cc9 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,14 +147,14 @@ 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.
 
 (* senza questo exT "fresco", universe inconsistency *)
-inductive exT (A:Type) (P:A→CProp) : CProp ≝
+inductive exT (A:Type0) (P:A→CProp0) : CProp0 ≝
   ex_introT: ∀w:A. P w → exT A P.
 
 definition big_union:
@@ -162,7 +162,7 @@ definition big_union:
  intros; constructor 1;
   [ intro; whd; whd in A; whd in I;
     apply ({x | (*∃i:carr I. x ∈ t i*) exT (carr I) (λi. x ∈ t i)});
-    simplify; intros; split; intros; cases H; clear H; exists; [1,3:apply w]
+    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]
@@ -208,8 +208,8 @@ definition SUBSETS: objs1 SET → OAlgebra.
         | (* senza questa change, universe inconsistency *)
           whd; change in ⊢ (? ? (λ_:%.?)) with (carr I);
           exists; [apply w] assumption]]
-  | intros; intros 2; cases (H (singleton ? a) ?);
+  | intros; intros 2; cases (f (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]]
-qed.
\ No newline at end of file
+qed.