]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/datatypes/categories.ma
Setoids are now more pervasive.
[helm.git] / helm / software / matita / library / datatypes / categories.ma
index 0eb9b681863648c21f29a48c2262c3b77398db8d..e90e1457d222d23702ec37792f252ac67a270b6a 100644 (file)
@@ -26,19 +26,90 @@ record setoid : Type ≝
    eq: equivalence_relation carr
  }.
 
+definition reflexive1 ≝ λA:Type.λR:A→A→CProp.∀x:A.R x x.
+definition symmetric1 ≝ λC:Type.λlt:C→C→CProp. ∀x,y:C.lt x y → lt y x.
+definition transitive1 ≝ λA:Type.λR:A→A→CProp.∀x,y,z:A.R x y → R y z → R x z.
+
+record equivalence_relation1 (A:Type) : Type ≝
+ { eq_rel1:2> A → A → CProp;
+   refl1: reflexive1 ? eq_rel1;
+   sym1: symmetric1 ? eq_rel1;
+   trans1: transitive1 ? eq_rel1
+ }.
+
+record setoid1: Type ≝
+ { carr1:> Type;
+   eq1: equivalence_relation1 carr1
+ }.
+
+definition setoid1_of_setoid: setoid → setoid1.
+ intro;
+ constructor 1;
+  [ apply (carr s)
+  | constructor 1;
+    [ apply (eq_rel s);
+      apply (eq s)
+    | apply (refl s)
+    | apply (sym s)
+    | apply (trans s)]]
+qed.
+
+coercion setoid1_of_setoid.
+
+(*
+definition Leibniz: Type → setoid.
+ intro;
+ constructor 1;
+  [ apply T
+  | constructor 1;
+     [ apply (λx,y:T.cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? x y)
+     | alias id "refl_eq" = "cic:/matita/logic/equality/eq.ind#xpointer(1/1/1)".
+       apply refl_eq
+     | alias id "sym_eq" = "cic:/matita/logic/equality/sym_eq.con".
+       apply sym_eq
+     | alias id "trans_eq" = "cic:/matita/logic/equality/trans_eq.con".
+       apply trans_eq ]]
+qed.
+
+coercion Leibniz.
+*)
+
+interpretation "setoid1 eq" 'eq x y = (eq_rel1 _ (eq1 _) x y).
 interpretation "setoid eq" 'eq x y = (eq_rel _ (eq _) x y).
-notation ".= r" with precedence 50 for @{trans ????? $r}.
+interpretation "setoid1 symmetry" 'invert r = (sym1 ____ r).
 interpretation "setoid symmetry" 'invert r = (sym ____ r).
+notation ".= r" with precedence 50 for @{'trans $r}.
+interpretation "trans1" 'trans r = (trans1 _____ r).
+interpretation "trans" 'trans r = (trans _____ r).
 
-record binary_morphism (A,B,C: setoid) : Type ≝
+record unary_morphism (A,B: setoid1) : Type ≝
+ { fun_1:1> A → B;
+   prop_1: ∀a,a'. eq1 ? a a' → eq1 ? (fun_1 a) (fun_1 a')
+ }.
+
+record binary_morphism (A,B,C:setoid) : Type ≝
  { fun:2> A → B → C;
    prop: ∀a,a',b,b'. eq ? a a' → eq ? b b' → eq ? (fun a b) (fun a' b')
  }.
 
-notation "# r" with precedence 60 for @{prop ???????? (refl ???) $r}.
-notation "r #" with precedence 60 for @{prop ???????? $r (refl ???)}.
+record binary_morphism1 (A,B,C:setoid1) : Type ≝
+ { fun1:2> A → B → C;
+   prop1: ∀a,a',b,b'. eq1 ? a a' → eq1 ? b b' → eq1 ? (fun1 a b) (fun1 a' b')
+ }.
 
-definition CPROP: setoid.
+notation "hbox(a break ⇒ b)" right associative with precedence 20 for @{ 'Imply $a $b }.
+interpretation "unary morphism" 'Imply a b = (unary_morphism a b).
+
+notation "† c" with precedence 90 for @{'prop1 $c }.
+notation "l ‡ r" with precedence 90 for @{'prop $l $r }.
+notation "#" with precedence 90 for @{'refl}.
+interpretation "prop_1" 'prop1 c = (prop_1 _____ c).
+interpretation "prop1" 'prop l r = (prop1 ________ l r).
+interpretation "prop" 'prop l r = (prop ________ l r).
+interpretation "refl1" 'refl = (refl1 ___).
+interpretation "refl" 'refl = (refl ___).
+
+definition CPROP: setoid1.
  constructor 1;
   [ apply CProp
   | constructor 1;
@@ -49,6 +120,46 @@ definition CPROP: setoid.
         [ apply (H4 (H2 H6)) | apply (H3 (H5 H6))]]]
 qed.
 
+definition if': ∀A,B:CPROP. A = B → A → B.
+ intros; apply (if ?? H); assumption.
+qed.
+
+notation ". r" with precedence 50 for @{'if $r}.
+interpretation "if" 'if r = (if' __ r).
+
+definition and_morphism: binary_morphism1 CPROP CPROP CPROP.
+ constructor 1;
+  [ apply And
+  | intros; split; intro; cases H2; split;
+     [ apply (if ?? H a1)
+     | apply (if ?? H1 b1)
+     | apply (fi ?? H a1)
+     | apply (fi ?? H1 b1)]]
+qed.
+
+interpretation "and_morphism" 'and a b = (fun1 ___ and_morphism a b).
+
+definition or_morphism: binary_morphism1 CPROP CPROP CPROP.
+ constructor 1;
+  [ apply Or
+  | intros; split; intro; cases H2; [1,3:left |2,4: right]
+     [ apply (if ?? H a1)
+     | apply (fi ?? H a1)
+     | apply (if ?? H1 b1)
+     | apply (fi ?? H1 b1)]]
+qed.
+
+interpretation "or_morphism" 'or a b = (fun1 ___ or_morphism a b).
+
+definition if_morphism: binary_morphism1 CPROP CPROP CPROP.
+ constructor 1;
+  [ apply (λA,B. A → B)
+  | intros; split; intros;
+     [ apply (if ?? H1); apply H2; apply (fi ?? H); assumption
+     | apply (fi ?? H1); apply H2; apply (if ?? H); assumption]]
+qed.
+
+(*
 definition eq_morphism: ∀S:setoid. binary_morphism S S CPROP.
  intro;
  constructor 1;
@@ -61,6 +172,7 @@ definition eq_morphism: ∀S:setoid. binary_morphism S S CPROP.
        apply (.= H2);
        apply (H1 \sup -1)]]
 qed.
+*)
 
 record category : Type ≝
  { objs:> Type;
@@ -73,5 +185,21 @@ record category : Type ≝
    id_neutral_right: ∀o1,o2. ∀a: arrows o1 o2. comp ??? a (id o2) = a
  }.
 
+record category1 : Type ≝
+ { objs1:> Type;
+   arrows1: objs1 → objs1 → setoid1;
+   id1: ∀o:objs1. arrows1 o o;
+   comp1: ∀o1,o2,o3. binary_morphism1 (arrows1 o1 o2) (arrows1 o2 o3) (arrows1 o1 o3);
+   comp_assoc1: ∀o1,o2,o3,o4. ∀a12,a23,a34.
+    comp1 o1 o3 o4 (comp1 o1 o2 o3 a12 a23) a34 = comp1 o1 o2 o4 a12 (comp1 o2 o3 o4 a23 a34);
+   id_neutral_left1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? (id1 o1) a = a;
+   id_neutral_right1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? a (id1 o2) = a
+ }.
+
+notation "'ASSOC'" with precedence 90 for @{'assoc}.
+notation "'ASSOC1'" with precedence 90 for @{'assoc1}.
+
+interpretation "category1 composition" 'compose x y = (fun1 ___ (comp1 ____) x y).
+interpretation "category1 assoc" 'assoc1 = (comp_assoc1 ________).
 interpretation "category composition" 'compose x y = (fun ___ (comp ____) x y).
-notation "'ASSOC'" with precedence 90 for @{comp_assoc ????????}.
\ No newline at end of file
+interpretation "category assoc" 'assoc = (comp_assoc ________).