]> matita.cs.unibo.it Git - helm.git/commitdiff
Niceness was just a temporary illusion :-(
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Sat, 3 Jan 2009 18:27:07 +0000 (18:27 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Sat, 3 Jan 2009 18:27:07 +0000 (18:27 +0000)
helm/software/matita/contribs/formal_topology/overlap/categories.ma
helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma

index f09e0ee6c109af1f9ee1572aeaa77ca394634d97..1452652fa6e8c30a45e2c007f9b6d4a78b7d293a 100644 (file)
@@ -404,3 +404,8 @@ coercion Type1_OF_SET1.
 interpretation "SET dagger" 'prop1 h = (prop11_SET1 _ _ _ _ _ h).
 interpretation "unary morphism1" 'Imply a b = (arrows2 SET1 a b).
 interpretation "SET1 eq" 'eq x y = (eq_rel1 _ (eq'' _) x y).
+
+lemma unary_morphism1_of_arrows1_SET1: ∀S,T. (S ⇒ T) → unary_morphism1 S T.
+ intros; apply t;
+qed.
+coercion unary_morphism1_of_arrows1_SET1.
\ No newline at end of file
index 640a936ee21a194e1cd35a291ebce95dd2870d15..be985a43cefd551e01fd6ea942dcc220dea8b622 100644 (file)
@@ -135,11 +135,6 @@ definition continuous_relation_comp:
      | apply refl1]]
 qed.
 
-lemma hintx: ∀S,T. (S ⇒ T) → unary_morphism1 S T.
- intros; apply t;
-qed.
-coercion hintx.
-
 definition BTop: category2.
  constructor 1;
   [ apply basic_topology
@@ -160,13 +155,16 @@ definition BTop: category2.
        change with (b⎻* (a'⎻* (A o1 x)) = b'⎻*(a'⎻* (A o1 x)));
        alias symbol "trans" = "trans1".
        alias symbol "prop1" = "prop11".
-       apply (.= †(saturated o1 o2 a' (A o1 x) : ?));
-        [ apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1); ]
+       alias symbol "invert" = "setoid1 symmetry".
+       lapply (.= †(saturated o1 o2 a' (A o1 x) : ?));
+        [3: apply (b⎻* ); | 5: apply Hletin; |1,2: skip;
+        |apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1); ]
        change in e1 with (∀x.b⎻* (A o2 x) = b'⎻* (A o2 x));
        apply (.= (e1 (a'⎻* (A o1 x))));
        alias symbol "invert" = "setoid1 symmetry".
-       apply (†((saturated ?? a' (A o1 x) : ?) ^ -1));
-       apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1);]
+       lapply (†((saturated ?? a' (A o1 x) : ?) ^ -1));
+        [2: apply (b'⎻* ); |4: apply Hletin; | skip;
+        |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);