]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/categories.ma
equivalence_relations made uniform w.r.t. universe level
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / categories.ma
index 5afda73d99865ae143aa257d93a1e4e006996dc4..ea246ef6c8e501c58a50b571763ccb50ee693812 100644 (file)
@@ -57,7 +57,7 @@ definition reflexive1 ≝ λA:Type1.λR:A→A→CProp1.∀x:A.R x x.
 definition symmetric1 ≝ λC:Type1.λlt:C→C→CProp1. ∀x,y:C.lt x y → lt y x.
 definition transitive1 ≝ λA:Type1.λR:A→A→CProp1.∀x,y,z:A.R x y → R y z → R x z.
 
-record equivalence_relation1 (A:Type1) : Type1 ≝
+record equivalence_relation1 (A:Type1) : Type2 ≝
  { eq_rel1:2> A → A → CProp1;
    refl1: reflexive1 ? eq_rel1;
    sym1: symmetric1 ? eq_rel1;
@@ -88,7 +88,7 @@ definition reflexive2 ≝ λA:Type2.λR:A→A→CProp2.∀x:A.R x x.
 definition symmetric2 ≝ λC:Type2.λlt:C→C→CProp2. ∀x,y:C.lt x y → lt y x.
 definition transitive2 ≝ λA:Type2.λR:A→A→CProp2.∀x,y,z:A.R x y → R y z → R x z.
 
-record equivalence_relation2 (A:Type2) : Type2 ≝
+record equivalence_relation2 (A:Type2) : Type3 ≝
  { eq_rel2:2> A → A → CProp2;
    refl2: reflexive2 ? eq_rel2;
    sym2: symmetric2 ? eq_rel2;
@@ -100,6 +100,20 @@ record setoid2: Type3 ≝
    eq2: equivalence_relation2 carr2
  }.
 
+definition setoid2_of_setoid1: setoid1 → setoid2.
+ intro;
+ constructor 1;
+  [ apply (carr1 s)
+  | constructor 1;
+    [ apply (eq_rel1 s);
+      apply (eq1 s)
+    | apply (refl1 s)
+    | apply (sym1 s)
+    | apply (trans1 s)]]
+qed.
+
+(*coercion setoid2_of_setoid1.*)
+
 (*
 definition Leibniz: Type → setoid.
  intro;
@@ -182,6 +196,7 @@ definition CPROP: setoid1.
         [ apply (H4 (H2 x1)) | apply (H3 (H5 z1))]]]
 qed.
 
+alias symbol "eq" = "setoid1 eq".
 definition if': ∀A,B:CPROP. A = B → A → B.
  intros; apply (if ?? e); assumption.
 qed.
@@ -282,12 +297,12 @@ interpretation "category assoc" 'assoc = (comp_assoc ________).
 
 (* bug grande come una casa?
    Ma come fa a passare la quantificazione larga??? *)
-definition unary_morphism_setoid: setoid → setoid → setoid.
+definition unary_morphism_setoid: setoid → setoid → setoid1.
  intros;
  constructor 1;
   [ apply (unary_morphism s s1);
   | constructor 1;
-     [ intros (f g); apply (∀a:s. f a = g a);
+     [ intros (f g); apply (∀a:s. eq ? (f a) (g a));
      | intros 1; simplify; intros; apply refl;
      | simplify; intros; apply sym; apply H;
      | simplify; intros; apply trans; [2: apply H; | skip | apply H1]]]
@@ -336,7 +351,9 @@ definition unary_morphism1_setoid1: setoid1 → setoid1 → setoid2.
  constructor 1;
   [ apply (unary_morphism1 s s1);
   | constructor 1;
-     [ intros (f g); apply (∀a: carr1 s. f a = g a);
+     [ intros (f g);
+       alias symbol "eq" = "setoid1 eq".
+       apply (∀a: carr1 s. f a = g a);
      | intros 1; simplify; intros; apply refl1;
      | simplify; intros; apply sym1; apply H;
      | simplify; intros; apply trans1; [2: apply H; | skip | apply H1]]]
@@ -370,7 +387,11 @@ definition prop11_SET1 :
 intros; apply (prop11 A B w a b e);
 qed.
 
+definition hint: Type_OF_category2 SET1 → Type1.
+ intro; whd in t; apply (carr1 t);
+qed.
+coercion hint.
+
 interpretation "SET dagger" 'prop1 h = (prop11_SET1 _ _ _ _ _ h).
-notation "hbox(a break ⇒ b)" right associative with precedence 20 for @{ 'Imply $a $b }.
 interpretation "unary morphism1" 'Imply a b = (arrows2 SET1 a b).
-interpretation "SET1 eq" 'eq x y = (eq_rel1 _ (eq'' _) x y).
\ No newline at end of file
+interpretation "SET1 eq" 'eq x y = (eq_rel1 _ (eq'' _) x y).