]> matita.cs.unibo.it Git - helm.git/commitdiff
Some fixes.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Fri, 26 Dec 2008 18:12:47 +0000 (18:12 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Fri, 26 Dec 2008 18:12:47 +0000 (18:12 +0000)
helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma

index 5b08b59170acbd478bc720c8f646e102ba09b256..9cbaada80b0b899cc74e43b0c530d097f9d334ce 100644 (file)
@@ -47,7 +47,8 @@ lemma xxx : ∀x.carr x → carr1 (setoid1_of_setoid x). intros; assumption; qed
 coercion xxx nocomposites.
 
 lemma down_p : ∀S,I:SET.∀u:S⇒S.∀c:arrows1 SET I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
-intros; unfold uncurry_arrows; apply (†(†H));
+intros; unfold uncurry_arrows; change in c with (I ⇒ S);
+apply (†(†H));
 qed.
 
 alias symbol "eq" = "setoid eq".
@@ -122,7 +123,7 @@ definition convergent_relation_space_composition:
  ∀o1,o2,o3: concrete_space.
   binary_morphism1
    (convergent_relation_space_setoid o1 o2)
-   (convergentin ⊢ (? (? ? ? (? ? ? (? ? ? ? ? (? ? ? (? ? ? (% ? ?))) ?)) ?) ? ? ?)_relation_space_setoid o2 o3)
+   (convergent_relation_space_setoid o2 o3)
    (convergent_relation_space_setoid o1 o3).
  intros; constructor 1;
      [ intros; whd in c c1 ⊢ %;