Skip to content

Commit

Permalink
prod notation for ereal and lemmas
Browse files Browse the repository at this point in the history
  • Loading branch information
affeldt-aist committed Dec 2, 2024
1 parent 61b7572 commit 1ee095f
Show file tree
Hide file tree
Showing 2 changed files with 39 additions and 0 deletions.
4 changes: 4 additions & 0 deletions CHANGELOG_UNRELEASED.md
Original file line number Diff line number Diff line change
Expand Up @@ -10,6 +10,10 @@
- in `normedtype.v`:
+ lemmas `continuous_within_itvcyP`, `continuous_within_itvNycP`

- in `constructive_ereal.v`:
+ notations `\prod` in scope ereal_scope
+ lemmas `prode_ge0`, `prode_fin_num`

### Changed

### Renamed
Expand Down
35 changes: 35 additions & 0 deletions reals/constructive_ereal.v
Original file line number Diff line number Diff line change
Expand Up @@ -46,6 +46,7 @@ From mathcomp Require Import mathcomp_extra signed.
(* (_ <= _)%E, (_ < _)%E, == comparison relations for extended reals *)
(* (_ >= _)%E, (_ > _)%E *)
(* (\sum_(i in A) f i)%E == bigop-like notation in scope %E *)
(* (\prod_(i in A) f i)%E == bigop-like notation in scope %E *)
(* maxe x y, mine x y == notation for the maximum/minimum of two *)
(* extended real numbers *)
(* ``` *)
Expand Down Expand Up @@ -551,6 +552,31 @@ Notation "\sum_ ( i 'in' A ) F" :=
Notation "\sum_ ( i 'in' A ) F" :=
(\big[+%E/0%E]_(i in A) F%E) : ereal_scope.

Notation "\prod_ ( i <- r | P ) F" :=
(\big[*%E/1%:E]_(i <- r | P%B) F%E) : ereal_scope.
Notation "\prod_ ( i <- r ) F" :=
(\big[*%E/1%:E]_(i <- r) F%E) : ereal_scope.
Notation "\prod_ ( m <= i < n | P ) F" :=
(\big[*%E/1%:E]_(m <= i < n | P%B) F%E) : ereal_scope.
Notation "\prod_ ( m <= i < n ) F" :=
(\big[*%E/1%:E]_(m <= i < n) F%E) : ereal_scope.
Notation "\prod_ ( i | P ) F" :=
(\big[*%E/1%:E]_(i | P%B) F%E) : ereal_scope.
Notation "\prod_ i F" :=
(\big[*%E/1%:E]_i F%E) : ereal_scope.
Notation "\prod_ ( i : t | P ) F" :=
(\big[*%E/1%:E]_(i : t | P%B) F%E) (only parsing) : ereal_scope.
Notation "\prod_ ( i : t ) F" :=
(\big[*%E/1%:E]_(i : t) F%E) (only parsing) : ereal_scope.
Notation "\prod_ ( i < n | P ) F" :=
(\big[*%E/1%:E]_(i < n | P%B) F%E) : ereal_scope.
Notation "\prod_ ( i < n ) F" :=
(\big[*%E/1%:E]_(i < n) F%E) : ereal_scope.
Notation "\prod_ ( i 'in' A | P ) F" :=
(\big[*%E/1%:E]_(i in A | P%B) F%E) : ereal_scope.
Notation "\prod_ ( i 'in' A ) F" :=
(\big[*%E/1%:E]_(i in A) F%E) : ereal_scope.

Section ERealOrderTheory.
Context {R : numDomainType}.
Implicit Types x y z : \bar R.
Expand Down Expand Up @@ -928,6 +954,11 @@ Lemma fin_numM x y : x \is a fin_num -> y \is a fin_num ->
x * y \is a fin_num.
Proof. by move: x y => [x| |] [y| |]. Qed.

Lemma prode_fin_num (I : Type) (s : seq I) (P : pred I) (f : I -> \bar R) :
(forall i, P i -> f i \is a fin_num) ->
\prod_(i <- s | P i) f i \is a fin_num.
Proof. by move=> ffin; elim/big_ind : _ => // x y x0 y0; rewrite fin_numM. Qed.

Lemma fin_numX x n : x \is a fin_num -> x ^+ n \is a fin_num.
Proof. by elim: n x => // n ih x finx; rewrite expeS fin_numM// ih. Qed.

Expand Down Expand Up @@ -1164,6 +1195,10 @@ move: x y => [x||] [y||]//=; rewrite /mule/= ?(lee_fin, eqe, lte_fin, lt0y)//.
by rewrite gt_eqF // y0 le0y.
Qed.

Lemma prode_ge0 (I : Type) (s : seq I) (P : pred I) (f : I -> \bar R) :
(forall i, P i -> 0 <= f i) -> 0 <= \prod_(i <- s | P i) f i.
Proof. by move=> f0; elim/big_ind : _ => // x y x0 y0; rewrite mule_ge0. Qed.

Lemma mule_gt0 x y : 0 < x -> 0 < y -> 0 < x * y.
Proof.
by rewrite !lt_def => /andP[? ?] /andP[? ?]; rewrite mule_neq0 ?mule_ge0.
Expand Down

0 comments on commit 1ee095f

Please sign in to comment.