(*-----------------------------------------------------------------*) ⊢
list S ≡ carr1 TT.
-notation > "B ⇒_0 C" right associative with precedence 72 for @{'umorph0 $B $C}.
-notation > "B ⇒_1 C" right associative with precedence 72 for @{'umorph1 $B $C}.
-notation "B ⇒\sub 0 C" right associative with precedence 72 for @{'umorph0 $B $C}.
-notation "B ⇒\sub 1 C" right associative with precedence 72 for @{'umorph1 $B $C}.
-
-interpretation "unary morphism 0" 'umorph0 A B = (unary_morphism A B).
-interpretation "unary morphism 1" 'umorph1 A B = (unary_morphism1 A B).
-
nlemma Sig: ∀S,T:setoid.∀P: S → (T → CPROP).
∀y,z:T.y = z → (∀x.y=z → P x y = P x z) → (Ex S (λx.P x y)) =_1 (Ex S (λx.P x z)).
nlemma cupA : ∀S.∀a,b,c:lang S.a ∪ b ∪ c = a ∪ (b ∪ c).
#S a b c; @; #w; *; /3/; *; /3/; nqed.
-nlemma setP : ∀S.∀A,B:Ω^S.∀x:S. A = B → (x ∈ A) = (x ∈ B).
-#S A B x; *; #H1 H2; @; ##[ napply H1 | napply H2] nqed.
-
-nlemma pset_ext : ∀S.∀A,B:Ω^S.(∀x:S. (x ∈ A) = (x ∈ B)) → A = B.
-#S A B H; @; #x; ncases (H x); #H1 H2; ##[ napply H1 | napply H2] nqed.
-
ndefinition if': ∀A,B:CPROP. A = B → A → B.
#A B; *; /2/. nqed.
ncoercion if : ∀A,B:CPROP. ∀p:A = B. A → B ≝ if' on _p : eq_rel1 ???? to ∀_:?.?.
-(* move in sets.ma? *)
-nlemma union_morphism : ∀A.Ω^A ⇒_1 (Ω^A ⇒_1 Ω^A).
-#X; napply (mk_binary_morphism1 … (λA,B.A ∪ B));
-#A1 A2 B1 B2 EA EB; napply pset_ext; #x;
-nchange in match (x ∈ (A1 ∪ B1)) with (?∨?);
-napply (.= (setP ??? x EA)‡#);
-napply (.= #‡(setP ??? x EB)); //;
-nqed.
-
-nlemma union_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A.
- #S A B; @ (A ∪ B); #x y Exy; @; *; #H1;
-##[##1,3: @; ##|##*: @2 ]
-##[##1,3: napply (. (Exy^-1)╪_1#)
-##|##2,4: napply (. Exy╪_1#)]
-nassumption;
-nqed.
-
-alias symbol "hint_decl" = "hint_decl_Type1".
-unification hint 0 ≔
- A : setoid, B,C : 𝛀^A;
- R ≟ (mk_ext_powerclass ? (B ∪ C) (ext_prop ? (union_is_ext ? B C)))
- (* ----------------------------------------------------------------*) ⊢
- ext_carr A R ≡ union ? (ext_carr ? B) (ext_carr ? C).
-
-unification hint 0 ≔ S:Type[0], A,B:Ω^S;
- MM ≟ mk_unary_morphism1 ??
- (λA.mk_unary_morphism1 ?? (λB.A ∪ B) (prop11 ?? (union_morphism S A)))
- (prop11 ?? (union_morphism S))
- (*-----------------------------------------------------*) ⊢
- fun11 ?? (fun11 ?? MM A) B ≡ A ∪ B.
-
-nlemma union_is_ext_morph:∀A:setoid.𝛀^A ⇒_1 (𝛀^A ⇒_1 𝛀^A).
-#A; napply (mk_binary_morphism1 … (union_is_ext …));
-#x1 x2 y1 y2 Ex Ey; napply (prop11 … (union_morphism A)); nassumption.
-nqed.
-
-unification hint 1 ≔
- AA : setoid, B,C : 𝛀^AA;
- A ≟ carr AA,
- R ≟ (mk_unary_morphism1 ??
- (λS:(*ext_powerclass_setoid AA*)𝛀^AA.
- mk_unary_morphism1 ??
- (λS':(*ext_powerclass_setoid AA*)𝛀^AA.
- mk_ext_powerclass AA (S ∪ S') (ext_prop AA (union_is_ext ? S S')))
- (prop11 ?? (union_is_ext_morph AA S)))
- (prop11 ?? (union_is_ext_morph AA))) ,
- BB ≟ (ext_carr ? B),
- CC ≟ (ext_carr ? C)
- (* ------------------------------------------------------*) ⊢
- ext_carr AA (R B C) ≡ union A BB CC.
-
-
(* theorem 16: 2 *)
nlemma oplus_cup : ∀S:Alpha.∀e1,e2:pre S.𝐋\p (e1 ⊕ e2) = 𝐋\p e1 ∪ 𝐋\p e2.
#S r1; ncases r1; #e1 b1 r2; ncases r2; #e2 b2;
nwhd in ⊢ (???(??%)?);
nchange in ⊢(???%?) with (𝐋\p (e1 + e2) ∪ ϵ (b1 || b2));
nchange in ⊢(???(??%?)?) with (𝐋\p (e1) ∪ 𝐋\p (e2));
-napply (.= #╪_1 (epsilon_or ???));
-napply (.= (cupA…)^-1);
-napply (.= (cupA…)╪_1#);
-napply (.= (#╪_1(cupC…))╪_1#);
-napply (.= (cupA…)^-1╪_1#);
-napply (.= (cupA…));
+napply (.=_1 #╪_1 (epsilon_or ???));
+napply (.=_1 (cupA…)^-1);
+napply (.=_1 (cupA…)╪_1#);
+napply (.=_1 (#╪_1(cupC…))╪_1#);
+napply (.=_1 (cupA…)^-1╪_1#);
+napply (.=_1 (cupA…));
//;
nqed.
include "properties/relations.ma".
include "hints_declaration.ma".
-nrecord setoid : Type[1] ≝
- { carr:> Type[0];
- eq0: equivalence_relation carr
- }.
+nrecord setoid : Type[1] ≝ {
+ carr:> Type[0];
+ eq0: equivalence_relation carr
+}.
interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq0 t) x y).
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] ≝
- { fun1:1> A → B;
- prop1: ∀a,a'. a = a' → fun1 a = fun1 a'
- }.
+nrecord unary_morphism (A,B: setoid) : Type[0] ≝ {
+ fun1:1> A → B;
+ prop1: ∀a,a'. a = a' → fun1 a = fun1 a'
+}.
+
+notation > "B ⇒_0 C" right associative with precedence 72 for @{'umorph0 $B $C}.
+notation "hvbox(B break ⇒\sub 0 C)" right associative with precedence 72 for @{'umorph0 $B $C}.
+interpretation "unary morphism 0" 'umorph0 A B = (unary_morphism A B).
notation "† c" with precedence 90 for @{'prop1 $c }.
notation "l ‡ r" with precedence 90 for @{'prop2 $l $r }.
interpretation "prop1_x0" 'prop1_x0 c = (prop1 ????? c).
ndefinition unary_morph_setoid : setoid → setoid → setoid.
-#S1; #S2; @ (unary_morphism S1 S2); @;
+#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);
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.
+ (* ----------------------------- *) ⊢
+ carr X ≡ o1 ⇒_0 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: unary_morphism A B. (∀x. f x = g x) → f=g.
+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.
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') →
- unary_morphism A (unary_morph_setoid B C).
+ 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.
interpretation "function composition" 'compose f g = (composition ??? f g).
ndefinition comp_unary_morphisms:
- ∀o1,o2,o3:setoid.
- unary_morphism o2 o3 → unary_morphism o1 o2 →
- unary_morphism o1 o3.
+ ∀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.
-unification hint 0 ≔ o1,o2,o3:setoid,f:unary_morphism o2 o3,g:unary_morphism o1 o2;
+unification hint 0 ≔ o1,o2,o3:setoid,f:o2 ⇒_0 o3,g:o1 ⇒_0 o2;
R ≟ mk_unary_morphism ?? (composition … f g)
(prop1 ?? (comp_unary_morphisms o1 o2 o3 f g))
(* -------------------------------------------------------------------- *) ⊢
ndefinition comp_binary_morphisms:
∀o1,o2,o3.
- unary_morphism (unary_morph_setoid o2 o3)
+ (unary_morph_setoid o2 o3) ⇒_0
(unary_morph_setoid (unary_morph_setoid o1 o2) (unary_morph_setoid o1 o3)).
#o1; #o2; #o3; napply mk_binary_morphism
[ #f; #g; napply (comp_unary_morphisms … f g) (*CSC: why not ∘?*)
include "sets/setoids.ma".
include "hints_declaration.ma".
-nrecord setoid1: Type[2] ≝
- { carr1:> Type[1];
- eq1: equivalence_relation1 carr1
- }.
+nrecord setoid1: Type[2] ≝ {
+ carr1:> Type[1];
+ eq1: equivalence_relation1 carr1
+}.
ndefinition setoid1_of_setoid: setoid → setoid1.
- #s; napply mk_setoid1
- [ napply (carr s)
- | napply (mk_equivalence_relation1 s)
- [ napply eq0
- | napply refl
- | napply sym
- | napply trans]##]
+ #s; @ (carr s); @ (eq0…) (refl…) (sym…) (trans…);
nqed.
-(*ncoercion setoid1_of_setoid : ∀s:setoid. setoid1 ≝ setoid1_of_setoid
- on _s: setoid to setoid1.*)
-(*prefer coercion Type_OF_setoid.*)
-
interpretation "setoid1 eq" 'eq t x y = (eq_rel1 ? (eq1 t) x y).
interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq0 t) x y).
interpretation "setoid1 symmetry" 'invert r = (sym1 ???? r).
interpretation "setoid symmetry" 'invert r = (sym ???? r).
-notation ".= r" with precedence 50 for @{'trans $r}.
notation ".=_1 r" with precedence 50 for @{'trans_x1 $r}.
interpretation "trans1" 'trans r = (trans1 ????? r).
interpretation "trans" 'trans r = (trans ????? r).
interpretation "trans1_x1" 'trans_x1 r = (trans1 ????? r).
-nrecord unary_morphism1 (A,B: setoid1) : Type[1] ≝
- { fun11:1> A → B;
- prop11: ∀a,a'. eq1 ? a a' → eq1 ? (fun11 a) (fun11 a')
- }.
+nrecord unary_morphism1 (A,B: setoid1) : Type[1] ≝ {
+ fun11:1> A → B;
+ prop11: ∀a,a'. eq1 ? a a' → eq1 ? (fun11 a) (fun11 a')
+}.
+
+notation > "B ⇒_1 C" right associative with precedence 72 for @{'umorph1 $B $C}.
+notation "hvbox(B break ⇒\sub 1 C)" right associative with precedence 72 for @{'umorph1 $B $C}.
+interpretation "unary morphism 1" 'umorph1 A B = (unary_morphism1 A B).
notation "┼_1 c" with precedence 89 for @{'prop1_x1 $c }.
interpretation "prop11" 'prop1 c = (prop11 ????? c).
interpretation "refl1" 'refl = (refl1 ???).
ndefinition unary_morphism1_setoid1: setoid1 → setoid1 → setoid1.
- #s; #s1; @ (unary_morphism1 s s1); @
+ #s; #s1; @ (s ⇒_1 s1); @
[ #f; #g; napply (∀a,a':s. a=a' → f a = g a')
| #x; #a; #a'; #Ha; napply (.= †Ha); napply refl1
| #x; #y; #H; #a; #a'; #Ha; napply (.= †Ha); napply sym1; /2/
nqed.
unification hint 0 ≔ S, T ;
- R ≟ (unary_morphism1_setoid1 S T)
+ R ≟ (unary_morphism1_setoid1 S T)
(* --------------------------------- *) ⊢
- carr1 R ≡ unary_morphism1 S T.
+ carr1 R ≡ S ⇒_1 T.
notation "l ╪_1 r" with precedence 89 for @{'prop2_x1 $l $r }.
interpretation "prop21" 'prop2 l r = (prop11 ? (unary_morphism1_setoid1 ??) ? ?? l ?? r).
interpretation "prop21_x1" 'prop2_x1 l r = (prop11 ? (unary_morphism1_setoid1 ??) ? ?? l ?? r).
-nlemma unary_morph1_eq1: ∀A,B.∀f,g: unary_morphism1 A B. (∀x. f x = g x) → f=g.
+nlemma unary_morph1_eq1: ∀A,B.∀f,g: A ⇒_1 B. (∀x. f x = g x) → f = g.
/3/. nqed.
nlemma mk_binary_morphism1:
∀A,B,C: setoid1. ∀f: A → B → C. (∀a,a',b,b'. a=a' → b=b' → f a b = f a' b') →
- unary_morphism1 A (unary_morphism1_setoid1 B C).
+ A ⇒_1 (unary_morphism1_setoid1 B C).
#A; #B; #C; #f; #H; @ [ #x; @ (f x) ] #a; #a'; #Ha [##2: napply unary_morph1_eq1; #y]
/2/.
nqed.
interpretation "function composition" 'compose f g = (composition ??? f g).
interpretation "function composition1" 'compose f g = (composition1 ??? f g).
-ndefinition comp1_unary_morphisms:
- ∀o1,o2,o3:setoid1.
- unary_morphism1 o2 o3 → unary_morphism1 o1 o2 →
- unary_morphism1 o1 o3.
+ndefinition comp1_unary_morphisms:
+ ∀o1,o2,o3:setoid1.o2 ⇒_1 o3 → o1 ⇒_1 o2 → o1 ⇒_1 o3.
#o1; #o2; #o3; #f; #g; @ (f ∘ g);
#a; #a'; #e; nnormalize; napply (.= †(†e)); napply #.
nqed.
-unification hint 0 ≔ o1,o2,o3:setoid1,f:unary_morphism1 o2 o3,g:unary_morphism1 o1 o2;
+unification hint 0 ≔ o1,o2,o3:setoid1,f:o2 ⇒_1 o3,g:o1 ⇒_1 o2;
R ≟ (mk_unary_morphism1 ?? (composition1 … f g)
(prop11 ?? (comp1_unary_morphisms o1 o2 o3 f g)))
(* -------------------------------------------------------------------- *) ⊢
ndefinition comp1_binary_morphisms:
∀o1,o2,o3.
- unary_morphism1 (unary_morphism1_setoid1 o2 o3)
- (unary_morphism1_setoid1 (unary_morphism1_setoid1 o1 o2) (unary_morphism1_setoid1 o1 o3)).
+ (unary_morphism1_setoid1 o2 o3) ⇒_1
+ (unary_morphism1_setoid1 (unary_morphism1_setoid1 o1 o2) (unary_morphism1_setoid1 o1 o3)).
#o1; #o2; #o3; napply mk_binary_morphism1
[ #f; #g; napply (comp1_unary_morphisms … f g) (*CSC: why not ∘?*)
| #a; #a'; #b; #b'; #ea; #eb; #x; #x'; #Hx; nnormalize; /3/ ]
include "properties/relations1.ma".
ndefinition seteq: ∀A. equivalence_relation1 (Ω^A).
- #A; @
- [ napply (λS,S'. S ⊆ S' ∧ S' ⊆ S)
- | /2/
- | #S; #S'; *; /3/
- | #S; #T; #U; *; #H1; #H2; *; /4/]
+#A; @(λS,S'. S ⊆ S' ∧ S' ⊆ S); /2/; ##[ #A B; *; /3/]
+#S T U; *; #H1 H2; *; /4/;
nqed.
include "sets/setoids1.ma".
include "hints_declaration.ma".
alias symbol "hint_decl" = "hint_decl_Type2".
-unification hint 0 ≔ A ⊢ carr1 (mk_setoid1 (Ω^A) (eq1 (powerclass_setoid A))) ≡ Ω^A.
+unification hint 0 ≔ A;
+ R ≟ (mk_setoid1 (Ω^A) (eq1 (powerclass_setoid A)))
+(*--------------------------------------------------*)⊢
+ carr1 R ≡ Ω^A.
(************ SETS OVER SETOIDS ********************)
include "logic/cprop.ma".
-nrecord ext_powerclass (A: setoid) : Type[1] ≝
- { ext_carr:> Ω^A; (* qui pc viene dichiarato con un target preciso...
- forse lo si vorrebbe dichiarato con un target più lasco
- ma la sintassi :> non lo supporta *)
+nrecord ext_powerclass (A: setoid) : Type[1] ≝ {
+ ext_carr:> Ω^A; (* qui pc viene dichiarato con un target preciso...
+ forse lo si vorrebbe dichiarato con un target più lasco
+ ma la sintassi :> non lo supporta *)
ext_prop: ∀x,x':A. x=x' → (x ∈ ext_carr) = (x' ∈ ext_carr)
- }.
+}.
notation > "𝛀 ^ term 90 A" non associative with precedence 70
for @{ 'ext_powerclass $A }.
(* ----------------------------------------------------- *) ⊢
carr1 R ≡ ext_powerclass A.
-(*
-interpretation "prop21 mem" 'prop2 l r = (prop21 (setoid1_of_setoid ?) (ext_powerclass_setoid ?) ? ? ???? l r).
-*)
-
-(*
-ncoercion ext_carr' : ∀A.∀x:ext_powerclass_setoid A. Ω^A ≝ ext_carr
-on _x : (carr1 (ext_powerclass_setoid ?)) to (Ω^?).
-*)
-
nlemma mem_ext_powerclass_setoid_is_morph:
- ∀A. unary_morphism1 (setoid1_of_setoid A) (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
+ ∀A. (setoid1_of_setoid A) ⇒_1 (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
#A; napply (mk_binary_morphism1 … (λx:setoid1_of_setoid A.λS: 𝛀^A. x ∈ S));
#a; #a'; #b; #b'; #Ha; *; #Hb1; #Hb2; @; #H
[ napply (. (ext_prop … Ha^-1)) | napply (. (ext_prop … Ha)) ] /2/.
unification hint 0 ≔ AA, x, S;
A ≟ carr AA,
SS ≟ (ext_carr ? S),
- TT ≟ (mk_unary_morphism1 …
+ TT ≟ (mk_unary_morphism1 ??
(λx:setoid1_of_setoid ?.
- mk_unary_morphism1 …
+ mk_unary_morphism1 ??
(λS:ext_powerclass_setoid ?. x ∈ S)
- (prop11 … (mem_ext_powerclass_setoid_is_morph AA x)))
- (prop11 … (mem_ext_powerclass_setoid_is_morph AA))),
+ (prop11 ?? (mem_ext_powerclass_setoid_is_morph AA x)))
+ (prop11 ?? (mem_ext_powerclass_setoid_is_morph AA))),
XX ≟ (ext_powerclass_setoid AA)
(*-------------------------------------*) ⊢
fun11 (setoid1_of_setoid AA)
(unary_morphism1_setoid1 XX CPROP) TT x S
≡ mem A SS x.
-nlemma subseteq_is_morph: ∀A. unary_morphism1 (ext_powerclass_setoid A)
- (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
+nlemma set_ext : ∀S.∀A,B:Ω^S.A =_1 B → ∀x:S.(x ∈ A) =_1 (x ∈ B).
+#S A B; *; #H1 H2 x; @; ##[ napply H1 | napply H2] nqed.
+
+nlemma ext_set : ∀S.∀A,B:Ω^S.(∀x:S. (x ∈ A) = (x ∈ B)) → A = B.
+#S A B H; @; #x; ncases (H x); #H1 H2; ##[ napply H1 | napply H2] nqed.
+
+nlemma subseteq_is_morph: ∀A.
+ (ext_powerclass_setoid A) ⇒_1
+ (unary_morphism1_setoid1 (ext_powerclass_setoid A) CPROP).
#A; napply (mk_binary_morphism1 … (λS,S':𝛀^A. S ⊆ S'));
#a; #a'; #b; #b'; *; #H1; #H2; *; /5/ by mk_iff, sym1, subseteq_trans;
nqed.
alias symbol "hint_decl" (instance 1) = "hint_decl_Type2".
-unification hint 0 ≔ A,a,a'
- (*-----------------------------------------------------------------*) ⊢
- eq_rel ? (eq0 A) a a' ≡ eq_rel1 ? (eq1 (setoid1_of_setoid A)) a a'.
+unification hint 0 ≔ A,x,y
+(*-----------------------------------------------*) ⊢
+ eq_rel ? (eq0 A) x y ≡ eq_rel1 ? (eq1 (setoid1_of_setoid A)) x y.
+
+(* XXX capire come mai questa hint non funziona se porto su (setoid1_of_setoid A) *)
nlemma intersect_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A.
- #A; #S; #S'; @ (S ∩ S');
- #a; #a'; #Ha; @; *; #H1; #H2; @
- [##1,2: napply (. Ha^-1‡#); nassumption;
-##|##3,4: napply (. Ha‡#); nassumption]
+#S A B; @ (A ∩ B); #x y Exy; @; *; #H1 H2; @;
+##[##1,2: napply (. Exy^-1‡#); nassumption;
+##|##3,4: napply (. Exy‡#); nassumption]
nqed.
alias symbol "hint_decl" = "hint_decl_Type1".
unification hint 0 ≔
A : setoid, B,C : ext_powerclass A;
- R ≟ (mk_ext_powerclass ? (B ∩ C) (ext_prop ? (intersect_is_ext ? B C)))
-
+ R ≟ (mk_ext_powerclass ? (B ∩ C) (ext_prop ? (intersect_is_ext ? B C)))
(* ------------------------------------------*) ⊢
ext_carr A R ≡ intersect ? (ext_carr ? B) (ext_carr ? C).
nlemma intersect_is_morph:
- ∀A. unary_morphism1 (powerclass_setoid A) (unary_morphism1_setoid1 (powerclass_setoid A) (powerclass_setoid A)).
+ ∀A. (powerclass_setoid A) ⇒_1 (unary_morphism1_setoid1 (powerclass_setoid A) (powerclass_setoid A)).
#A; napply (mk_binary_morphism1 … (λS,S'. S ∩ S'));
#a; #a'; #b; #b'; *; #Ha1; #Ha2; *; #Hb1; #Hb2; @; #x; nnormalize; *;/3/.
nqed.
alias symbol "hint_decl" = "hint_decl_Type1".
-unification hint 0 ≔
- A : Type[0], B,C : Ω^A;
- R ≟ (mk_unary_morphism1 …
- (λS. mk_unary_morphism1 … (λS'.S ∩ S') (prop11 … (intersect_is_morph A S)))
- (prop11 … (intersect_is_morph A)))
- ⊢
- R B C ≡ intersect ? B C.
+unification hint 0 ≔ A : Type[0], B,C : Ω^A;
+ R ≟ mk_unary_morphism1 ??
+ (λS. mk_unary_morphism1 ?? (λS'.S ∩ S') (prop11 ?? (intersect_is_morph A S)))
+ (prop11 ?? (intersect_is_morph A))
+(*------------------------------------------------------------------------*) ⊢
+ fun11 ?? (fun11 ?? R B) C ≡ intersect A B C.
interpretation "prop21 ext" 'prop2 l r =
(prop11 (ext_powerclass_setoid ?)
(unary_morphism1_setoid1 (ext_powerclass_setoid ?) ?) ? ?? l ?? r).
nlemma intersect_is_ext_morph:
- ∀A. unary_morphism1 (ext_powerclass_setoid A)
+ ∀A.
+ (ext_powerclass_setoid A) ⇒_1
(unary_morphism1_setoid1 (ext_powerclass_setoid A) (ext_powerclass_setoid A)).
#A; napply (mk_binary_morphism1 … (intersect_is_ext …));
#a; #a'; #b; #b'; #Ha; #Hb; napply (prop11 … (intersect_is_morph A)); nassumption.
unification hint 1 ≔
AA : setoid, B,C : 𝛀^AA;
A ≟ carr AA,
- R ≟ (mk_unary_morphism1 …
- (λS:ext_powerclass_setoid AA.
+ R ≟ (mk_unary_morphism1 ??
+ (λS:𝛀^AA.
mk_unary_morphism1 ??
- (λS':ext_powerclass_setoid AA.
+ (λS':𝛀^AA.
mk_ext_powerclass AA (S∩S') (ext_prop AA (intersect_is_ext ? S S')))
- (prop11 … (intersect_is_ext_morph AA S)))
- (prop11 … (intersect_is_ext_morph AA))) ,
+ (prop11 ?? (intersect_is_ext_morph AA S)))
+ (prop11 ?? (intersect_is_ext_morph AA))) ,
BB ≟ (ext_carr ? B),
CC ≟ (ext_carr ? C)
(* ------------------------------------------------------*) ⊢
ext_carr AA (R B C) ≡ intersect A BB CC.
+nlemma union_is_morph :
+ ∀A. (powerclass_setoid A) ⇒_1 (unary_morphism1_setoid1 (powerclass_setoid A) (powerclass_setoid A)).
+(*XXX ∀A.Ω^A ⇒_1 (Ω^A ⇒_1 Ω^A). avec non-unif-coerc*)
+#X; napply (mk_binary_morphism1 … (λA,B.A ∪ B));
+#A1 A2 B1 B2 EA EB; napply ext_set; #x;
+nchange in match (x ∈ (A1 ∪ B1)) with (?∨?);
+napply (.= (set_ext ??? EA x)‡#);
+napply (.= #‡(set_ext ??? EB x)); //;
+nqed.
+
+nlemma union_is_ext: ∀A. 𝛀^A → 𝛀^A → 𝛀^A.
+ #S A B; @ (A ∪ B); #x y Exy; @; *; #H1;
+##[##1,3: @; ##|##*: @2 ]
+##[##1,3: napply (. (Exy^-1)╪_1#)
+##|##2,4: napply (. Exy╪_1#)]
+nassumption;
+nqed.
+
+alias symbol "hint_decl" = "hint_decl_Type1".
+unification hint 0 ≔
+ A : setoid, B,C : 𝛀^A;
+ R ≟ (mk_ext_powerclass ? (B ∪ C) (ext_prop ? (union_is_ext ? B C)))
+(*-------------------------------------------------------------------------*) ⊢
+ ext_carr A R ≡ union ? (ext_carr ? B) (ext_carr ? C).
+
+unification hint 0 ≔ S:Type[0], A,B:Ω^S;
+ MM ≟ mk_unary_morphism1 ??
+ (λA.mk_unary_morphism1 ?? (λB.A ∪ B) (prop11 ?? (union_is_morph S A)))
+ (prop11 ?? (union_is_morph S))
+(*--------------------------------------------------------------------------*) ⊢
+ fun11 ?? (fun11 ?? MM A) B ≡ A ∪ B.
+
+nlemma union_is_ext_morph:∀A.
+ (ext_powerclass_setoid A) ⇒_1
+ (unary_morphism1_setoid1 (ext_powerclass_setoid A) (ext_powerclass_setoid A)).
+(*XXX ∀A:setoid.𝛀^A ⇒_1 (𝛀^A ⇒_1 𝛀^A). with coercion non uniformi *)
+#A; napply (mk_binary_morphism1 … (union_is_ext …));
+#x1 x2 y1 y2 Ex Ey; napply (prop11 … (union_is_morph A)); nassumption.
+nqed.
+
+unification hint 1 ≔
+ AA : setoid, B,C : 𝛀^AA;
+ A ≟ carr AA,
+ R ≟ (mk_unary_morphism1 ??
+ (λS:𝛀^AA.
+ mk_unary_morphism1 ??
+ (λS':𝛀^AA.
+ mk_ext_powerclass AA (S ∪ S') (ext_prop AA (union_is_ext ? S S')))
+ (prop11 ?? (union_is_ext_morph AA S)))
+ (prop11 ?? (union_is_ext_morph AA))) ,
+ BB ≟ (ext_carr ? B),
+ CC ≟ (ext_carr ? C)
+(*------------------------------------------------------*) ⊢
+ ext_carr AA (R B C) ≡ union A BB CC.
+
(*
alias symbol "hint_decl" = "hint_decl_Type2".
unification hint 0 ≔
(******************* first omomorphism theorem for sets **********************)
ndefinition eqrel_of_morphism:
- ∀A,B. unary_morphism A B → compatible_equivalence_relation A.
+ ∀A,B. A ⇒_0 B → compatible_equivalence_relation A.
#A; #B; #f; @
[ @ [ napply (λx,y. f x = f y) ] /2/;
##| #x; #x'; #H; nwhd; alias symbol "prop1" = "prop1".
napply (.= (†H)); // ]
nqed.
-ndefinition canonical_proj: ∀A,R. unary_morphism A (quotient A R).
+ndefinition canonical_proj: ∀A,R. A ⇒_0 (quotient A R).
#A; #R; @
[ napply (λx.x) | #a; #a'; #H; napply (compatibility … R … H) ]
nqed.
ndefinition quotiented_mor:
- ∀A,B.∀f:unary_morphism A B.
- unary_morphism (quotient … (eqrel_of_morphism … f)) B.
+ ∀A,B.∀f:A ⇒_0 B.(quotient … (eqrel_of_morphism … f)) ⇒_0 B.
#A; #B; #f; @ [ napply f ] //.
nqed.
alias symbol "eq" = "setoid eq".
ndefinition surjective ≝
- λA,B.λS: ext_powerclass A.λT: ext_powerclass B.λf:unary_morphism A B.
+ λA,B.λS: ext_powerclass A.λT: ext_powerclass B.λf:A ⇒_0 B.
∀y. y ∈ T → ∃x. x ∈ S ∧ f x = y.
ndefinition injective ≝
- λA,B.λS: ext_powerclass A.λf:unary_morphism A B.
+ λA,B.λS: ext_powerclass A.λf:A ⇒_0 B.
∀x,x'. x ∈ S → x' ∈ S → f x = f x' → x = x'.
nlemma first_omomorphism_theorem_functions2:
- ∀A,B.∀f: unary_morphism A B.
+ ∀A,B.∀f:A ⇒_0 B.
surjective … (Full_set ?) (Full_set ?) (canonical_proj ? (eqrel_of_morphism … f)).
/3/. nqed.
nlemma first_omomorphism_theorem_functions3:
- ∀A,B.∀f: unary_morphism A B.
+ ∀A,B.∀f:A ⇒_0 B.
injective … (Full_set ?) (quotiented_mor … f).
#A; #B; #f; nwhd; #x; #x'; #Hx; #Hx'; #K; nassumption.
nqed.
nrecord isomorphism (A, B : setoid) (S: ext_powerclass A) (T: ext_powerclass B) : Type[0] ≝
- { iso_f:> unary_morphism A B;
+ { iso_f:> A ⇒_0 B;
f_closed: ∀x. x ∈ S → iso_f x ∈ T;
f_sur: surjective … S T iso_f;
f_inj: injective … S iso_f