]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/categories.ma
...
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / categories.ma
index f09e0ee6c109af1f9ee1572aeaa77ca394634d97..86bc2529cda85de608a0ab1087396d277af75efe 100644 (file)
@@ -286,13 +286,11 @@ record category2 : Type3 ≝
  }.
 
 notation "'ASSOC'" with precedence 90 for @{'assoc}.
-notation "'ASSOC1'" with precedence 90 for @{'assoc1}.
-notation "'ASSOC2'" with precedence 90 for @{'assoc2}.
 
 interpretation "category2 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
-interpretation "category2 assoc" 'assoc1 = (comp_assoc2 ________).
+interpretation "category2 assoc" 'assoc = (comp_assoc2 ________).
 interpretation "category1 composition" 'compose x y = (fun21 ___ (comp1 ____) y x).
-interpretation "category1 assoc" 'assoc1 = (comp_assoc1 ________).
+interpretation "category1 assoc" 'assoc = (comp_assoc1 ________).
 interpretation "category composition" 'compose x y = (fun2 ___ (comp ____) y x).
 interpretation "category assoc" 'assoc = (comp_assoc ________).
 
@@ -404,3 +402,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.