]> matita.cs.unibo.it Git - helm.git/commitdiff
Some clean up.
authorEnrico Tassi <enrico.tassi@inria.fr>
Mon, 22 Dec 2008 18:07:42 +0000 (18:07 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Mon, 22 Dec 2008 18:07:42 +0000 (18:07 +0000)
helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma
helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
helm/software/matita/library/datatypes/categories.ma

index 4c61f0fb47753fc191b383a9f74bff27c0d718de..0cb3fe1172a341f7ac5fe1ba53530c7d5d13adff 100644 (file)
@@ -26,11 +26,6 @@ constructor 1; [apply bool] constructor 1;
   try assumption; apply I]
 qed.
 
-definition setoid_OF_SET: objs1 SET → setoid.
- intros; apply o; qed.
-
-coercion setoid_OF_SET.
-
 lemma IF_THEN_ELSE_p :
   ∀S:setoid.∀a,b:S.∀x,y:BOOL.x = y → 
     (λm.match m with [ true ⇒ a | false ⇒ b ]) x =
@@ -297,4 +292,4 @@ split;
    | apply ((comp_assoc1 ????? H* G* F* ));]
 | intros; split; unfold ORelation_composition; simplify; apply id_neutral_left1;
 | intros; split; unfold ORelation_composition; simplify; apply id_neutral_right1;]
-qed.
\ No newline at end of file
+qed.
index 80ee4649fee1f149370a456e5c81d28f9375d26e..5b08b59170acbd478bc720c8f646e102ba09b256 100644 (file)
@@ -118,13 +118,6 @@ definition rp'': ∀CS1,CS2.convergent_relation_space_setoid CS1 CS2 → arrows1
 
 coercion rp''.
 
-definition prop_1_SET : 
- ∀A,B:SET.∀w:arrows1 SET A B.∀a,b:A.eq1 ? a b→eq1 ? (w a) (w b).
-intros; apply (prop_1 A B w a b H);
-qed.
-
-interpretation "SET dagger" 'prop1 h = (prop_1_SET _ _ _ _ _ h).
-
 definition convergent_relation_space_composition:
  ∀o1,o2,o3: concrete_space.
   binary_morphism1
@@ -170,4 +163,4 @@ definition CSPA: category1.
   | intros; simplify;
     change with (id1 ? o2 ∘ a = a);
     apply (id_neutral_left1 : ?);]
-qed.
\ No newline at end of file
+qed.
index ca49cbdd327f2909da9a3fca5b7e9ce8bbf39730..65bc8ac6bb98c077f175223ca61814a17c69efe5 100644 (file)
@@ -236,4 +236,17 @@ definition SET: category1.
   | intros; simplify; whd; intros; simplify; apply refl;
   | intros; simplify; whd; intros; simplify; apply refl;
   ]
-qed.
\ No newline at end of file
+qed.
+
+definition setoid_OF_SET: objs1 SET → setoid.
+ intros; apply o; qed.
+
+coercion setoid_OF_SET.
+
+
+definition prop_1_SET : 
+ ∀A,B:SET.∀w:arrows1 SET A B.∀a,b:A.eq1 ? a b→eq1 ? (w a) (w b).
+intros; apply (prop_1 A B w a b H);
+qed.
+
+interpretation "SET dagger" 'prop1 h = (prop_1_SET _ _ _ _ _ h).