Skip to content

Commit 0aa108b

Browse files
affeldt-aistIshiguroYoshihirot6s
committed
basic facts about Gdelta sets
Co-authored-by: IshiguroYoshihiro <[email protected]> Co-authored-by: Takafumi Saikawa <[email protected]>
1 parent 9129537 commit 0aa108b

File tree

8 files changed

+187
-6
lines changed

8 files changed

+187
-6
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -18,8 +18,26 @@
1818
- in `exp.v:
1919
+ lemma `norm_expR`
2020

21+
- in `reals.v`:
22+
+ definitions `rational`, `irrational`
23+
+ lemmas `irrationalE`, `rationalP`
24+
25+
- in `topology_structure.v`:
26+
+ lemmas `denseI`, `dense0`
27+
28+
- in `pseudometric_normed_Zmodule.v`:
29+
+ lemma `dense_set1C`
30+
31+
- new file `borel_hierarchy.v`:
32+
+ definitions `Gdelta`, `Fsigma`
33+
+ lemmas `closed_Fsigma`, `Gdelta_measurable`, `Gdelta_subspace_open`,
34+
`irrational_Gdelta`, `not_rational_Gdelta`
35+
2136
### Changed
2237

38+
- moved from `pi_irrational.v` to `reals.v` and changed
39+
+ definition `rational`
40+
2341
### Renamed
2442

2543
- in `lebesgue_stieltjes_measure.v`:

_CoqProject

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -96,6 +96,7 @@ theories/lebesgue_stieltjes_measure.v
9696
theories/derive.v
9797
theories/measure.v
9898
theories/numfun.v
99+
theories/borel_hierarchy.v
99100

100101
theories/lebesgue_integral_theory/simple_functions.v
101102
theories/lebesgue_integral_theory/lebesgue_integral_definition.v

reals/reals.v

Lines changed: 35 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -36,6 +36,11 @@
3636
(* Rceil x == the ceil of x as a real number, i.e., - Rfloor (- x) *)
3737
(* ``` *)
3838
(* *)
39+
(* ``` *)
40+
(* @rational R == set of rational numbers of type R : realType *)
41+
(* @irrational R == the complement of rational R *)
42+
(* ``` *)
43+
(* *)
3944
(******************************************************************************)
4045

4146
From Corelib Require Import Setoid.
@@ -708,3 +713,33 @@ by rewrite 2!ratr_int mulr_natr (le_lt_trans _ c).
708713
Qed.
709714

710715
End rat_in_itvoo.
716+
717+
Section rational.
718+
Context {R : realType}.
719+
720+
Definition rational : set R := ratr @` [set: rat].
721+
722+
Lemma rationalP (r : R) :
723+
rational r <-> exists (a : int) (b : nat), r = a%:~R / b%:R.
724+
Proof.
725+
split=> [[q _ <-{r}]|[a [b ->]]]; last first.
726+
exists (a%:~R / b%:R) => //.
727+
by rewrite ratr_is_multiplicative ratr_int fmorphV//= ratr_nat.
728+
have [n d nd] := ratP q.
729+
have [n0|n0] := leP 0 n.
730+
exists n, d.+1.
731+
by rewrite ratr_is_multiplicative ratr_int fmorphV/= ratr_nat.
732+
exists (- `|n|), d.+1.
733+
by rewrite ratr_is_multiplicative ltr0_norm// opprK fmorphV//= !ratr_int.
734+
Qed.
735+
736+
Definition irrational : set R := ~` rational.
737+
738+
Lemma irrationalE : irrational = \bigcap_(q : rat) ~` (ratr @` [set q]).
739+
Proof.
740+
apply/seteqP; split => [r/= rE q _/=|r/= rE [q] _ qr]; last first.
741+
by apply: (rE q Logic.I) => /=; exists q.
742+
by apply: contra_not rE => -[_ -> <-{r}]; exists q.
743+
Qed.
744+
745+
End rational.

theories/Make

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -60,6 +60,7 @@ lebesgue_measure.v
6060
derive.v
6161
measure.v
6262
numfun.v
63+
borel_hierarchy.v
6364

6465
lebesgue_integral_theory/simple_functions.v
6566
lebesgue_integral_theory/lebesgue_integral_definition.v

theories/borel_hierarchy.v

Lines changed: 96 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,96 @@
1+
From mathcomp Require Import all_ssreflect all_algebra archimedean finmap.
2+
From mathcomp Require Import mathcomp_extra unstable boolp classical_sets.
3+
From mathcomp Require Import functions cardinality fsbigop interval_inference.
4+
From mathcomp Require Import reals ereal topology normedtype sequences.
5+
From mathcomp Require Import real_interval numfun esum measure lebesgue_measure.
6+
7+
(**md**************************************************************************)
8+
(* # Basic facts about G-delta and F-sigma sets *)
9+
(* *)
10+
(* ``` *)
11+
(* Gdelta S == S is countable intersection of open sets *)
12+
(* Gdelta_dense S == S is a countable intersection of dense open sets *)
13+
(* Fsigma S == S is countable union of closed sets *)
14+
(* ``` *)
15+
(* *)
16+
(******************************************************************************)
17+
18+
Set Implicit Arguments.
19+
Unset Strict Implicit.
20+
Unset Printing Implicit Defensive.
21+
22+
Import Order.TTheory GRing.Theory Num.Theory.
23+
Import numFieldNormedType.Exports.
24+
25+
Local Open Scope classical_set_scope.
26+
27+
Section Gdelta_Fsigma.
28+
Context {T : topologicalType}.
29+
Implicit Type S : set T.
30+
31+
Definition Gdelta S :=
32+
exists2 F : (set T)^nat, (forall i, open (F i)) & S = \bigcap_i (F i).
33+
34+
Lemma open_Gdelta S : open S -> Gdelta S.
35+
Proof. by exists (fun=> S)=> //; rewrite bigcap_const. Qed.
36+
37+
Definition Gdelta_dense S :=
38+
exists2 F : (set T)^nat,
39+
(forall i, open (F i) /\ dense (F i)) & S = \bigcap_i (F i).
40+
41+
Definition Fsigma S :=
42+
exists2 F : (set T)^nat, (forall i, closed (F i)) & S = \bigcup_i (F i).
43+
44+
Lemma closed_Fsigma S : closed S -> Fsigma S.
45+
Proof. by exists (fun=> S)=> //; rewrite bigcup_const. Qed.
46+
47+
End Gdelta_Fsigma.
48+
49+
Lemma Gdelta_measurable {R : realType} (S : set R) : Gdelta S -> measurable S.
50+
Proof.
51+
move=> [] B oB ->; apply: bigcapT_measurable => i.
52+
exact: open_measurable.
53+
Qed.
54+
55+
Lemma Gdelta_subspace_open {R : realType} (A : set R) (S : set (subspace A)) :
56+
open A -> Gdelta S -> Gdelta (A `&` S).
57+
Proof.
58+
move=> oA [/= S_ oS_ US]; exists (fun n => A `&` (S_ n)).
59+
by move=> ?; rewrite open_setSI.
60+
by rewrite bigcapIr// US.
61+
Qed.
62+
63+
Section irrational_Gdelta.
64+
Context {R : realType}.
65+
66+
Lemma irrational_Gdelta : Gdelta_dense (@irrational R).
67+
Proof.
68+
rewrite irrationalE.
69+
have /pcard_eqP/bijPex[f bijf] := card_rat.
70+
pose f1 := 'pinv_(fun => 0%R) [set: rat] f.
71+
exists (fun n => ~` [set ratr (f1 n)]).
72+
move=> n; rewrite openC; split; last exact: dense_set1C.
73+
exact/accessible_closed_set1/hausdorff_accessible/Rhausdorff.
74+
apply/seteqP; split => [/= r/= rE n _/= rf1n|/=r rE q _/= [_ -> qr]].
75+
by apply: (rE (f1 n)) => //=; exists (f1 n).
76+
apply: (rE (f q)) => //=.
77+
by rewrite /f1 pinvKV ?inE//=; exact: set_bij_inj bijf.
78+
Qed.
79+
80+
End irrational_Gdelta.
81+
82+
Lemma not_rational_Gdelta (R : realType) : ~ Gdelta (@rational R).
83+
Proof.
84+
apply/forall2NP => A; apply/not_andP => -[oA ratrA].
85+
have dense_A n : dense (A n).
86+
move=> D D0 /(@dense_rat R D D0); apply/subset_nonempty; apply: setIS.
87+
by rewrite -/(@rational R) ratrA; exact: bigcap_inf.
88+
have [/= B oB irratB] := @irrational_Gdelta R.
89+
pose C : (set R)^nat := fun n => A n `&` B n.
90+
have C0 : set0 = \bigcap_i C i by rewrite bigcapI -ratrA -irratB setICr.
91+
have /Baire : forall n, open (C n) /\ dense (C n).
92+
move=> n; split.
93+
- by apply: openI => //; apply oB.
94+
- by apply: denseI => //; apply oB.
95+
by rewrite -C0; exact: dense0.
96+
Qed.

theories/normedtype_theory/pseudometric_normed_Zmodule.v

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -859,6 +859,22 @@ Qed.
859859

860860
End open_itv_subset.
861861

862+
Lemma dense_set1C {R : realType} (r : R) : dense (~` [set r]).
863+
Proof.
864+
move=> /= O O0 oO.
865+
suff [s Os /eqP sr] : exists2 s, O s & s != r by exists s; split.
866+
case: O0 => o Oo.
867+
have [->{r}|ro] := eqVneq r o; last by exists o => //; rewrite eq_sym.
868+
have [e' /= e'0 e'o] := open_itvoo_subset oO Oo.
869+
near (0:R)^'+ => e.
870+
exists (o + e/2)%R; last by rewrite gt_eqF// ltrDl// divr_gt0.
871+
apply: (e'o e) => //=.
872+
by rewrite sub0r normrN gtr0_norm.
873+
rewrite in_itv/= !ltrD2l; apply/andP; split.
874+
by rewrite (@lt_le_trans _ _ 0%R) ?divr_ge0// ltrNl oppr0.
875+
by rewrite gtr_pMr// invf_lt1// ltr1n.
876+
Unshelve. all: by end_near. Qed.
877+
862878
Section pseudoMetricNormedZmod_numFieldType.
863879
Variables (R : numFieldType) (V : pseudoMetricNormedZmodType R).
864880
Variables (I : Type) (F : set_system I) (FF : Filter F) (f : I -> V) (y : V).

theories/pi_irrational.v

Lines changed: 3 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -28,9 +28,6 @@ apply/measurable_funTS; apply: continuous_measurable_fun.
2828
exact/continuous_horner.
2929
Qed.
3030

31-
(* TODO: move somewhere to classical *)
32-
Definition rational {R : realType} (x : R) := exists m n, x = (m%:~R / n%:R)%R.
33-
3431
Module pi_irrational.
3532
Local Open Scope ring_scope.
3633

@@ -409,10 +406,10 @@ End pi_irrational.
409406

410407
Lemma pi_irrationnal {R : realType} : ~ rational (pi : R).
411408
Proof.
412-
move=> [a [b]]; have [->|b0 piratE] := eqVneq b O.
409+
move/rationalP => [a [b]]; have [->|b0 piratE] := eqVneq b O.
413410
by rewrite invr0 mulr0; apply/eqP; rewrite gt_eqF// pi_gt0.
414-
have [na ana] : exists na, (a%:~R = na %:R :> R)%R.
415-
exists `|a|; rewrite natr_absz gtr0_norm//.
411+
have [na ana] : exists na, (a%:~R = na%:R :> R)%R.
412+
exists `|a|%N; rewrite natr_absz gtr0_norm//.
416413
by have := @pi_gt0 R; rewrite piratE pmulr_lgt0 ?invr_gt0 ?ltr0n ?lt0n// ltr0z.
417414
rewrite {}ana in piratE.
418415
have [N _] := pi_irrational.intfsin_small b0 (esym piratE) (@ltr01 R).

theories/topology_theory/topology_structure.v

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -935,6 +935,23 @@ move=> /existsNP[X /not_implyP[[x Xx] /not_implyP[ Ox /forallNP A]]].
935935
by exists X; split; [exists x | rewrite -subset0; apply/A].
936936
Qed.
937937

938+
Lemma denseI (T : topologicalType) (A B : set T) :
939+
open A -> dense A -> dense B -> dense (A `&` B).
940+
Proof.
941+
move=> oA dA dB C C0 oC.
942+
have CA0 : C `&` A !=set0 by exact: dA.
943+
suff: (C `&` A) `&` B !=set0 by rewrite setIA.
944+
by apply: dB => //; exact: openI.
945+
Qed.
946+
947+
Lemma dense0 {R : ptopologicalType} : ~ dense (@set0 R).
948+
Proof.
949+
apply/existsNP; exists setT.
950+
apply/not_implyP; split; first exact/set0P/setT0.
951+
apply/not_implyP; split; first exact: openT.
952+
by rewrite setTI => -[].
953+
Qed.
954+
938955
Section ClopenSets.
939956
Implicit Type T : topologicalType.
940957

0 commit comments

Comments
 (0)