definition o_basic_pair_of_basic_pair: basic_pair → Obasic_pair.
intro b;
constructor 1;
- [ apply (map_objs2 ?? SUBSETS' (concr b));
- | apply (map_objs2 ?? SUBSETS' (form b));
- | apply (map_arrows2 ?? SUBSETS' (concr b) (form b) (rel b)); ]
+ [ apply (map_objs2 ?? POW (concr b));
+ | apply (map_objs2 ?? POW (form b));
+ | apply (map_arrows2 ?? POW (concr b) (form b) (rel b)); ]
qed.
definition o_relation_pair_of_relation_pair:
Orelation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2).
intros;
constructor 1;
- [ apply (map_arrows2 ?? SUBSETS' (concr BP1) (concr BP2) (r \sub \c));
- | apply (map_arrows2 ?? SUBSETS' (form BP1) (form BP2) (r \sub \f));
- | apply (.= (respects_comp2 ?? SUBSETS' (concr BP1) (concr BP2) (form BP2) r\sub\c (⊩\sub BP2) )^-1);
- cut (⊩ \sub BP2∘r \sub \c = r\sub\f ∘ ⊩ \sub BP1) as H;
+ [ apply (map_arrows2 ?? POW (concr BP1) (concr BP2) (r \sub \c));
+ | apply (map_arrows2 ?? POW (form BP1) (form BP2) (r \sub \f));
+ | apply (.= (respects_comp2 ?? POW (concr BP1) (concr BP2) (form BP2) r\sub\c (⊩\sub BP2) )^-1);
+ cut ( ⊩ \sub BP2 ∘ r \sub \c =_12 r\sub\f ∘ ⊩ \sub BP1) as H;
[ apply (.= †H);
- apply (respects_comp2 ?? SUBSETS' (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f);
+ apply (respects_comp2 ?? POW (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f);
| apply commute;]]
qed.
| change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
simplify;
apply (prop12);
- apply (.= (respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1);
+ apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1);
apply sym2;
- apply (.= (respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) (b\sub\c) (⊩\sub T))^-1);
+ apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (b\sub\c) (⊩\sub T))^-1);
apply sym2;
apply prop12;
apply Eab;
simplify;
apply prop12;
apply prop22;[2,4,6,8: apply rule #;]
- apply (respects_id2 ?? SUBSETS' (concr o));
+ apply (respects_id2 ?? POW (concr o));
| simplify; intros; whd; split;
[ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
| change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
simplify;
apply prop12;
apply prop22;[2,4,6,8: apply rule #;]
- apply (respects_comp2 ?? SUBSETS' (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);]
+ apply (respects_comp2 ?? POW (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);]
qed.
theorem BP_to_OBP_faithful:
∀S,T.∀f,g:arrows2 (category2_of_category1 BP) S T.
map_arrows2 ?? BP_to_OBP ?? f = map_arrows2 ?? BP_to_OBP ?? g → f=g.
intros; change with ( (⊩) ∘ f \sub \c = (⊩) ∘ g \sub \c);
- apply (SUBSETS_faithful);
- apply (.= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) f \sub \c (⊩ \sub T));
+ apply (POW_faithful);
+ apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) f \sub \c (⊩ \sub T));
apply sym2;
- apply (.= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) g \sub \c (⊩ \sub T));
+ apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) g \sub \c (⊩ \sub T));
apply sym2;
apply e;
qed.
theorem BP_to_OBP_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f).
intros;
- cases (SUBSETS_full (concr S) (concr T) (Oconcr_rel ?? f)) (gc Hgc);
- cases (SUBSETS_full (form S) (form T) (Oform_rel ?? f)) (gf Hgf);
+ cases (POW_full (concr S) (concr T) (Oconcr_rel ?? f)) (gc Hgc);
+ cases (POW_full (form S) (form T) (Oform_rel ?? f)) (gf Hgf);
exists[
constructor 1; [apply gc|apply gf]
- apply (SUBSETS_faithful);
- apply (let xxxx ≝SUBSETS' in .= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) gc (rel T));
+ apply (POW_faithful);
+ apply (let xxxx ≝POW in .= respects_comp2 ?? POW (concr S) (concr T) (form T) gc (rel T));
apply rule (.= Hgc‡#);
apply (.= Ocommute ?? f);
apply (.= #‡Hgf^-1);
- apply (let xxxx ≝SUBSETS' in (respects_comp2 ?? SUBSETS' (concr S) (form S) (form T) (rel S) gf)^-1)]
+ apply (let xxxx ≝POW in (respects_comp2 ?? POW (concr S) (form S) (form T) (rel S) gf)^-1)]
split;
[ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
| change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
interpretation "setoid2 eq" 'eq x y = (eq_rel2 _ (eq2 _) x y).
interpretation "setoid1 eq" 'eq x y = (eq_rel1 _ (eq1 _) x y).
interpretation "setoid eq" 'eq x y = (eq_rel _ (eq _) x y).
+
+notation < "hvbox(a break = \sub 1 b)" non associative with precedence 45
+for @{ 'eq1 $a $b }.
+
+notation < "hvbox(a break = \sub 2 b)" non associative with precedence 45
+for @{ 'eq2 $a $b }.
+
+notation < "hvbox(a break = \sub 3 b)" non associative with precedence 45
+for @{ 'eq3 $a $b }.
+
+notation > "hvbox(a break =_12 b)" non associative with precedence 45
+for @{ eq_rel2 (carr2 (setoid2_of_setoid1 ?)) (eq2 (setoid2_of_setoid1 ?)) $a $b }.
+notation > "hvbox(a break =_0 b)" non associative with precedence 45
+for @{ eq_rel ? (eq ?) $a $b }.
+notation > "hvbox(a break =_1 b)" non associative with precedence 45
+for @{ eq_rel1 ? (eq1 ?) $a $b }.
+notation > "hvbox(a break =_2 b)" non associative with precedence 45
+for @{ eq_rel2 ? (eq2 ?) $a $b }.
+notation > "hvbox(a break =_3 b)" non associative with precedence 45
+for @{ eq_rel3 ? (eq3 ?) $a $b }.
+
+interpretation "setoid3 eq explicit" 'eq3 x y = (eq_rel3 _ (eq3 _) x y).
+interpretation "setoid2 eq explicit" 'eq2 x y = (eq_rel2 _ (eq2 _) x y).
+interpretation "setoid1 eq explicit" 'eq1 x y = (eq_rel1 _ (eq1 _) x y).
+
interpretation "setoid3 symmetry" 'invert r = (sym3 ____ r).
interpretation "setoid2 symmetry" 'invert r = (sym2 ____ r).
interpretation "setoid1 symmetry" 'invert r = (sym1 ____ r).
oa_overlap_sym: ∀a,b:oa_P.oa_overlap a b → oa_overlap b a;
oa_meet_inf:
∀I:SET.∀p_i:I ⇒ oa_P.∀p:oa_P.
- oa_leq p (oa_meet I p_i) = ∀i:I.oa_leq p (p_i i);
- oa_join_sup: ∀I:SET.∀p_i:I ⇒ oa_P.∀p:oa_P.oa_leq (oa_join I p_i) p = ∀i:I.oa_leq (p_i i) p;
+ oa_leq p (oa_meet I p_i) = (∀i:I.oa_leq p (p_i i));
+ oa_join_sup: ∀I:SET.∀p_i:I ⇒ oa_P.∀p:oa_P.oa_leq (oa_join I p_i) p = (∀i:I.oa_leq (p_i i) p);
oa_zero_bot: ∀p:oa_P.oa_leq oa_zero p;
oa_one_top: ∀p:oa_P.oa_leq p oa_one;
oa_overlap_preserves_meet_:
(oa_meet ? { x ∈ BOOL | match x with [ true ⇒ p | false ⇒ q ] | IF_THEN_ELSE_p oa_P p q });
oa_join_split:
∀I:SET.∀p.∀q:I ⇒ oa_P.
- oa_overlap p (oa_join I q) = ∃i:I.oa_overlap p (q i);
+ oa_overlap p (oa_join I q) = (∃i:I.oa_overlap p (q i));
(*oa_base : setoid;
1) enum non e' il nome giusto perche' non e' suriettiva
2) manca (vedere altro capitolo) la "suriettivita'" come immagine di insiemi di oa_base
| constructor 1;
(* tenere solo una uguaglianza e usare la proposizione 9.9 per
le altre (unicita' degli aggiunti e del simmetrico) *)
- [ apply (λp,q. And42 (eq2 ? (or_f_minus_star_ ?? p) (or_f_minus_star_ ?? q))
- (eq2 ? (or_f_minus_ ?? p) (or_f_minus_ ?? q))
- (eq2 ? (or_f_ ?? p) (or_f_ ?? q))
- (eq2 ? (or_f_star_ ?? p) (or_f_star_ ?? q)));
+ [ apply (λp,q. And42
+ (or_f_minus_star_ ?? p = or_f_minus_star_ ?? q)
+ (or_f_minus_ ?? p = or_f_minus_ ?? q)
+ (or_f_ ?? p = or_f_ ?? q)
+ (or_f_star_ ?? p = or_f_star_ ?? q));
| whd; simplify; intros; repeat split; intros; apply refl2;
| whd; simplify; intros; cases a; clear a; split;
intro a; apply sym1; generalize in match a;assumption;
}.
(* FIX *)
-interpretation "basic pair relation indexed" 'Vdash2 x y c = (Orel c x y).
-interpretation "basic pair relation (non applied)" 'Vdash c = (Orel c).
+interpretation "o-basic pair relation indexed" 'Vdash2 x y c = (Orel c x y).
+interpretation "o-basic pair relation (non applied)" 'Vdash c = (Orel c).
alias symbol "eq" = "setoid1 eq".
alias symbol "compose" = "category1 composition".
}.
(* FIX *)
-interpretation "concrete relation" 'concr_rel r = (Oconcr_rel __ r).
-interpretation "formal relation" 'form_rel r = (Oform_rel __ r).
+interpretation "o-concrete relation" 'concr_rel r = (Oconcr_rel __ r).
+interpretation "o-formal relation" 'form_rel r = (Oform_rel __ r).
definition Orelation_pair_equality:
∀o1,o2. equivalence_relation2 (Orelation_pair o1 o2).
include "relations.ma".
include "o-algebra.ma".
-definition SUBSETS: objs1 SET → OAlgebra.
+definition POW': objs1 SET → OAlgebra.
intro A; constructor 1;
[ apply (Ω \sup A);
| apply subseteq;
| intros; split; intro;
[ cases f; cases x1; exists [apply w1] exists [apply w] assumption;
| cases e; cases x; exists; [apply w1] [ assumption | exists; [apply w] assumption]]
- | intros; intros 2; cases (f (singleton ? a) ?);
+ | intros; intros 2; cases (f {(a)} ?);
[ exists; [apply a] [assumption | change with (a = a); apply refl1;]
| change in x1 with (a = w); change with (mem A a q); apply (. (x1‡#));
assumption]]
qed.
-definition powerset_of_SUBSETS: ∀A.oa_P (SUBSETS A) → Ω \sup A ≝ λA,x.x.
-coercion powerset_of_SUBSETS.
+definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω \sup A ≝ λA,x.x.
+coercion powerset_of_POW'.
-definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2).
+definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (POW' o1) (POW' o2).
intros;
constructor 1;
[ constructor 1;
qed.
lemma orelation_of_relation_preserves_identity:
- ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (SUBSETS o1)).
+ ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (POW' o1)).
intros; split; intro; split; whd; intro;
[ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros;
apply (f a1); change with (a1 = a1); apply refl1;
lemma orelation_of_relation_preserves_composition:
∀o1,o2,o3:REL.∀F: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3.
orelation_of_relation ?? (G ∘ F) =
- comp2 OA (SUBSETS o1) (SUBSETS o2) (SUBSETS o3)
+ comp2 OA (POW' o1) (POW' o2) (POW' o3)
?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*).
[ apply (orelation_of_relation ?? F); | apply (orelation_of_relation ?? G); ]
intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
| cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
qed.
-definition SUBSETS': carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
+definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
constructor 1;
- [ apply SUBSETS;
+ [ apply POW';
| intros; constructor 1;
[ apply (orelation_of_relation S T);
| intros; apply (orelation_of_relation_preserves_equality S T a a' e); ]
| apply orelation_of_relation_preserves_composition; ]
qed.
-theorem SUBSETS_faithful:
+theorem POW_faithful:
∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
- map_arrows2 ?? SUBSETS' ?? f = map_arrows2 ?? SUBSETS' ?? g → f=g.
- intros; unfold SUBSETS' in e; simplify in e; cases e;
+ map_arrows2 ?? POW ?? f = map_arrows2 ?? POW ?? g → f=g.
+ intros; unfold POW in e; simplify in e; cases e;
unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
- intros 2; lapply (e3 (singleton ? x)); cases Hletin;
+ intros 2; cases (e3 {(x)});
split; intro; [ lapply (s y); | lapply (s1 y); ]
[2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
- |*: cases Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;]
+ |*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;]
qed.
+interpretation "lifting singl" 'singl x =
+ (fun11 _ (objs2 (POW _)) (singleton _) x).
-theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? SUBSETS' S T g = f).
+theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
intros; exists;
[ constructor 1; constructor 1;
- [ apply (λx:carr S.λy:carr T. y ∈ f (singleton S x));
+ [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
| intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
[4: apply mem; |6: apply Hletin;|1,2,3,5: skip]
lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]
(* qua non riesco a mettere set *)
definition singleton: ∀A:setoid. unary_morphism1 A (Ω \sup A).
intros; constructor 1;
- [ apply (λa:A.{b | eq ? a b}); unfold setoid1_of_setoid; simplify;
+ [ apply (λa:A.{b | a =_0 b}); unfold setoid1_of_setoid; simplify;
intros; simplify;
split; intro;
apply (.= e1);