apply ((id_neutral_left1 ????)‡#);]
qed.
-definition BPext: ∀o: BP. form o ⇒ Ω \sup (concr o) ≝ λo.ext ? ? (rel o).
+definition BPext: ∀o: BP. form o ⇒ Ω \sup (concr o).
+ intros; constructor 1;
+ [ apply (ext ? ? (rel o));
+ | intros;
+ apply (.= #‡H);
+ apply refl1]
+qed.
definition BPextS: ∀o: BP. Ω \sup (form o) ⇒ Ω \sup (concr o) ≝
λo.extS ?? (rel o).
definition cont_rel': ∀S,T: basic_topology. continuous_relation_setoid S T → arrows1 ? S T ≝ cont_rel.
coercion cont_rel'.
+
(*
definition BTop: category1.
constructor 1;
apply (.= (saturated ?????)\sup -1);
[ apply (.= (saturated ?????)); [ assumption | apply refl1 ]
| apply (.= (minus_star_image_comp ??????)\sup -1);
- apply refl1]]]
+ apply refl1]]
+ | intros; simplify; intro; simplify; whd in H H1;
+ apply (.= †(ext_comp ???));
+ ]
| intros; simplify; intro; simplify;
+ apply (.= †(ASSOC1‡#));
+ apply refl1
| intros; simplify; intro; simplify;
+ apply (.= †((id_neutral_right1 ????)‡#));
+ apply refl1
| intros; simplify; intro; simplify;
- ]
-qed.
-*)
\ No newline at end of file
+ apply (.= †((id_neutral_left1 ????)‡#));
+ apply refl1]
+qed.*)
\ No newline at end of file
interpretation "subset comprehension" 'comprehension s p =
(comprehension s (mk_unary_morphism __ p _)).
-definition ext: ∀X,S:REL. ∀r: arrows1 ? X S. S ⇒ Ω \sup X.
- apply (λX,S,r.mk_unary_morphism ?? (λf.{x ∈ X | x ♮r f}) ?);
+definition ext: ∀X,S:REL. binary_morphism1 (arrows1 ? X S) S (Ω \sup X).
+ apply (λX,S.mk_binary_morphism1 ??? (λr:arrows1 ? X S.λf:S.{x ∈ X | x ♮r f}) ?);
[ intros; simplify; apply (.= (H‡#)); apply refl1
- | intros; simplify; split; intros; simplify; intros;
- [ apply (. #‡(#‡H)); assumption
- | apply (. #‡(#‡H\sup -1)); assumption]]
+ | intros; simplify; split; intros; simplify; intros; cases f; split; try assumption;
+ [ apply (. (#‡H1)); whd in H; apply (if ?? (H ??)); assumption
+ | apply (. (#‡H1\sup -1)); whd in H; apply (fi ?? (H ??));assumption]]
qed.
definition extS: ∀X,S:REL. ∀r: arrows1 ? X S. Ω \sup S ⇒ Ω \sup X.