FP: map_objs2 ?? F F1 =_\ID F2
}.
+notation "ℱ\sub 1 x" non associative with precedence 60 for @{'F1 $x}.
+notation > "ℱ_1" non associative with precedence 90 for @{F1 ???}.
+interpretation "F1" 'F1 x = (F1 ___ x).
+
+notation "ℱ\sub 2 x" non associative with precedence 60 for @{'F2 $x}.
+notation > "ℱ_2" non associative with precedence 90 for @{F2 ???}.
+interpretation "F2" 'F2 x = (F2 ___ x).
+
lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F.
- arrows2 C2 (F (F1 ??? X)) (F (F1 ??? Y)) →
- arrows2 C2 (F2 ??? X) (F2 ??? Y).
+ arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) →
+ arrows2 C2 (ℱ_2 X) (ℱ_2 Y).
intros 5; cases X; cases Y; clear X Y;
cases H; cases H1; intros; assumption;
qed.
FmP: REW ?? F X Y (map_arrows2 ?? F ?? Fm1) = Fm2
}.
+notation "ℳ\sub 1 x" non associative with precedence 60 for @{'Fm1 $x}.
+notation > "ℳ_1" non associative with precedence 90 for @{Fm1 ?????}.
+interpretation "Fm1" 'Fm1 x = (Fm1 _____ x).
+
+notation "ℳ\sub 2 x" non associative with precedence 60 for @{'Fm2 $x}.
+notation > "ℳ_2" non associative with precedence 90 for @{Fm2 ?????}.
+interpretation "Fm2" 'Fm2 x = (Fm2 _____ x).
+
definition Fm :
∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
Fo ?? F → Fo ?? F → setoid2.
binary_morphism2 (Fm ?? F o1 o2) (Fm ?? F o2 o3) (Fm ?? F o1 o3).
intros; constructor 1;
[ intros (f g); constructor 1;
- [ apply (comp2 C2 ??? (Fm2 ????? f) (Fm2 ????? g));
- | apply (comp2 C1 ??? (Fm1 ????? f) (Fm1 ????? g));
+ [ apply (comp2 C2 ??? (ℳ_2 f) (ℳ_2 g));
+ | apply (comp2 C1 ??? (ℳ_1 f) (ℳ_1 g));
| apply hide; cases o1 in f; cases o2 in g; cases o3; clear o1 o2 o3;
cases H; cases H1; cases H2; intros 2; cases c; cases c1; clear c c1;
simplify; apply (.= (respects_comp2:?)); apply (e1‡e);]
-| intros 6; change with
- ((Fm2 C1 C2 F o2 o3 b∘Fm2 C1 C2 F o1 o2 a) =
- (Fm2 C1 C2 F o2 o3 b'∘Fm2 C1 C2 F o1 o2 a'));
- change in e1 with (Fm2 ?? F ?? b = Fm2 ?? F ?? b');
- change in e with (Fm2 ?? F ?? a = Fm2 ?? F ?? a');
+| intros 6; change with ((ℳ_2 b ∘ ℳ_2 a) = (ℳ_2 b' ∘ ℳ_2 a'));
+ change in e1 with (ℳ_2 b = ℳ_2 b');
+ change in e with (ℳ_2 a = ℳ_2 a');
apply (e‡e1);]
qed.
| apply (Fm ?? F);
| apply F_id;
| apply F_comp;
-| intros; apply (comp_assoc2 C2 ???? (Fm2 ????? a12) (Fm2 ????? a23) (Fm2 ????? a34));
-| intros; apply (id_neutral_right2 C2 ?? (Fm2 ????? a));
-| intros; apply (id_neutral_left2 C2 ?? (Fm2 ????? a));]
+| intros; apply (comp_assoc2 C2 ???? (ℳ_2 a12) (ℳ_2 a23) (ℳ_2 a34));
+| intros; apply (id_neutral_right2 C2 ?? (ℳ_2 a));
+| intros; apply (id_neutral_left2 C2 ?? (ℳ_2 a));]
qed.
include "o-basic_pairs_to_o-basic_topologies.ma".
lemma rOR_full :
- ∀rS,rT:rOBP.∀f:arrows2 ? (OR (F2 ??? rS)) (OR (F2 ??? rT)).
- exT22 ? (λg:arrows2 OBP (F1 ??? rS) (F1 ??? rT).
- map_arrows2 ? ? OR rS rT g = f).
- ∀S,T.∀f. exT22 ? (λg. map_arrows2 ? ? OR S T g = f).
-arrows2 OBP S T
-unary_morphism1_setoid1 (OR S) (OR T)
+ ∀rS,rT:rOBP.∀f:arrows2 BTop (OR (ℱ_2 rS)) (OR (ℱ_2 rT)).
+ exT22 ? (λg:arrows2 rOBP rS rT.
+ map_arrows2 ?? OR ?? (ℳ_2 g) = f).
+intros; cases f (c H1 H2); simplify;
+STOP;
+
(* Todo: rename BTop → OBTop *)