Skip to content
Draft
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
44 commits
Select commit Hold shift + click to select a range
171af4c
sorgenfreyline and properties
motikaku Apr 23, 2025
fc56955
make
motikaku Apr 23, 2025
3825cc9
add Definition perfectly_normal_space
motikaku May 7, 2025
cc5886f
add Definition sdist
motikaku May 7, 2025
0049b38
simplify b_join
t6s May 9, 2025
56ff7d6
simplify open_separatend and subset_separated
t6s May 10, 2025
91fdb6d
simplify sorgenfrey_totally_disconnected
t6s May 10, 2025
fddffe8
almost complete proof of continuous_sdist
garrigue May 12, 2025
187db8b
simplify abs_subr_min
garrigue May 13, 2025
c75057f
fix Definition perfectly_normal_space
motikaku Jun 4, 2025
4e36b69
prove zeroset_sdist
garrigue May 14, 2025
7159793
clarify
garrigue May 14, 2025
39100b7
dl_shift
garrigue May 14, 2025
5f31565
add section
garrigue Jun 4, 2025
10b62e3
remove Admitted
motikaku Jun 20, 2025
4f05afc
add sorgenfrey epsilon delta
motikaku Jun 25, 2025
0e840fb
refactor and add inf_dlxz
motikaku Jul 2, 2025
a6ac2b7
better use of subrK/addKr
garrigue Jul 2, 2025
935b89b
refactor
motikaku Jul 9, 2025
940c0e9
remore duplication and shorten
garrigue Jul 9, 2025
ef027ad
extract sdist_in_lt and restructure
garrigue Jul 9, 2025
61e6c3a
extract lerPnormB
garrigue Jul 9, 2025
a9941f0
renaming and style
garrigue Jul 9, 2025
455ba84
simplify zeroset_sdist
garrigue Jul 10, 2025
f4223b3
typo
garrigue Jul 10, 2025
341f6fb
inf_shift
garrigue Jul 10, 2025
8661816
dlE drE
t6s Jul 10, 2025
517fbb8
shiftN, centerN
t6s Jul 10, 2025
a6f5ee7
simplify dlE drE
t6s Jul 10, 2025
dcbfa70
lt_(dl,dr)_set0
t6s Jul 10, 2025
e66798b
dr_shift/dl_shift
garrigue Jul 11, 2025
e00e50e
elim
garrigue Jul 11, 2025
90899ab
small clean up
garrigue Jul 22, 2025
fa36e57
move Section perfectlynormalspace separation_axioms.v to borel_hierar…
motikaku Jul 30, 2025
5c96415
fix indent
motikaku Jul 30, 2025
a1bd73e
WIP : add Vedenissoff theorem
motikaku Sep 3, 2025
8edfc7e
rebase
motikaku Sep 8, 2025
cf3b823
add proof of Vedenissoff
motikaku Sep 8, 2025
db97ab4
add proof of perfectly_normal_space_41
motikaku Sep 13, 2025
256f266
proof Lemmas perfectly_normal_space01_normal and perfectly_normal_spa…
motikaku Sep 17, 2025
d8f1632
rename and add Vedenissof proof
motikaku Oct 1, 2025
6a8942c
add proof of perfectly_normal_space_12
motikaku Oct 1, 2025
cfcadcb
wip
garrigue Oct 2, 2025
5702fc1
1 -> 2 almost done
garrigue Oct 2, 2025
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
2 changes: 1 addition & 1 deletion _CoqProject
Original file line number Diff line number Diff line change
Expand Up @@ -30,7 +30,6 @@ reals/constructive_ereal.v
reals/reals.v
reals/real_interval.v
reals/signed.v
reals/interval_inference.v
reals/prodnormedzmodule.v
reals/all_reals.v
experimental_reals/xfinmap.v
Expand Down Expand Up @@ -120,5 +119,6 @@ theories/kernel.v
theories/pi_irrational.v
theories/gauss_integral.v
theories/showcase/summability.v
theories/showcase/sorgenfreyline.v
analysis_stdlib/Rstruct_topology.v
analysis_stdlib/showcase/uniform_bigO.v
2 changes: 1 addition & 1 deletion experimental_reals/discrete.v
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@
(* Copyright (c) - 2016--2018 - Polytechnique *)

(* -------------------------------------------------------------------- *)
From Corelib Require Setoid.
From Coq Require Setoid.
From HB Require Import structures.
From mathcomp Require Import all_ssreflect all_algebra.
From mathcomp.classical Require Import boolp.
Expand Down
1 change: 0 additions & 1 deletion reals/all_reals.v
Original file line number Diff line number Diff line change
@@ -1,4 +1,3 @@
From mathcomp Require Export interval_inference.
From mathcomp Require Export constructive_ereal.
From mathcomp Require Export reals.
From mathcomp Require Export real_interval.
Expand Down
1,642 changes: 0 additions & 1,642 deletions reals/interval_inference.v

This file was deleted.

2 changes: 1 addition & 1 deletion reals/reals.v
Original file line number Diff line number Diff line change
Expand Up @@ -43,7 +43,7 @@
(* *)
(******************************************************************************)

From Corelib Require Import Setoid.
From Coq Require Import Setoid.
From HB Require Import structures.
From mathcomp Require Import all_ssreflect all_algebra archimedean.
From mathcomp Require Import boolp classical_sets set_interval.
Expand Down
1 change: 1 addition & 0 deletions theories/Make
Original file line number Diff line number Diff line change
Expand Up @@ -86,3 +86,4 @@ pi_irrational.v
gauss_integral.v
all_analysis.v
showcase/summability.v
showcase/sorgenfreyline.v
290 changes: 289 additions & 1 deletion theories/borel_hierarchy.v
Original file line number Diff line number Diff line change
Expand Up @@ -19,10 +19,13 @@ Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.

Import Order.TTheory GRing.Theory Num.Theory.
Import Order.TTheory GRing.Theory Num.Theory Num.Def.
Import numFieldNormedType.Exports.
Import numFieldTopology.Exports.
Import exp.

Local Open Scope classical_set_scope.
Local Open Scope ring_scope.

Section Gdelta_Fsigma.
Context {T : topologicalType}.
Expand Down Expand Up @@ -94,3 +97,288 @@ have /Baire : forall n, open (C n) /\ dense (C n).
- by apply: denseI => //; apply oB.
by rewrite -C0; exact: dense0.
Qed.

Section perfectlynormalspace.
Context (R : realType) (T : topologicalType).

Definition perfectly_normal_space (x : R) :=
forall E : set T, closed E ->
exists f : T -> R, continuous f /\ E = f @^-1` [set x].

Lemma perfectly_normal_spaceP x y : perfectly_normal_space x -> perfectly_normal_space y.
Proof.
move=>px E cE.
case:(px E cE) => f [] cf ->.
pose f' := f + cst (y - x).
exists f'.
split.
rewrite /f'.
move=> z.
apply: continuousD.
exact:cf.
exact:cst_continuous.
apply/seteqP.
rewrite /f' /cst /=.
split => z /=.
rewrite addrfctE => ->.
by rewrite subrKC.
rewrite addrfctE.
move/eqP.
by rewrite eq_sym -subr_eq opprB subrKC eq_sym => /eqP.
Qed.

Definition perfectly_normal_space' (x : R) :=
forall E : set T, open E ->
exists f : T -> R, continuous f /\ E = f @^-1` ~`[set x].

Definition perfectly_normal_space01 :=
forall E F : set T, closed E -> closed F -> [disjoint E & F] ->
exists f : T -> R, continuous f /\ E = f @^-1` [set 0] /\ F = f @^-1` [set 1]
/\ f @` [set: T] = `[0, 1]%classic.

Definition perfectly_normal_space_Gdelta :=
normal_space T /\ forall E : set T, closed E -> Gdelta E.

Lemma perfectly_normal_space01_normal :
perfectly_normal_space01 -> normal_space T.
Proof.
move=> pns01 A cA B /set_nbhsP[C] [oC AC CB].
case: (pns01 A (~` C) cA).
- by rewrite closedC.
- exact/disj_setPCl.
move=> f [/continuousP /= cf] [f0] [f1] f01.
exists (f @^-1` `]-oo, 1/2]).
apply/set_nbhsP.
exists (f @^-1` `]-oo, 1/2[).
split => //.
- exact: cf.
- by rewrite f0 => x /= ->; rewrite in_itv /=.
- by apply: preimage_subset => x /=; rewrite !in_itv /=; apply: ltW.
apply: subset_trans CB.
have<-:= proj1 (closure_id _).
have<-:= (setCK C).
rewrite f1 preimage_setC.
apply: preimage_subset => x /=; rewrite in_itv /=.
apply: contraTnot => ->.
by rewrite -ltNge ltr_pdivrMr // mul1r ltr1n.
have/continuousP /continuous_closedP:= cf.
apply.
exact: lray_closed.
Qed.

Lemma EFin_series (f : R^nat) : EFin \o series f = eseries (EFin \o f).
Proof.
apply/boolp.funext => n.
rewrite /series /eseries /=.
elim: n => [|n IH]; first by rewrite !big_geq.
by rewrite !big_nat_recr //= EFinD IH.
Qed.

Let perfectly_normal_space_12 : perfectly_normal_space_Gdelta -> perfectly_normal_space 0.
Proof.
move=> pnsGd E cE.
case: (pnsGd) => nT cEGdE.
have[U oU HE]:= cEGdE E cE.
have/boolp.choice[f_n Hn]: forall n, exists f : T -> R,
[/\ continuous f, range f `<=` `[0, 1], f @` E `<=` [set 0] & f @` (~` U n) `<=` [set 1]].
move=> n.
apply/uniform_separatorP.
apply: normal_uniform_separator => //.
- by rewrite closedC.
- rewrite HE -subsets_disjoint.
exact: bigcap_inf.
pose f_sum := fun n => \sum_(0 <= k < n) (f_n k \* cst (2^-k.+1)).
have cf_sum n : continuous (f_sum n).
rewrite /f_sum => x; elim: n => [|n IH].
rewrite big_geq //; exact: cst_continuous.
rewrite big_nat_recr //=; apply: continuousD => //.
apply/continuousM/cst_continuous.
by case: (Hn n) => /(_ x).
have f_sumE x : f_sum ^~ x = [series f_n k x * 2^-k.+1]_k.
apply/boolp.funext => n.
rewrite /f_sum /series /mk_sequence.
elim: n => [|n IH]; first by rewrite !big_geq.
by rewrite !big_nat_recr //= [X in X _ _ x]/GRing.add /= IH.
pose f := fun x => limn (f_sum^~ x).
rewrite /= in f.
exists f.
have ndf_sum y : {homo f_sum^~ y : a b / (a <= b)%N >-> a <= b}.
move=> a b ab.
rewrite /f_sum.
rewrite (big_cat_nat _ ab) //= lerDl.
elim/big_rec: _ => //= i f0 _.
rewrite /GRing.add /= => /le_trans; apply.
rewrite lerDr mulr_ge0 //.
case: (Hn i) => _ Hr _ _.
have /Hr /= := imageT (f_n i) y.
by rewrite in_itv /= => /andP[].
by rewrite invr_ge0 exprn_ge0.
(*have Hub y m k : (\sum_(m <= i < k) f_n i \* cst (2 ^- i.+1)) y <= 2^-m.
apply: (@le_trans _ _ (2^-m - 2^-k)); last first.
by rewrite gerBl invr_ge0 exprn_ge0.*)
have Hcvg y : cvgn (f_sum^~ y).
apply: nondecreasing_is_cvgn => //.
exists 1 => z /= [m] _ <-.
rewrite /f_sum.
apply: (@le_trans _ _ (1 - 2^-m)); last by rewrite gerBl invr_ge0 exprn_ge0.
elim: m => [|m IH]; first by rewrite big_geq // expr0 invr1 subrr.
rewrite big_nat_recr //=.
have Hm : f_n m y * 2 ^- m.+1 <= 2 ^- m.+1.
case: (Hn m) => _ Hr _ _.
have /Hr /= := imageT (f_n m) y.
rewrite in_itv /= => /andP[f0 f1].
by rewrite -[leRHS]mul1r ler_pM.
apply: (le_trans (lerD IH Hm)).
rewrite -addrA.
have -> : 2^- m = 2 * 2^-m.+1 :> R.
by rewrite exprS invfM mulrA mulfV // mul1r.
rewrite -{1}add1n {1}mulrnDr mulrDl !mul1r opprD !addrA -addrA.
by rewrite [_ + 2^- _]addrC subrr addr0.
split.
move=> x Nfx.
rewrite -filter_from_ballE.
case => eps /= eps0 HB.
pose n := (2 + truncn (- ln eps / ln 2))%N.
have eps0' : eps / 2 > 0 by exact: divr_gt0.
move/continuousP/(_ _ (ball_open (f_sum n x) eps0')) : (cf_sum n) => /= ofs.
rewrite nbhs_filterE.
rewrite fmapE.
rewrite nbhsE /=.
set B := _ @^-1` _ in ofs.
exists B.
split => //.
exact: ballxx.
apply: subset_trans (preimage_subset (f:=f) HB).
rewrite /B /preimage /ball => t /=.
have Hf y :
f y = f_sum n y + limn (fun n' => \sum_(n <= k < n') f_n k y * 2^- k.+1).
have /= := nondecreasing_telescope_sumey n _ (ndf_sum y).
rewrite EFin_lim // fin_numE /= => /(_ isT).
rewrite (eq_eseriesr (g:=fun i => ((f_n i \* cst (2 ^- i.+1)) y)%:E));
last first.
move => i _.
rewrite /f_sum -EFinD big_nat_recr //=.
by rewrite [X in X _ _ y]/GRing.add /= addrAC subrr add0r.
move/(f_equal (fun x => (f_sum n y)%:E + x)).
rewrite addrA addrAC -EFinB subrr add0r => H.
apply: EFin_inj.
rewrite -H EFinD; congr (_ + _).
rewrite -EFin_lim.
apply/congr_lim/boolp.funext => k /=.
elim: k => [|k IH]; first by rewrite !big_geq.
case: (leP n k) => nk.
by rewrite !big_nat_recr //= IH EFinD.
by rewrite !big_geq.
move: (Hcvg y).
by rewrite is_cvg_series_restrict f_sumE.
move=> feps.
rewrite !Hf opprD addrA (addrAC (f_sum n x)) -(addrA (_ - _)).
apply: (le_lt_trans (ler_normD _ _)).
rewrite (splitr eps).
apply: ltr_leD => //.
admit.
apply/seteqP; split => x /= Hx.
apply: EFin_inj.
rewrite -EFin_lim // f_sumE EFin_series /= eseries0 // => i _ _ /=.
case: (Hn i) => _ _ /(_ (f_n i x)) /= => ->.
by rewrite mul0r.
by exists x.
apply: contraPP Hx.
rewrite (_ : (~ E x) = setC E x) // HE setC_bigcap /= => -[] i _ HU.
case: (Hn i) => _ _ _ /(_ (f_n i x)) /= => Hfix.
rewrite /f => Hf.
have := (nondecreasing_cvgn_le (ndf_sum x) (Hcvg x) i.+1).
have := ndf_sum x _ _ (leq0n i.+1).
rewrite Hf {1}/f_sum big_geq // => /[swap] fi_le0.
rewrite le0r ltNge fi_le0 orbF.
rewrite /f_sum big_nat_recr //= [X in X _ _ x]/GRing.add /= -/(f_sum i).
rewrite Hfix; last by exists x.
rewrite mul1r /cst => /eqP fi0.
have := ndf_sum x _ _ (leq0n i).
rewrite {1}/f_sum big_geq // -[0 x]fi0 gerDl.
by rewrite leNgt invr_gt0 exprn_gt0.
Admitted.

Let perfectly_normal_space_13 : perfectly_normal_space_Gdelta -> perfectly_normal_space' 0.
move=> pnsGd E oE.
case: (pnsGd) => nT cEGdE.
have clcpE: closed (~`E).
by rewrite closedC.
have[U oU hE]:= cEGdE (~` E) clcpE.
have/boolp.choice[f_n Hn]: forall n, exists f : T -> R,
[/\ continuous f, range f `<=` `[0, 1], f @` (~` E) `<=` [set 0] & f @` (~` U n) `<=` [set 1]].
move=> n.
apply/uniform_separatorP.
apply: normal_uniform_separator => //.
- by rewrite closedC.
- rewrite hE.
rewrite -subsets_disjoint.
exact: bigcap_inf.
Admitted.

Let perfectly_normal_space_23 : perfectly_normal_space 0 -> perfectly_normal_space' 0.
Proof.
Admitted.

Let perfectly_normal_space_32 : perfectly_normal_space' 0 -> perfectly_normal_space 0.
Proof.
move=> pns' E cE; case: (pns' (~`E)).
by rewrite openC.
move=> f [cf f0]; exists f.
split.
by [].
by rewrite -[RHS]setCK preimage_setC -f0 setCK.
Qed.

Let perfectly_normal_space_24 : perfectly_normal_space 0 -> perfectly_normal_space01.
Proof.
Admitted.

Let perfectly_normal_space_34 : perfectly_normal_space' 0 -> perfectly_normal_space01.
Proof.
Admitted.

Local Lemma perfectly_normal_space_42 :
perfectly_normal_space01 -> perfectly_normal_space 0.
Proof.
move=> + E cE => /(_ E set0 cE closed0).
rewrite disj_set2E setI0 eqxx => /(_ erefl).
case=> f [] cf [] Ef [] _ _.
by exists f; split.
Qed.

Let perfectly_normal_space_41 :
perfectly_normal_space01 -> perfectly_normal_space_Gdelta.
Proof.
move=> pns01; split; first exact: perfectly_normal_space01_normal.
move=> E cE; case:(perfectly_normal_space_42 pns01 cE) => // f [] cf Ef.
exists (fun n => f @^-1` `]-n.+1%:R^-1,n.+1%:R^-1[%classic).
by move=> n; move/continuousP: cf; apply; exact: itv_open.
rewrite -preimage_bigcap (_ : bigcap _ _ = [set 0])//.
rewrite eqEsubset; split; last first.
by move=> x -> n _ /=; rewrite in_itv//= oppr_lt0 andbb invr_gt0.
apply: subsetC2; rewrite setC_bigcap => x /= /eqP x0.
exists (trunc `|x^-1|) => //=; rewrite in_itv/= -ltr_norml ltNge.
apply/negP; rewrite negbK.
by rewrite invf_ple ?posrE// ?normr_gt0// -normfV ltW// truncnS_gt.
Qed.

Theorem Vedenissoff_closed : perfectly_normal_space_Gdelta <-> perfectly_normal_space 0.
Proof.
move: perfectly_normal_space_12 perfectly_normal_space_23 perfectly_normal_space_34 perfectly_normal_space_41.
tauto.
Qed.

Theorem Vedenissoff_open : perfectly_normal_space_Gdelta <-> perfectly_normal_space' 0.
Proof.
move: perfectly_normal_space_12 perfectly_normal_space_23 perfectly_normal_space_34 perfectly_normal_space_41.
tauto.
Qed.

Theorem Vedenissoff01 : perfectly_normal_space_Gdelta <-> perfectly_normal_space01.
Proof.
move: perfectly_normal_space_12 perfectly_normal_space_23 perfectly_normal_space_34 perfectly_normal_space_41.
tauto.
Qed.

End perfectlynormalspace.
16 changes: 16 additions & 0 deletions theories/normedtype_theory/pseudometric_normed_Zmodule.v
Original file line number Diff line number Diff line change
Expand Up @@ -95,6 +95,22 @@ Proof. by rewrite funeqE => x /=; rewrite addr0. Qed.
Lemma center0 : center 0 = id.
Proof. by rewrite oppr0 shift0. Qed.

Lemma centerN (x y : R) : center x (- y) = - shift x y.
Proof. by rewrite /shift opprD. Qed.

Lemma shiftN (x y : R) : shift x (- y) = - center x y.
Proof. by rewrite /shift opprD opprK. Qed.

Lemma image_centerN (E : set R) (x : R) :
[set center x (- y) | y in E] =[set - (shift x y) | y in E].
Proof.
by apply/seteqP; split => y [] u Eu <- /=; exists u => //; rewrite opprD.
Qed.

Lemma image_shiftN (E : set R) (x : R) :
[set shift x (- y) | y in E] = [set - (center x y) | y in E].
Proof. by rewrite -(opprK x) image_centerN opprK. Qed.

End Shift.
Arguments shift {R} x / y.
Notation center c := (shift (- c)).
Expand Down
Loading
Loading