X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matitaB%2Fmatita%2Fnlibrary%2Fsets%2Fsetoids.ma;h=29f4c910c8963e55dcedfda64ea2c0d045a3ebf8;hb=f34f2623a3133e235331d0c0c1830ec213dd09f1;hp=e40dad6f6d20e4e3a1f7624b6f28b5d6056a6ef3;hpb=cacbe3c6493ddce76c4c13379ade271d8dd172e8;p=helm.git diff --git a/matitaB/matita/nlibrary/sets/setoids.ma b/matitaB/matita/nlibrary/sets/setoids.ma index e40dad6f6..29f4c910c 100644 --- a/matitaB/matita/nlibrary/sets/setoids.ma +++ b/matitaB/matita/nlibrary/sets/setoids.ma @@ -16,7 +16,7 @@ include "logic/connectives.ma". include "properties/relations.ma". include "hints_declaration.ma". -nrecord setoid : Type[1] ≝ { +record setoid : Type[1] ≝ { carr:> Type[0]; eq0: equivalence_relation carr }. @@ -46,7 +46,7 @@ interpretation "trans" 'trans r = (trans ????? r). notation > ".=_0 r" with precedence 50 for @{'trans_x0 $r}. interpretation "trans_x0" 'trans_x0 r = (trans ????? r). -nrecord unary_morphism (A,B: setoid) : Type[0] ≝ { +record unary_morphism (A,B: setoid) : Type[0] ≝ { fun1:1> A → B; prop1: ∀a,a'. a = a' → fun1 a = fun1 a' }. @@ -64,13 +64,13 @@ notation "┼_0 c" with precedence 89 for @{'prop1_x0 $c }. notation "l ╪_0 r" with precedence 89 for @{'prop2_x0 $l $r }. interpretation "prop1_x0" 'prop1_x0 c = (prop1 ????? c). -ndefinition unary_morph_setoid : setoid → setoid → setoid. -#S1; #S2; @ (S1 ⇒_0 S2); @; -##[ #f; #g; napply (∀x,x'. x=x' → f x = g x'); -##| #f; #x; #x'; #Hx; napply (.= †Hx); napply #; -##| #f; #g; #H; #x; #x'; #Hx; napply ((H … Hx^-1)^-1); -##| #f; #g; #h; #H1; #H2; #x; #x'; #Hx; napply (trans … (H1 …) (H2 …)); //; ##] -nqed. +definition unary_morph_setoid : setoid → setoid → setoid. +#S1 #S2 @(mk_setoid … (S1 ⇒_0 S2)) % +[ #f #g @(∀x,x'. x=x' → f x = g x') +| #f #x #x' #Hx @(.= †Hx) @# +| #f #g #H #x #x' #Hx @((H … Hx^-1)^-1) +| #f #g #h #H1 #H2 #x #x' #Hx @(trans … (H1 …) (H2 …)) // ] +qed. alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". unification hint 0 ≔ o1,o2 ; @@ -81,26 +81,27 @@ unification hint 0 ≔ o1,o2 ; interpretation "prop2" 'prop2 l r = (prop1 ? (unary_morph_setoid ??) ? ?? l ?? r). interpretation "prop2_x0" 'prop2_x0 l r = (prop1 ? (unary_morph_setoid ??) ? ?? l ?? r). -nlemma unary_morph_eq: ∀A,B.∀f,g:A ⇒_0 B. (∀x. f x = g x) → f = g. -#A B f g H x1 x2 E; napply (.= †E); napply H; nqed. +lemma unary_morph_eq: ∀A,B.∀f,g:A ⇒_0 B. (∀x. f x = g x) → f = g. +#A #B #f #g #H #x1 #x2 #E @(.= †E) @H +qed. -nlemma mk_binary_morphism: +lemma 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') → A ⇒_0 (unary_morph_setoid B C). - #A; #B; #C; #f; #H; @; ##[ #x; @ (f x) ] #a; #a'; #Ha [##2: napply unary_morph_eq; #y] - /2/. -nqed. + #A #B #C #f #H % [ #x % [@(f x)|/2/] |#a #a' #Ha @unary_morph_eq #y] + /2/ +qed. -ndefinition composition ≝ +definition composition ≝ λo1,o2,o3:Type[0].λf:o2 → o3.λg: o1 → o2.λx.f (g x). interpretation "function composition" 'compose f g = (composition ??? f g). -ndefinition comp_unary_morphisms: +definition comp_unary_morphisms: ∀o1,o2,o3:setoid.o2 ⇒_0 o3 → o1 ⇒_0 o2 → o1 ⇒_0 o3. -#o1; #o2; #o3; #f; #g; @ (f ∘ g); - #a; #a'; #e; nnormalize; napply (.= †(†e)); napply #. -nqed. +#o1 #o2 #o3 #f #g % [@(f ∘ g)] + #a #a' #e normalize @(.= †(†e)) @# +qed. unification hint 0 ≔ o1,o2,o3:setoid,f:o2 ⇒_0 o3,g:o1 ⇒_0 o2; R ≟ mk_unary_morphism o1 o3 @@ -109,13 +110,13 @@ unification hint 0 ≔ o1,o2,o3:setoid,f:o2 ⇒_0 o3,g:o1 ⇒_0 o2; (* -------------------------------------------------------------------- *) ⊢ fun1 o1 o3 R ≡ composition ??? (fun1 o2 o3 f) (fun1 o1 o2 g). -ndefinition comp_binary_morphisms: +definition comp_binary_morphisms: ∀o1,o2,o3.(o2 ⇒_0 o3) ⇒_0 ((o1 ⇒_0 o2) ⇒_0 (o1 ⇒_0 o3)). -#o1; #o2; #o3; napply mk_binary_morphism - [ #f; #g; napply (comp_unary_morphisms ??? f g) +#o1 #o2 #o3 @mk_binary_morphism + [ #f #g @(comp_unary_morphisms ??? f g) (* CSC: why not ∘? GARES: because the coercion to FunClass is not triggered if there are no "extra" arguments. We could fix that in the refiner *) - | #a; #a'; #b; #b'; #ea; #eb; #x; #x'; #Hx; nnormalize; /3/ ] -nqed. + | #a #a' #b #b' #ea #eb #x #x' #Hx normalize /3/ ] +qed.