|
| 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. |
0 commit comments