@@ -6385,14 +6385,14 @@ Unshelve. all: by end_near. Qed.
63856385
63866386(* It turns out {family compact, U -> V} can be generalized to only assume *)
63876387(* `topologicalType` on V. This topology is called the compact-open topology. *)
6388- (* This topology is special becaise its the _only_ topology that will allow *)
6388+ (* This topology is special becaise it _only_ topology that will allow *)
63896389(* curry/uncurry to be continuous. *)
63906390
63916391Section compact_open.
63926392
63936393Context {T U : topologicalType}.
63946394
6395- Definition compact_open := ( T->U)^o .
6395+ Definition compact_open : Type := T->U.
63966396
63976397Section compact_open_setwise.
63986398Context {K : set T}.
@@ -6461,7 +6461,7 @@ Qed.
64616461Lemma compact_open_open (K : set T) (O : set U) :
64626462 compact K -> open O -> open ([set g | g @` K `<=` O] : set compact_open).
64636463Proof .
6464- pose C := [set g | g @` K `<=` O]; move=> cptK oO.
6464+ pose C := [set g | g @` K `<=` O]; move=> cptK oO.
64656465exists [set C]; last by rewrite bigcup_set1.
64666466move=> ? ->; exists (fset1 C) => //; last by rewrite set_fset1 bigcap_set1.
64676467by move=> ?; rewrite ?inE => ->; exists (existT _ _ cptK) => // z Cz; exists O.
@@ -6491,43 +6491,41 @@ Lemma compact_open_fam_compactP (f : U -> V) (F : set (set (U -> V))) :
64916491 continuous f -> (Filter F) ->
64926492 {compact-open, F --> f} <-> {family compact, F --> f}.
64936493Proof .
6494- move=> ctsf FF; split.
6495- move/compact_open_cvgP=> cptOF; apply/cvg_sup; case=> K cptK R.
6496- case=> D /= [[E oE <- /=] Ekf] /filterS; apply.
6497- move: oE; rewrite openE => /(_ _ Ekf); case => ? [J entJ /=] EKR KfE.
6498- near=> z; apply/KfE/EKR; case => u /= Kp.
6499- rewrite /= /set_val /= /eqincl /incl/set_val /= /comp.
6500- (have Ku : K u by rewrite inE in Kp); move: u Ku {D Kp}; near: z.
6501- move/compact_near_coveringP/near_covering_withinP:(cptK); apply.
6502- move=> u Ku; near (powerset_filter_from (@entourage V)) => E'.
6503- have entE' : entourage E' by exact: (near (near_small_set _)).
6504- pose C := f@^-1`(to_set E' (f u))%classic.
6505- pose B := \bigcup_(z in K `&` closure C) interior (to_set E' (f z)).
6506- have oB : open B by apply: bigcup_open => ? ?; exact: open_interior.
6507- have fKB : f @` (K `&` closure C) `<=` B.
6508- move=> ? [z ? <-]; exists z => //; rewrite /interior.
6509- by rewrite -nbhs_entourageE; exists E'.
6510- have cptKC : compact (K `&` closure C).
6511- apply: compact_closedI => //; exact: closed_closure.
6512- have := cptOF (K `&` closure C) B cptKC oB fKB.
6513- exists (C, [set g | [set g x | x in K `&` closure C] `<=` B]).
6514- split.
6515- - by apply: (ctsf) => //; rewrite /filter_of -nbhs_entourageE; exists E'.
6516- - exact: cptOF.
6517- - case=> z h /= [Cz KB Kz].
6518- case: (KB (h z)); first by (exists z; split => //; apply: subset_closure).
6519- move=> w [Kw Cw /interior_subset Jfwhz]; apply:subset_split_ent => //.
6520- exists (f w); last apply:(near (small_ent_sub _) E') => //.
6521- apply: subset_split_ent => //; exists (f u).
6522- apply entourage_sym; apply:(near (small_ent_sub _) E') => //.
6523- have [] := Cw (f@^-1` (to_set E' (f w))).
6524- by apply: ctsf; rewrite /= -nbhs_entourageE; exists E'.
6525- move=> r [Cr /= Ewr]; apply: subset_split_ent => //; exists (f r).
6526- by apply:(near (small_ent_sub _) E') => //.
6527- by apply entourage_sym; apply:(near (small_ent_sub _) E') => //.
6528- move=> cptF; apply/compact_open_cvgP => K O cptK oO fKO.
6529- apply: cptF; have := fam_compact_nbhs oO fKO cptK ctsf; apply: filter_app.
6530- by near=> g => /= gKO ? [z Kx <-]; apply: gKO.
6494+ move=> ctsf FF; split; first last.
6495+ move=> cptF; apply/compact_open_cvgP => K O cptK oO fKO.
6496+ apply: cptF; have := fam_compact_nbhs oO fKO cptK ctsf; apply: filter_app.
6497+ by near=> g => /= gKO ? [z Kx <-]; apply: gKO.
6498+ move/compact_open_cvgP=> cptOF; apply/cvg_sup; case=> K cptK R.
6499+ case=> D [[E oE <-] Ekf] /filterS; apply.
6500+ move: oE; rewrite openE => /(_ _ Ekf); case => ? [J entJ] EKR KfE.
6501+ near=> z; apply/KfE/EKR; case => u Kp; rewrite /= /set_val /= /eqincl /incl.
6502+ (have Ku : K u by rewrite inE in Kp); move: u Ku {D Kp}; near: z.
6503+ move/compact_near_coveringP/near_covering_withinP:(cptK); apply.
6504+ move=> u Ku; near (powerset_filter_from (@entourage V)) => E'.
6505+ have entE' : entourage E' by exact: (near (near_small_set _)).
6506+ pose C := f@^-1`(to_set E' (f u)).
6507+ pose B := \bigcup_(z in K `&` closure C) interior (to_set E' (f z)).
6508+ have oB : open B by apply: bigcup_open => ? ?; exact: open_interior.
6509+ have fKB : f @` (K `&` closure C) `<=` B.
6510+ move=> ? [z ? <-]; exists z => //; rewrite /interior.
6511+ by rewrite -nbhs_entourageE; exists E'.
6512+ have cptKC : compact (K `&` closure C).
6513+ by apply: compact_closedI => //; exact: closed_closure.
6514+ have := cptOF (K `&` closure C) B cptKC oB fKB.
6515+ exists (C, [set g | [set g x | x in K `&` closure C] `<=` B]).
6516+ split; last exact: cptOF.
6517+ by apply: (ctsf) => //; rewrite /filter_of -nbhs_entourageE; exists E'.
6518+ case=> z h /= [Cz KB Kz].
6519+ case: (KB (h z)); first by (exists z; split => //; apply: subset_closure).
6520+ move=> w [Kw Cw /interior_subset Jfwhz]; apply:subset_split_ent => //.
6521+ exists (f w); last apply:(near (small_ent_sub _) E') => //.
6522+ apply: subset_split_ent => //; exists (f u).
6523+ apply entourage_sym; apply:(near (small_ent_sub _) E') => //.
6524+ have [] := Cw (f@^-1` (to_set E' (f w))).
6525+ by apply: ctsf; rewrite /= -nbhs_entourageE; exists E'.
6526+ move=> r [Cr /= Ewr]; apply: subset_split_ent => //; exists (f r).
6527+ by apply:(near (small_ent_sub _) E') => //.
6528+ by apply entourage_sym; apply:(near (small_ent_sub _) E') => //.
65316529Unshelve. all: by end_near. Qed .
65326530
65336531End compact_open_uniform.
@@ -8105,43 +8103,36 @@ Local Notation "U '~>' V" :=
81058103Context {U V W : topologicalType}.
81068104
81078105Lemma continuous_curry (f : U * V ~> W) :
8108- continuous f -> continuous (( curry : ((U * V) ~> W) -> (U ~> V ~> W)) f ).
8106+ continuous f -> continuous (curry f : (U ~> V ~> W)).
81098107Proof .
81108108move=> ctsf x; apply/compact_open_cvgP => K O /= cptK oO fKO.
8111- near=> z => /= w /= [+ + <-]; near: z.
8109+ near=> z => w /= [+ + <-]; near: z.
81128110move/compact_near_coveringP/near_covering_withinP: cptK; apply.
81138111move=> v Kv; have [[P Q] [Px Qv] PQfO]: nbhs (x,v) (f@^-1` O).
81148112 by apply: ctsf; move: oO; rewrite openE; apply; apply: fKO; exists v.
81158113by exists (Q,P); [split => // |]; case=> b a /= [Qb Pa] Kb; apply: PQfO.
81168114Unshelve. all: by end_near. Qed .
81178115
81188116Lemma continuous_uncurry_regular (f : U ~> V ~> W):
8119- locally_compact [set: V] ->
8120- @regular V ->
8121- continuous f ->
8122- (forall u, continuous (f u)) ->
8123- continuous ((uncurry : (U ~> V ~> W) -> ((U * V) ~> W)) f).
8117+ locally_compact [set: V] -> @regular V -> continuous f ->
8118+ (forall u, continuous (f u)) -> continuous (uncurry f : ((U * V) ~> W)).
81248119Proof .
8125- move=> lcV reg ctsf ctsfp /= [u v] D; rewrite /= nbhsE; case=> O [oO Ofuv].
8126- move/filterS; apply.
8127- have [B] := @lcV v I; rewrite withinET; move=> Bv [cptB clB].
8120+ move=> lcV reg cf cfp /= [u v] D; rewrite /= nbhsE; case=> O [oO Ofuv] /filterS.
8121+ apply; have [B] := @lcV v I; rewrite withinET; move=> Bv [cptB clB].
81288122have [R Rv RO] : exists2 R, nbhs v R & (forall z, closure R z -> O (f u z)).
8129- have [] := reg v ((f u)@^-1` O); first by apply: ctsfp ; exact: open_nbhs_nbhs.
8130- by move=> R nR cRO ; exists R.
8123+ have [] := reg v ((f u)@^-1` O); first by apply: cfp ; exact: open_nbhs_nbhs.
8124+ by move=> R ? ? ; exists R.
81318125exists (f @^-1` ([set g | g @` (B `&` closure R) `<=` O]), (B `&` closure R)).
8132- split; [apply/ctsf /open_nbhs_nbhs; split | apply: filterI].
8126+ split; [apply/cf /open_nbhs_nbhs; split | apply: filterI] => // .
81338127 - apply: compact_open_open => //; apply: compact_closedI => //.
81348128 exact: closed_closure.
81358129 - by move=> ? [x [? + <-]]; apply: RO.
8136- - done.
81378130 - by apply: filterS; [exact: subset_closure|].
81388131by case=> a r /= [fBMO [Br] cmR]; apply: fBMO; exists r.
81398132Qed .
81408133
81418134Lemma continuous_uncurry (f : U ~> V ~> W):
8142- locally_compact [set: V] ->
8143- hausdorff_space V ->
8144- continuous f ->
8135+ locally_compact [set: V] -> hausdorff_space V -> continuous f ->
81458136 (forall u, continuous (f u)) ->
81468137 continuous ((uncurry : (U ~> V ~> W) -> ((U * V) ~> W)) f).
81478138Proof .
@@ -8150,9 +8141,7 @@ move=> v; have [B] := @lcV v I; rewrite withinET; move=> Bv [cptB clB].
81508141by move=> z; apply: (@compact_regular V hsdf v B).
81518142Qed .
81528143
8153- Lemma curry_continuous (f : U * V ~> W) :
8154- continuous f ->
8155- @regular U ->
8144+ Lemma curry_continuous (f : U * V ~> W) : continuous f -> @regular U ->
81568145 {for f, continuous ((curry : ((U * V) ~> W) -> (U ~> V ~> W)))}.
81578146Proof .
81588147move=> ctsf regU; apply/compact_open_cvgP.
@@ -8165,12 +8154,10 @@ suff : \forall x' \near u & i \near nbhs f, K x' ->
81658154 (\bigcap_(i in [set` N]) i) (curry i x').
81668155 apply: filter_app; near=> a b => + Ka => /(_ Ka) => ?.
81678156 by exists (\bigcap_(i in [set` N]) i).
8168- apply: filter_bigI_within=> R RN.
8169- have /set_mem [[M cptM _]] := Asub _ RN.
8157+ apply: filter_bigI_within=> R RN; have /set_mem [[M cptM _]] := Asub _ RN.
81708158have Rfu : R (curry f u) by exact: INf.
8171- case/(_ _ Rfu) => O [fMO oO] MOR.
8172- near=> p => /= Ki; apply: MOR => + [+ + <-] => _.
8173- move=> v Mv; move: v Mv Ki; near: p.
8159+ case/(_ _ Rfu) => O [fMO oO] MOR; near=> p => /= Ki; apply: MOR => + [+ + <-].
8160+ move=> _ v Mv; move: v Mv Ki; near: p.
81748161have umb : \forall y \near u, (forall b, M b -> nbhs (y, b) (f@^-1` O) ).
81758162 move/compact_near_coveringP/near_covering_withinP : (cptM); apply => v Mv.
81768163 have [[P Q] [Pu Qv] PQO] : nbhs (u,v) (f@^-1` O).
@@ -8181,26 +8168,23 @@ move/compact_near_coveringP/near_covering_withinP : (cptM); apply => v Mv.
81818168have [P' P'u cPO] := regU u _ umb.
81828169pose L := [set h | h @`((K `&` (closure P'))`*`M) `<=` O].
81838170exists (setT, (P' `*` L)).
8184- split => //; first exact: filterT.
8185- exists (P', L) => //; split => //.
8171+ split => //; [exact: filterT|]; exists (P', L) => //; split => //.
81868172 apply: open_nbhs_nbhs; split; first apply: compact_open_open => //.
8187- apply: compact_setM => //; apply: compact_closedI => //; exact: closed_closure.
8173+ apply: compact_setM => //; apply: compact_closedI => //.
8174+ exact: closed_closure.
81888175 by move=> ? [[a b] [[Ka /cPO +] Mb <-]] => /(_ _ Mb)/nbhs_singleton.
81898176case => b [a h] [? [Pa] +] Ma Ka; apply.
81908177exists (a,b); split => //; split => //; apply/subset_closure => //.
81918178Unshelve. all: by end_near. Qed .
81928179
81938180Lemma uncurry_continuous (f : U ~> V ~> W) :
8194- locally_compact [set: V] ->
8195- @regular V ->
8196- @regular U ->
8197- continuous f ->
8198- (forall u, continuous (f u)) ->
8181+ locally_compact [set: V] -> @regular V -> @regular U ->
8182+ continuous f -> (forall u, continuous (f u)) ->
81998183 {for f, continuous ((uncurry : (U ~> V ~> W) -> (U * V ~> W)))}.
82008184Proof .
82018185move=> lcV regV regU ctsf ctsfp; apply/compact_open_cvgP.
82028186 by apply: fmap_filter; exact:nbhs_filter.
8203- move=> /= K O cptK oO fKO; near=> h => /= ? [+ + <-]; near: h.
8187+ move=> /= K O cptK oO fKO; near=> h => ? [+ + <-]; near: h.
82048188move/compact_near_coveringP/near_covering_withinP : (cptK); apply.
82058189case=> u v Kuv.
82068190have : exists P Q, [/\ closed P, compact Q, nbhs u P, nbhs v Q & P `*` Q `<=` uncurry f @^-1` O].
@@ -8209,18 +8193,15 @@ have : exists P Q, [/\ closed P, compact Q, nbhs u P, nbhs v Q & P `*` Q `<=` un
82098193 by apply: fKO; exists (u,v).
82108194 case=> /= P' Q' [P'u Q'v] PQO.
82118195 have [B] := @lcV v I; rewrite withinET; move=> Bv [cptB clB].
8212- have [P Pu cPP'] := regU u P' P'u.
8213- have [Q Qv cQQ'] := regV v Q' Q'v.
8196+ have [P Pu cPP'] := regU u P' P'u; have [Q Qv cQQ'] := regV v Q' Q'v.
82148197 exists (closure P), (B `&` closure Q); split.
82158198 - exact: closed_closure.
82168199 - by apply: compact_closedI => //; exact: closed_closure.
82178200 - by apply: filterS; first exact: subset_closure.
82188201 - by apply: filterI=> //; apply: filterS; first exact: subset_closure.
82198202 - by case => a b [/cPP' ?] [_ /cQQ' ?]; exact: PQO.
8220- case=> P [Q [clP cptQ Pu Qv PQfO]].
8221- pose R := [set g : V ~> W | g @` Q `<=` O].
8222- have oR : open R by exact: compact_open_open.
8223- pose P' := f@^-1` R.
8203+ case=> P [Q [clP cptQ Pu Qv PQfO]]; pose R := [set g : V ~> W | g @` Q `<=` O].
8204+ (have oR : open R by exact: compact_open_open); pose P' := f@^-1` R.
82248205pose L := [set h : U ~> V ~> W | h @` (fst @` K `&` P) `<=` R].
82258206exists (((P`&` P') `*` Q), L).
82268207 split => /=.
@@ -8233,9 +8214,8 @@ exists (((P`&` P') `*` Q), L).
82338214 exact: cvg_fst.
82348215 move=> /= ? [a [Pa ?] <-] ? [b Qb <-].
82358216 by apply: (PQfO (a,b)); split => //; apply: nbhs_singleton.
8236- case; case => a b h [/= [[Pa P'a] Bb Lh] Kab].
8237- move/(_ (h a)): Lh ; rewrite /L/=/R/=.
8238- apply; first by exists a => //; split => //; exists (a,b).
8217+ case; case => a b h [/= [[? ?] ? Lh] ?].
8218+ apply: (Lh (h a)); first by exists a => //; split => //; exists (a,b).
82398219by exists b.
82408220Unshelve. all: by end_near. Qed .
8241- End currying.
8221+ End currying.
0 commit comments