apply (respects_comp2 ?? SUBSETS' (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; unfold BP_to_OBP in e; simplify in e; cases e;
- unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
- intros 2; change in match or_f_ in e3 with (λq,w,x.fun12 ?? (or_f q w) x);
- simplify in e3; STOP lapply (e3 (singleton ? x)); cases Hletin;
- 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;]
+ 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 sym2;
+ apply (.= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) g \sub \c (⊩ \sub T));
+ apply sym2;
+ apply e;
qed.
-*)
-(*
-theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f).
- intros; exists;
-
-*)
\ No newline at end of file
+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);
+ 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 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)]
+ 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);
+ | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
+ | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
+ simplify; apply (†(Hgc‡#));
+qed.