]> 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 655c5076034328872447c628855bf65b1479dba9..86bc2529cda85de608a0ab1087396d277af75efe 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;
@@ -112,7 +112,7 @@ definition setoid2_of_setoid1: setoid1 → setoid2.
     | apply (trans1 s)]]
 qed.
 
-coercion setoid2_of_setoid1.
+(*coercion setoid2_of_setoid1.*)
 
 (*
 definition Leibniz: Type → setoid.
@@ -181,6 +181,7 @@ interpretation "prop11" 'prop1 c = (prop11 _____ c).
 interpretation "prop12" 'prop1 c = (prop12 _____ c).
 interpretation "prop2" 'prop2 l r = (prop2 ________ l r).
 interpretation "prop21" 'prop2 l r = (prop21 ________ l r).
+interpretation "prop22" 'prop2 l r = (prop22 ________ l r).
 interpretation "refl" 'refl = (refl ___).
 interpretation "refl1" 'refl = (refl1 ___).
 interpretation "refl2" 'refl = (refl2 ___).
@@ -285,19 +286,17 @@ 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 "category1 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
-interpretation "category1 assoc" 'assoc1 = (comp_assoc2 ________).
+interpretation "category2 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
+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 ________).
 
 (* 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);
@@ -387,12 +386,24 @@ definition prop11_SET1 :
 intros; apply (prop11 A B w a b e);
 qed.
 
-definition hint: Type_OF_category2 SET1 → Type1.
+definition setoid2_OF_category2: Type_OF_category2 SET1 → setoid2.
+ intro; apply (setoid2_of_setoid1 t); qed.
+coercion setoid2_OF_category2.
+
+definition objs2_OF_category1: Type_OF_category1 SET → objs2 SET1.
+ intro; apply (setoid1_of_setoid t); qed.
+coercion objs2_OF_category1.
+
+definition Type1_OF_SET1: Type_OF_category2 SET1 → Type1.
  intro; whd in t; apply (carr1 t);
 qed.
-coercion hint.
+coercion Type1_OF_SET1.
 
 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).
+
+lemma unary_morphism1_of_arrows1_SET1: ∀S,T. (S ⇒ T) → unary_morphism1 S T.
+ intros; apply t;
+qed.
+coercion unary_morphism1_of_arrows1_SET1.