]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/sets/setoids.ma
eq -> eq0 renaming
[helm.git] / helm / software / matita / nlibrary / sets / setoids.ma
index 58fe2184456dc39dddfa35854563ac686c960b82..c159be66a6633198187c9ee9cd5ae79ed888b131 100644 (file)
@@ -16,25 +16,15 @@ include "logic/connectives.ma".
 include "properties/relations.ma".
 include "hints_declaration.ma".
 
-(*
-notation "hvbox(a break = \sub \ID b)" non associative with precedence 45
-for @{ 'eqID $a $b }.
-
-notation > "hvbox(a break =_\ID b)" non associative with precedence 45
-for @{ cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? $a $b }.
-
-interpretation "ID eq" 'eqID x y = (cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? x y).
-*)
-
 nrecord setoid : Type[1] ≝
  { carr:> Type[0];
-   eq: equivalence_relation carr
+   eq0: equivalence_relation carr
  }.
 
-interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq t) x y).
+interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq0 t) x y).
 
 notation > "hvbox(a break =_0 b)" non associative with precedence 45
-for @{ eq_rel ? (eq ?) $a $b }.
+for @{ eq_rel ? (eq0 ?) $a $b }.
 
 interpretation "setoid symmetry" 'invert r = (sym ???? r).
 notation ".= r" with precedence 50 for @{'trans $r}.
@@ -42,7 +32,7 @@ interpretation "trans" 'trans r = (trans ????? r).
 
 nrecord unary_morphism (A,B: setoid) : Type[0] ≝
  { fun1:1> A → B;
-   prop1: ∀a,a'. eq ? a a' → eq ? (fun1 a) (fun1 a')
+   prop1: ∀a,a'. a = a' → fun1 a = fun1 a'
  }.
 
 notation "† c" with precedence 90 for @{'prop1 $c }.
@@ -59,13 +49,16 @@ ndefinition unary_morph_setoid : setoid → setoid → setoid.
 ##| #f; #g; #h; #H1; #H2; #x; #x'; #Hx; napply (trans … (H1 …) (H2 …)); //; ##]
 nqed.
 
-unification hint 0
- (∀o1,o2. (λx,y:Type[0].True) (carr (unary_morph_setoid o1 o2)) (unary_morphism o1 o2)).
+alias symbol "hint_decl" (instance 1) = "hint_decl_Type1".
+unification hint 0 ≔ o1,o2 ; 
+     X ≟ unary_morph_setoid o1 o2
+  (* ------------------------------ *) ⊢ 
+     carr X ≡ unary_morphism o1 o2.
 
 interpretation "prop2" 'prop2 l r = (prop1 ? (unary_morph_setoid ??) ? ?? l ?? r).
 
 nlemma unary_morph_eq: ∀A,B.∀f,g: unary_morphism A B. (∀x. f x = g x) → f=g.
-/3/. nqed.
+#A B f g H x1 x2 E; napply (.= †E); napply H; nqed.
 
 nlemma mk_binary_morphism:
  ∀A,B,C: setoid. ∀f: A → B → C. (∀a,a',b,b'. a=a' → b=b' → f a b = f a' b') →
@@ -88,8 +81,8 @@ ndefinition comp_unary_morphisms:
 nqed.
 
 unification hint 0 ≔ o1,o2,o3:setoid,f:unary_morphism o2 o3,g:unary_morphism o1 o2;
-   R ≟ (mk_unary_morphism ?? (composition … f g)
-        (prop1 ?? (comp_unary_morphisms o1 o2 o3 f g)))
+   R ≟ mk_unary_morphism ?? (composition … f g)
+        (prop1 ?? (comp_unary_morphisms o1 o2 o3 f g))
  (* -------------------------------------------------------------------- *) ⊢
                               fun1 ?? R ≡ (composition … f g).
 
@@ -100,4 +93,4 @@ ndefinition comp_binary_morphisms:
 #o1; #o2; #o3; napply mk_binary_morphism
  [ #f; #g; napply (comp_unary_morphisms … f g) (*CSC: why not ∘?*)
  | #a; #a'; #b; #b'; #ea; #eb; #x; #x'; #Hx; nnormalize; /3/ ]
-nqed.
\ No newline at end of file
+nqed.