Skip to content

Commit

Permalink
at_left version of an at_right lemma
Browse files Browse the repository at this point in the history
  • Loading branch information
affeldt-aist committed Jul 18, 2024
1 parent 689880d commit 6a1c4a4
Show file tree
Hide file tree
Showing 2 changed files with 19 additions and 0 deletions.
3 changes: 3 additions & 0 deletions CHANGELOG_UNRELEASED.md
Original file line number Diff line number Diff line change
Expand Up @@ -36,6 +36,9 @@
- in `mathcomp_extra.v`:
+ lemmas `intr1D`, `intrD1`, `floor_eq`, `floorN`

- in `realfun.v`:
+ lemma `nondecreasing_at_left_is_cvgr`

### Changed

- in `topology.v`:
Expand Down
16 changes: 16 additions & 0 deletions theories/realfun.v
Original file line number Diff line number Diff line change
Expand Up @@ -365,6 +365,22 @@ by eexists; apply: (@nondecreasing_at_right_cvgr _ _ (BLeft b));
[rewrite bnd_simp|near: b..].
Unshelve. all: by end_near. Qed.

Lemma nondecreasing_at_left_is_cvgr f a :
(\forall x \near a^'-, {in `]x, a[ &, {homo f : n m / n <= m}}) ->
(\forall x \near a^'-, has_ubound [set f y | y in `]x, a[]) ->
cvg (f x @[x --> a^'-]).
Proof.
move=> ndf ubf; suff: cvg ((f \o -%R) x @[x --> (- a)^'+]).
move=> /cvg_ex[/= l fal].
by apply/cvg_ex; exists l; exact/cvg_at_leftNP.
apply: @nonincreasing_at_right_is_cvgr.
- rewrite at_rightN near_simpl; apply: filterS ndf => x ndf y z.
by rewrite -2!oppr_itvoo => yxa zxa yz; rewrite ndf// lerNr opprK.
- rewrite at_rightN near_simpl; apply: filterS ubf => x [r ubf].
exists r => _/= [s sax <-]; rewrite ubf//=; exists (- s) => //.
by rewrite oppr_itvoo.
Qed.

End fun_cvg_realType.
Arguments nondecreasing_at_right_cvgr {R f a} b.
Arguments nondecreasing_at_right_cvgr {R f a} b.
Expand Down

0 comments on commit 6a1c4a4

Please sign in to comment.