Skip to content

Commit b513e81

Browse files
affeldt-aistproux01
authored andcommitted
minor generalizations, additions, fixes (#974)
- fixes #938
1 parent 9b8468d commit b513e81

File tree

7 files changed

+116
-80
lines changed

7 files changed

+116
-80
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -51,6 +51,9 @@
5151
- in file `lebesgue_integral.v`,
5252
+ new lemma `approximation_sfun_integrable`.
5353

54+
- in `classical_sets.v`:
55+
+ lemmas `properW`, `properxx`
56+
5457
### Changed
5558

5659
- moved from `lebesgue_measure.v` to `real_interval.v`:
@@ -119,6 +122,15 @@
119122
+ `powere_posM` -> `poweRM`
120123
+ `powere12_sqrt` -> `poweR12_sqrt`
121124

125+
- in `lebesgue_integral.v`:
126+
+ `ge0_integralM_EFin` -> `ge0_integralZl_EFin`
127+
+ `ge0_integralM` -> `ge0_integralZl`
128+
+ `integralM_indic` -> `integralZl_indic`
129+
+ `integralM_indic_nnsfun` -> `integralZl_indic_nnsfun`
130+
+ `integrablerM` -> `integrableZl`
131+
+ `integrableMr` -> `integrableZr`
132+
+ `integralM` -> `integralZl`
133+
122134
### Generalized
123135

124136
- in `exp.v`:
@@ -128,6 +140,11 @@
128140
+ lemma `ln_power_pos`
129141
- in file `lebesgue_integral.v`, updated `le_approx`.
130142

143+
- in `sequences.v`:
144+
+ lemmas `is_cvg_nneseries_cond`, `is_cvg_npeseries_cond`
145+
+ lemmas `is_cvg_nneseries`, `is_cvg_npeseries`
146+
+ lemmas `nneseries_ge0`, `npeseries_le0`
147+
131148
### Deprecated
132149

133150
### Removed

classical/classical_sets.v

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -60,6 +60,7 @@ From mathcomp Require Import mathcomp_extra boolp.
6060
(* \bigcap_i F == same as before with T left implicit. *)
6161
(* smallest C G := \bigcap_(A in [set M | C M /\ G `<=` M]) A *)
6262
(* A `<=` B <-> A is included in B. *)
63+
(* A `<` B := A `<=` B /\ ~ (B `<=` A) *)
6364
(* A `<=>` B <-> double inclusion A `<=` B and B `<=` A. *)
6465
(* f @^-1` A == preimage of A by f. *)
6566
(* f @` A == image of A by f. Notation for `image A f`. *)
@@ -531,6 +532,10 @@ Proof. by move=> sAB sBC ? ?; apply/sBC/sAB. Qed.
531532

532533
Lemma sub0set A : set0 `<=` A. Proof. by []. Qed.
533534

535+
Lemma properW A B : A `<` B -> A `<=` B. Proof. by case. Qed.
536+
537+
Lemma properxx A : ~ A `<` A. Proof. by move=> [?]; apply. Qed.
538+
534539
Lemma setC0 : ~` set0 = setT :> set T.
535540
Proof. by rewrite predeqE; split => ?. Qed.
536541

classical/mathcomp_extra.v

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -830,3 +830,18 @@ Reserved Notation "f \min g" (at level 50, left associativity).
830830
Definition min_fun T (R : numDomainType) (f g : T -> R) x := Num.min (f x) (g x).
831831
Notation "f \min g" := (min_fun f g) : ring_scope.
832832
Arguments min_fun {T R} _ _ _ /.
833+
834+
(* NB: Coq 8.17.0 generalizes dependent_choice from Set to Type
835+
making the following lemma redundant *)
836+
Section dependent_choice_Type.
837+
Context X (R : X -> X -> Prop).
838+
839+
Lemma dependent_choice_Type : (forall x, {y | R x y}) ->
840+
forall x0, {f | f 0%N = x0 /\ forall n, R (f n) (f n.+1)}.
841+
Proof.
842+
move=> h x0.
843+
set (f := fix f n := if n is n'.+1 then proj1_sig (h (f n')) else x0).
844+
exists f; split => //.
845+
intro n; induction n; simpl; apply: proj2_sig.
846+
Qed.
847+
End dependent_choice_Type.

theories/charge.v

Lines changed: 1 addition & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -69,21 +69,6 @@ Unset Printing Implicit Defensive.
6969
Import Order.TTheory GRing.Theory Num.Def Num.Theory.
7070
Import numFieldTopology.Exports.
7171

72-
(* NB: in the next releases of Coq, dependent_choice will be
73-
generalized from Set to Type making the following lemma redundant *)
74-
Section dependent_choice_Type.
75-
Context X (R : X -> X -> Prop).
76-
77-
Lemma dependent_choice_Type : (forall x, {y | R x y}) ->
78-
forall x0, {f | f 0 = x0 /\ forall n, R (f n) (f n.+1)}.
79-
Proof.
80-
move=> h x0.
81-
set (f := fix f n := if n is n'.+1 then proj1_sig (h (f n')) else x0).
82-
exists f; split => //.
83-
intro n; induction n; simpl; apply: proj2_sig.
84-
Qed.
85-
End dependent_choice_Type.
86-
8772
Local Open Scope ring_scope.
8873
Local Open Scope classical_set_scope.
8974
Local Open Scope ereal_scope.
@@ -723,7 +708,7 @@ move=> /cvg_ex[[l| |]]; first last.
723708
have : nu N <= -oo by rewrite -limNoo// nuN.
724709
by rewrite leNgt => /negP; apply; rewrite ltNye_eq fin_num_measure.
725710
- move/cvg_lim => limoo.
726-
have := @npeseries_le0 _ (fun n => maxe (z_ (v n) * 2^-1%:E) (- 1%E)) xpredT.
711+
have := @npeseries_le0 _ (fun n => maxe (z_ (v n) * 2^-1%:E) (- 1%E)) xpredT 0.
727712
by rewrite limoo// leNgt => /(_ (fun n _ => max_le0 n))/negP; apply.
728713
move/fine_cvgP => [Hfin cvgl].
729714
have : cvg (series (fun n => fine (maxe (z_ (v n) * 2^-1%:E) (- 1%E))) n @[n --> \oo]).

0 commit comments

Comments
 (0)