@@ -117,12 +117,12 @@ Require Import reals signed.
117117(* predicates on natural numbers that are *)
118118(* eventually true. *)
119119(* separates_points_from_closed f == For a closed set U and point x outside *)
120- (* some member of the family `f` sends *)
121- (* f_i(x) outside (closure (f_i@` U)). *)
122- (* used together with ` join_product` *)
123- (* join_product f == The function (x i => f i x). When the *)
120+ (* some member of the family f sends *)
121+ (* f_i(x) outside (closure (f_i @` U)). *)
122+ (* Used together with join_product. *)
123+ (* join_product f == The function (x => f ^~ x). When the *)
124124(* family f separates points from closed *)
125- (* sets, join_product is an embedding *)
125+ (* sets, join_product is an embedding. *)
126126(* *)
127127(* * Near notations and tactics: *)
128128(* --> The purpose of the near notations and tactics is to make the *)
@@ -3326,7 +3326,7 @@ Qed.
33263326Definition hausdorff_accessible : hausdorff_space T -> accessible_space.
33273327Proof .
33283328rewrite open_hausdorff => hsdfT => x y /hsdfT [[U V] [xU yV]] [/= ? ? /eqP].
3329- rewrite setIC => /disjoints_subset VUc; exists U; repeat split => //.
3329+ rewrite setIC => /disjoints_subset VUc; exists U; repeat split => //.
33303330by rewrite inE; apply: VUc; rewrite -inE.
33313331Qed .
33323332
@@ -6678,21 +6678,20 @@ Qed.
66786678End SubspaceWeak.
66796679
66806680Definition separates_points_from_closed {I : Type } {T : topologicalType}
6681- {U_ : I -> topologicalType} (f_ : forall i, ( T -> U_ i) ) :=
6681+ {U_ : I -> topologicalType} (f_ : forall i, T -> U_ i) :=
66826682 forall (U : set T) x,
66836683 closed U -> ~ U x -> exists i, ~ (closure (f_ i @` U)) (f_ i x).
66846684
6685- (* A handy technique for emebdding a space T into a product. The key interface
6686- is 'separates_points_from_closed', which guarantees that the topologies
6685+ (* A handy technique for embedding a space T into a product. The key interface
6686+ is 'separates_points_from_closed', which guarantees that the topologies
66876687 - T's native topology
66886688 - sup (weak f_i) - the sup of all the weak topologies of f_i
66896689 - weak (x => (f_1 x, f_2 x,...)) - the weak topology from the product space
66906690 are equivalent (the last equivalence seems to require accessible_space).
66916691 *)
66926692Section product_embeddings.
6693- Context {I : choiceType} {T : topologicalType}.
6694- Context {U_ : I -> topologicalType}.
6695- Variable (f_ : forall i, (T -> U_ i)).
6693+ Context {I : choiceType} {T : topologicalType} {U_ : I -> topologicalType}.
6694+ Variable (f_ : forall i, T -> U_ i).
66966695
66976696Hypothesis sepf : separates_points_from_closed f_.
66986697Hypothesis ctsf : forall i, continuous (f_ i).
@@ -6712,7 +6711,7 @@ move=> FF; split.
67126711move/cvg_sup => wiFx U; rewrite /= nbhs_simpl nbhsE => [[B [[oB ?]]]].
67136712move/filterS; apply; have [//|i nclfix] := @sepf _ x (open_closedC oB).
67146713apply: (wiFx i); have /= -> := @nbhsE (weak_topologicalType (f_ i)) x.
6715- exists (f_ i @^-1` (~` closure [set f_ i x | x in ~` B])); repeat split => // .
6714+ exists (f_ i @^-1` (~` closure [set f_ i x | x in ~` B])); split; [split=>//|] .
67166715 apply: open_comp; last by rewrite ?openC; last apply: closed_closure.
67176716 by move=> + _; exact: weak_continuous.
67186717rewrite closureC preimage_bigcup => z [V [oV]] VnB => /VnB.
@@ -6732,7 +6731,7 @@ rewrite predeqE => A; rewrite ?openE /interior.
67326731by split => + z => /(_ z); rewrite weak_sep_nbhsE.
67336732Qed .
67346733
6735- Definition join_product (x : T) : PU := fun i => f_ i x.
6734+ Definition join_product (x : T) : PU := f_ ^~ x.
67366735
67376736Lemma join_product_continuous : continuous join_product.
67386737Proof .
@@ -6746,21 +6745,22 @@ apply: open_comp => // + _; rewrite /cvg_to => x U.
67466745by rewrite nbhs_simpl /= -weak_sep_nbhsE; move: x U; exact: ctsf.
67476746Qed .
67486747
6749- Lemma join_product_open : forall (A : set T), open A ->
6750- open ((join_product @` A) : set (subspace (join_product @` setT ))).
6748+ Lemma join_product_open (A : set T) : open A ->
6749+ open ((join_product @` A) : set (subspace (range join_product ))).
67516750Proof .
6752- move=> A oA; rewrite openE => y /= [x Ax] jxy.
6751+ move=> oA; rewrite openE => y /= [x Ax] jxy.
67536752have [// | i nAfiy] := @sepf (~` A) x (open_closedC oA).
67546753pose B := prod_topo_apply i @^-1` (~` closure (f_ i @` ~` A)).
6755- apply: (@filterS _ _ _ (( join_product @` setT) ` &` B)).
6754+ apply: (@filterS _ _ _ (range join_product ` &` B)).
67566755 move=> z [[w ?]] wzE Bz; exists w => //.
67576756 move: Bz; rewrite /B -wzE closureC; case=> K [oK KsubA] /KsubA.
67586757 have -> : prod_topo_apply i (join_product w) = f_ i w by [].
67596758 by move=> /exists2P/forallNP/(_ w)/not_andP [] // /contrapT.
67606759apply: open_nbhs_nbhs; split; last by rewrite -jxy.
6761- apply: openI; first exact: open_subspaceT; apply: open_subspaceW.
6762- apply: open_comp; first by move=> + _; exact: prod_topo_apply_continuous.
6763- by exact/closed_openC/closed_closure.
6760+ apply: openI; first exact: open_subspaceT.
6761+ apply: open_subspaceW; apply: open_comp.
6762+ by move=> + _; exact: prod_topo_apply_continuous.
6763+ exact/closed_openC/closed_closure.
67646764Qed .
67656765
67666766Lemma join_product_inj : accessible_space T -> set_inj [set: T] join_product.
0 commit comments