]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
The functor from BP to OBP has been defined (but no property proved yet).
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_topologies.ma
index 61334e54080fe84472942e6e88556d3ea02c9efe..1081f5beab493cd652ed0872749ffc8f6a9f1cc3 100644 (file)
@@ -108,7 +108,6 @@ theorem continuous_relation_eq_inv':
 qed.
 *)
 
-axiom daemon: False.
 definition continuous_relation_comp:
  ∀o1,o2,o3.
   continuous_relation_setoid o1 o2 →
@@ -167,7 +166,7 @@ definition BTop: category2.
         |apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1);]]
   | intros; simplify;
     change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1));
-    apply rule (#‡ASSOC1\sup -1);
+    apply rule (#‡ASSOC ^ -1);
   | intros; simplify;
     change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1);
     apply (#‡(id_neutral_right2 : ?));
@@ -176,6 +175,9 @@ definition BTop: category2.
     apply (#‡(id_neutral_left2 : ?));]
 qed.
 
+definition btop_carr: BTop → Type1 ≝ λo:BTop. carr1 (oa_P (carrbt o)).
+coercion btop_carr.
+
 (*
 (*CSC: unused! *)
 (* this proof is more logic-oriented than set/lattice oriented *)
@@ -201,4 +203,4 @@ theorem continuous_relation_eqS:
   [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
  apply Hcut2; assumption.
 qed.
-*)
\ No newline at end of file
+*)