Skip to content

Commit c9c9435

Browse files
authored
adjacent sets, cut, limit point (#1772)
* adjacent sets, cut, limit point
1 parent 61acc8d commit c9c9435

File tree

3 files changed

+158
-4
lines changed

3 files changed

+158
-4
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,13 @@
1313
- in `normed_module.v`:
1414
+ lemma `limit_point_infinite_setP`
1515

16+
- in `sequences.v`:
17+
+ lemma `increasing_seq_injective`
18+
+ definition `adjacent_set`
19+
+ lemmas `adjacent_sup_inf`, `adjacent_sup_inf_unique`
20+
+ definition `cut`
21+
+ lemmas `cut_adjacent`, `infinite_bounded_limit_point_nonempty`
22+
1623
### Changed
1724

1825
### Renamed

_CoqProject

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -84,8 +84,8 @@ theories/normedtype_theory/urysohn.v
8484
theories/normedtype_theory/vitali_lemma.v
8585
theories/normedtype_theory/normedtype.v
8686

87-
theories/realfun.v
8887
theories/sequences.v
88+
theories/realfun.v
8989
theories/exp.v
9090
theories/trigo.v
9191
theories/esum.v

theories/sequences.v

Lines changed: 150 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,9 +1,9 @@
11
(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *)
22
From HB Require Import structures.
33
From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint.
4-
From mathcomp Require Import interval archimedean.
5-
From mathcomp Require Import boolp classical_sets.
6-
From mathcomp Require Import functions set_interval reals interval_inference.
4+
From mathcomp Require Import interval interval_inference archimedean.
5+
From mathcomp Require Import mathcomp_extra boolp contra classical_sets.
6+
From mathcomp Require Import functions cardinality set_interval reals.
77
From mathcomp Require Import ereal topology tvs normedtype landau.
88

99
(**md**************************************************************************)
@@ -93,6 +93,12 @@ From mathcomp Require Import ereal topology tvs normedtype landau.
9393
(* of extended reals *)
9494
(* ``` *)
9595
(* *)
96+
(* ``` *)
97+
(* adjacent_set L R == L and R are two adjacent sets of real numbers *)
98+
(* cut L R == L and R are two sets of real numbers that form *)
99+
(* a cut *)
100+
(* ``` *)
101+
(* *)
96102
(* ## Bounded functions *)
97103
(* This section proves Baire's Theorem, stating that complete normed spaces *)
98104
(* are Baire spaces, and Banach-Steinhaus' theorem, stating that between a *)
@@ -202,6 +208,20 @@ split; first by move=> u_nondec; apply: le_mono; apply: homo_ltn lt_trans _.
202208
by move=> u_incr n; rewrite lt_neqAle eq_le !u_incr leqnSn ltnn.
203209
Qed.
204210

211+
Lemma increasing_seq_injective d {T : orderType d} (f : T^nat) :
212+
increasing_seq f -> injective f.
213+
Proof.
214+
move=> incrf a b fafb; apply: contrapT => /eqP; rewrite neq_lt => /orP[ab|ba].
215+
- have : (f a < f b)%O.
216+
rewrite (@lt_le_trans _ _ (f a.+1))//.
217+
by move/increasing_seqP : incrf; exact.
218+
by move: ab; rewrite incrf.
219+
by rewrite fafb ltxx.
220+
- have := incrf a b.
221+
rewrite fafb lexx => /esym.
222+
by rewrite -leEnat leNgt ba.
223+
Qed.
224+
205225
Lemma decreasing_seqP d (T : porderType d) (u_ : T ^nat) :
206226
(forall n, u_ n > u_ n.+1)%O <-> decreasing_seq u_.
207227
Proof.
@@ -2651,6 +2671,133 @@ apply: nondecreasing_cvgn_le; last exact: is_cvg_geometric_series.
26512671
by apply: nondecreasing_series => ? _ /=; rewrite pmulr_lge0 // exprn_gt0.
26522672
Qed.
26532673

2674+
Section adjacent_cut.
2675+
Context {R : realType}.
2676+
Implicit Types L D E : set R.
2677+
2678+
Definition adjacent_set A B :=
2679+
[/\ A !=set0, B !=set0, (forall x y, A x -> B y -> x <= y) &
2680+
forall e : {posnum R}, exists2 xy, xy \in A `*` B & xy.2 - xy.1 < e%:num].
2681+
2682+
Lemma adjacent_sup_inf A B : adjacent_set A B -> sup A = inf B.
2683+
Proof.
2684+
case=> A0 B0 AB_le AB_eps; apply/eqP; rewrite eq_le; apply/andP; split.
2685+
by apply: ge_sup => // x Ax; apply: lb_le_inf => // y By; exact: AB_le.
2686+
apply/ler_addgt0Pl => _ /posnumP[e]; rewrite -lerBlDr.
2687+
have [[x y]/=] := AB_eps e.
2688+
rewrite !inE => -[/= Ax By] /ltW yxe.
2689+
rewrite (le_trans _ yxe)// lerB//.
2690+
- by rewrite ge_inf//; exists x => // z; exact: AB_le.
2691+
- by rewrite ub_le_sup//; exists y => // z Az; exact: AB_le.
2692+
Qed.
2693+
2694+
Lemma adjacent_sup_inf_unique A B M : adjacent_set A B ->
2695+
ubound A M -> lbound B M -> M = sup A.
2696+
Proof.
2697+
move=> [A0 B0 AB_leq AB_eps] AM BM.
2698+
apply/eqP; rewrite eq_le ge_sup// andbT (@adjacent_sup_inf A B)//.
2699+
exact: lb_le_inf.
2700+
Qed.
2701+
2702+
Definition cut L B := [/\ L !=set0, B !=set0,
2703+
(forall x y, L x -> B y -> x < y) & L `|` B = [set: R] ].
2704+
2705+
Lemma cut_adjacent A B : cut A B -> adjacent_set A B.
2706+
Proof.
2707+
move=> [A0 B0 ABlt ABT]; split => //; first by move=> x y Ax By; exact/ltW/ABlt.
2708+
move: A0 B0 => [a aA] [b bB] e.
2709+
have ba0 : b - a > 0 by rewrite subr_gt0 ABlt.
2710+
have [N N0 baNe] : exists2 N, N != 0 & (b - a) / N%:R < e%:num.
2711+
exists (truncn ((b - a) / e%:num)).+1 => //.
2712+
by rewrite ltr_pdivrMr// mulrC -ltr_pdivrMr// truncnS_gt.
2713+
pose a_ i := a + i%:R * (b - a) / N%:R.
2714+
pose k : nat := [arg min_(i < @ord_max N | a_ i \in B) i].
2715+
have ? : a_ (@ord_max N) \in B.
2716+
by rewrite /a_ /= mulrAC divff ?pnatr_eq0// mul1r subrKC; exact/mem_set.
2717+
have k_gt0 : (0 < k)%N.
2718+
rewrite /k; case: arg_minnP => // /= i + aBi.
2719+
contra; rewrite leqn0 => /eqP ->.
2720+
rewrite /a_ !mul0r addr0; apply/negP => /set_mem/(ABlt _ _ aA).
2721+
by rewrite ltxx.
2722+
have akN1A : a_ k.-1 \in A.
2723+
rewrite /k; case: arg_minnP => // /= i aiB aBi.
2724+
have i0 : i != ord0.
2725+
contra: aiB => ->.
2726+
rewrite /a_ !mul0r addr0; apply/negP => /set_mem/(ABlt _ _ aA).
2727+
by rewrite ltxx.
2728+
apply/mem_set/boolp.notP => abs.
2729+
have {}abs : a_ i.-1 \in B.
2730+
by move/seteqP : ABT => [_ /(_ (a_ i.-1) Logic.I)] [//|/mem_set].
2731+
have iN : (i.-1 < N.+1)%N by rewrite prednK ?lt0n// ltnW.
2732+
have := aBi (Ordinal iN) abs.
2733+
apply/negP; rewrite -ltnNge/=.
2734+
by case: i => -[//|? ?] in i0 iN abs aiB aBi *.
2735+
have akB : a_ k \in B by rewrite /k; case: arg_minnP => // /= i aiB aBi.
2736+
exists (a_ k.-1, a_ k); first by rewrite !inE; split => //=; exact/set_mem.
2737+
rewrite /a_ opprD addrACA subrr add0r -!mulrA -!mulrBl.
2738+
by rewrite -natrB ?leq_pred// -subn1 subKn// mul1r.
2739+
Qed.
2740+
2741+
Lemma infinite_bounded_limit_point_nonempty E :
2742+
infinite_set E -> bounded_set E -> limit_point E !=set0.
2743+
Proof.
2744+
move=> infiniteE boundedE.
2745+
have E0 : E !=set0.
2746+
apply/set0P/negP => /eqP E0.
2747+
by move: infiniteE; rewrite E0; apply; exact: finite_set0.
2748+
have ? : ProperFilter (globally E).
2749+
by case: E0 => x Ex; exact: globally_properfilter Ex.
2750+
pose A := [set x | infinite_set (`[x, +oo[ `&` E)].
2751+
have A0 : A !=set0.
2752+
move/ex_bound : boundedE => [M EM]; exists (- M).
2753+
rewrite /A /= setIidr// => x Ex /=.
2754+
by rewrite in_itv/= andbT lerNnormlW// EM.
2755+
pose B := ~` A.
2756+
have B0 : B !=set0.
2757+
move/ex_bound : boundedE => [M EM]; exists (M + 1).
2758+
rewrite /B /A /= (_ : _ `&` _ = set0)// -subset0 => x []/=.
2759+
rewrite in_itv/= andbT => /[swap] /EM/= /ler_normlW xM.
2760+
by move/le_trans => /(_ _ xM); rewrite leNgt ltrDl ltr01.
2761+
have Ale_closed x y : A x -> y <= x -> A y.
2762+
rewrite /A /= => xE yx.
2763+
rewrite (@itv_bndbnd_setU _ _ _ (BLeft x))//.
2764+
by apply: contra_not xE; rewrite setIUl finite_setU => -[].
2765+
have ABlt x y : A x -> B y -> x < y.
2766+
by move=> Ax By; rewrite ltNge; apply/negP => /(Ale_closed _ _ Ax).
2767+
have AB : cut A B by split => //; rewrite /B setUv.
2768+
pose l := sup A. (* the real number defined by the cut (A, B) *)
2769+
have infleE (e : R) (e0 : e > 0) :infinite_set (`]l - e, +oo[ `&` E).
2770+
suff : A (l - e).
2771+
apply: contra_not => leE.
2772+
rewrite -setU1itv// setIUl finite_setU; split => //.
2773+
by apply/(sub_finite_set _ (finite_set1 (l - e))); exact: subIsetl.
2774+
have : has_sup A.
2775+
by split => //; case: B0 => d dB; exists d => z Az; exact/ltW/ABlt.
2776+
move/(sup_adherent e0) => [r Ar].
2777+
by rewrite -/l => /ltW ler; exact: (Ale_closed _ _ Ar).
2778+
have finleE (e : R) (e0 : e > 0) : finite_set (`[l + e, +oo[ `&` E).
2779+
suff : B (l + e) by rewrite /B/= /A/= => /contrapT.
2780+
have : has_inf B.
2781+
by split => //; case: A0 => g gA; exists g => z Bz; exact/ltW/ABlt.
2782+
move/(inf_adherent e0) => [r Br].
2783+
rewrite -(adjacent_sup_inf (cut_adjacent AB)) -/l => /ltW rle Ale.
2784+
have /ABlt := Ale_closed _ _ Ale rle.
2785+
by move/(_ _ Br); rewrite ltxx.
2786+
exists l; apply/limit_point_infinite_setP => /= U.
2787+
rewrite /nbhs/= /nbhs_ball_/= => -[e /= e0].
2788+
rewrite -[ball_ _ _ _]/(ball _ _) => leU.
2789+
have : infinite_set (`]l - e, l + e[ `&` E).
2790+
rewrite (_ : _ `&` _ =
2791+
`]l - e, +oo[ `&` E `\` `[l + e, +oo[ `&` E); last first.
2792+
rewrite setDE setCI setIUr -(setIA _ _ (~` E)) setICr setI0 setU0.
2793+
by rewrite setIAC -setDE [in LHS]set_itv_splitD.
2794+
by apply: infinite_setD; [exact: infleE|exact: finleE].
2795+
apply/contra_not/sub_finite_set; apply: setSI.
2796+
by move: leU; rewrite ball_itv.
2797+
Qed.
2798+
2799+
End adjacent_cut.
2800+
26542801
Section banach_contraction.
26552802

26562803
Context {R : realType} {X : completeNormedModType R} (U : set X).

0 commit comments

Comments
 (0)