]> matita.cs.unibo.it Git - helm.git/commitdiff
An hint moved to the right place.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 6 Jan 2009 01:39:36 +0000 (01:39 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 6 Jan 2009 01:39:36 +0000 (01:39 +0000)
helm/software/matita/contribs/formal_topology/overlap/categories.ma
helm/software/matita/contribs/formal_topology/overlap/relations.ma

index 3a7614fca8b652e8668375d3bc1024bc6b3166a4..3ee3915074dd3db68941cbe032c9b4bcb8f3fa40 100644 (file)
@@ -372,6 +372,12 @@ definition Type1_OF_SET1: Type_OF_category2 SET1 → Type1.
 qed.
 coercion Type1_OF_SET1.
 
+definition Type_OF_setoid1_of_carr: ∀U. carr U → Type_OF_setoid1 ?(*(setoid1_of_SET U)*).
+ [ apply setoid1_of_SET; apply U
+ | intros; apply c;]
+qed.
+coercion Type_OF_setoid1_of_carr.
+
 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).
index eb60386fe3d1abc0831f7f747d7a5093062845c9..ed8b83bcbd1f3b17538cfe595ccf37d05f02efb0 100644 (file)
@@ -175,12 +175,6 @@ lemma extS_com: ∀o1,o2,o3,c1,c2,S. extS o1 o3 (c2 ∘ c1) S = extS o1 o2 c1 (e
 qed.
 *)
 
-lemma hint: ∀U. carr U → Type_OF_setoid1 ?(*(setoid1_of_SET U)*).
- [ apply setoid1_of_SET; apply U
- | intros; apply c;]
-qed.
-coercion hint.
-
 (* the same as ⋄ for a basic pair *)
 definition image: ∀U,V:REL. binary_morphism1 (arrows1 ? U V) (Ω \sup U) (Ω \sup V).
  intros; constructor 1;
@@ -327,4 +321,4 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → ORelation (S
      | cases x1; cases x2; clear x1 x2; exists; [apply w1]
         [ exists; [apply w] split; assumption;
         | assumption; ]]]
-qed. (*sistemare anche l'hint da un'altra parte *)
+qed.