Skip to content

Commit 4c30ca2

Browse files
committed
tentative definition of bigmaxr with bigop
1 parent eca06c7 commit 4c30ca2

File tree

1 file changed

+45
-32
lines changed

1 file changed

+45
-32
lines changed

theories/Rstruct.v

Lines changed: 45 additions & 32 deletions
Original file line numberDiff line numberDiff line change
@@ -496,66 +496,80 @@ Section bigmaxr.
496496
Context {R : realDomainType}.
497497

498498
(* bigop pour le max pour des listes non vides ? *)
499-
Definition bigmaxr (x0 : R) lr :=
500-
foldr Num.max (head x0 lr) (behead lr).
499+
Definition bigmaxr (r : R) s := \big[Num.max/head r s]_(i <- s) i.
501500

502501
Lemma bigmaxr_nil (x0 : R) : bigmaxr x0 [::] = x0.
503-
Proof. by rewrite /bigmaxr. Qed.
502+
Proof. by rewrite /bigmaxr /= big_nil. Qed.
504503

505504
Lemma bigmaxr_un (x0 x : R) : bigmaxr x0 [:: x] = x.
506-
Proof. by rewrite /bigmaxr. Qed.
505+
Proof. by rewrite /bigmaxr /= big_cons big_nil joinxx. Qed.
507506

508-
Lemma bigmaxr_cons (x0 x y : R) lr :
509-
bigmaxr x0 (x :: y :: lr) = Num.max x (bigmaxr x0 (y :: lr)).
507+
(* previous definition *)
508+
Lemma bigmaxrE (r : R) s : bigmaxr r s = foldr Num.max (head r s) (behead s).
510509
Proof.
511-
rewrite /bigmaxr /=; elim: lr => [/= | a lr /=]; first by rewrite joinC.
512-
set b := foldr _ _ _; set c := foldr _ _ _ => H.
513-
by rewrite [Num.max a b]joinC joinA H -joinA (joinC c a).
510+
rewrite (_ : bigmaxr _ _ = if s isn't h :: t then r else \big[Num.max/h]_(i <- s) i).
511+
case: s => // ? t; rewrite big_cons /bigmaxr.
512+
by elim: t => //= [|? ? <-]; [rewrite big_nil joinxx | rewrite big_cons joinCA].
513+
by case: s => //=; rewrite /bigmaxr big_nil.
514514
Qed.
515515

516-
Lemma bigmaxr_ler (x0 : R) (lr : seq R) i :
517-
(i < size lr)%N -> (nth x0 lr i) <= (bigmaxr x0 lr).
516+
Lemma bigrmax_dflt (x y : R) s : Num.max x (\big[Num.max/x]_(j <- y :: s) j) =
517+
Num.max x (\big[Num.max/y]_(i <- y :: s) i).
518518
Proof.
519-
case: lr i => [i | x lr]; first by rewrite nth_nil bigmaxr_nil lexx.
520-
elim: lr x => [x i /= | x lr /= ihlr y i i_size].
521-
by rewrite ltnS leqn0 => /eqP ->; rewrite nth0 bigmaxr_un /=.
522-
rewrite bigmaxr_cons /=; case: i i_size => [_ /= | i].
523-
by rewrite lexU lexx.
524-
rewrite ltnS /=; move/(ihlr x); move/(le_trans)=> H; apply: H.
525-
by rewrite lexU lexx orbT.
519+
elim: s => /= [|h t IH] in x y *.
520+
by rewrite !big_cons !big_nil joinxx joinCA joinxx joinC.
521+
by rewrite big_cons joinCA IH joinCA [in RHS]big_cons IH.
522+
Qed.
523+
524+
Lemma bigmaxr_cons (x0 x y : R) lr :
525+
bigmaxr x0 (x :: y :: lr) = Num.max x (bigmaxr x0 (y :: lr)).
526+
Proof. by rewrite [y :: lr]lock /bigmaxr /= -lock big_cons bigrmax_dflt. Qed.
527+
528+
Lemma bigmaxr_ler (x0 : R) s i :
529+
(i < size s)%N -> (nth x0 s i) <= (bigmaxr x0 s).
530+
Proof.
531+
rewrite /bigmaxr; elim: s i => // h t IH [_|i] /=.
532+
by rewrite big_cons /= lexU lexx.
533+
rewrite ltnS => ti; case: t => [|h' t] // in IH ti *.
534+
by rewrite big_cons bigrmax_dflt lexU orbC IH.
526535
Qed.
527536

528537
(* Compatibilité avec l'addition *)
529538
Lemma bigmaxr_addr (x0 : R) lr (x : R) :
530539
bigmaxr (x0 + x) (map (fun y : R => y + x) lr) = (bigmaxr x0 lr) + x.
531540
Proof.
532-
case: lr => [/= | y lr]; first by rewrite bigmaxr_nil.
533-
elim: lr y => [y | y lr ihlr z]; first by rewrite /= !bigmaxr_un.
534-
by rewrite map_cons !bigmaxr_cons ihlr addr_maxl.
541+
rewrite /bigmaxr; case: lr => [|h t]; first by rewrite !big_nil.
542+
elim: t h => /= [|h' t IH] h; first by rewrite ?(big_cons,big_nil) -addr_maxl.
543+
by rewrite [in RHS]big_cons bigrmax_dflt addr_maxl -IH big_cons bigrmax_dflt.
544+
Qed.
545+
546+
Lemma bigmaxr_mem (x0 : R) lr : (0 < size lr)%N -> bigmaxr x0 lr \in lr.
547+
Proof.
548+
rewrite /bigmaxr; case: lr => // h t _.
549+
elim: t => //= [|h' t IH] in h *; first by rewrite big_cons big_nil inE joinxx.
550+
rewrite big_cons bigrmax_dflt inE eq_le; case: lerP => /=.
551+
- by rewrite lexU lexx.
552+
- by rewrite ltxU ltxx => ?; rewrite join_r ?IH // ltW.
535553
Qed.
536554

537555
(* TODO: bigmaxr_morph? *)
538556
Lemma bigmaxr_mulr (A : finType) (s : seq A) (k : R) (x : A -> R) :
539557
0 <= k -> bigmaxr 0 (map (fun i => k * x i) s) = k * bigmaxr 0 (map x s).
540558
Proof.
541-
move=> k0; elim: s => /= [|h [//|h' t ih]]; first by rewrite bigmaxr_nil mulr0.
559+
move=> k0; elim: s => /= [|h [/=|h' t ih]].
560+
by rewrite bigmaxr_nil mulr0.
561+
by rewrite !bigmaxr_un.
542562
by rewrite bigmaxr_cons {}ih bigmaxr_cons maxr_pmulr.
543563
Qed.
544564

545565
Lemma bigmaxr_index (x0 : R) lr :
546566
(0 < size lr)%N -> (index (bigmaxr x0 lr) lr < size lr)%N.
547567
Proof.
548-
case: lr => [//= | x l _].
549-
elim: l x => [x | x lr]; first by rewrite bigmaxr_un /= eq_refl.
550-
move/(_ x); set z := bigmaxr _ _ => /= ihl y; rewrite bigmaxr_cons -/z.
551-
rewrite eq_sym eq_joinl.
552-
by case: (leP z y).
568+
rewrite /bigmaxr; case: lr => //= h t _; case: ifPn => // /negbTE H.
569+
move: (@bigmaxr_mem x0 (h :: t) isT).
570+
by rewrite ltnS index_mem inE /= eq_sym H.
553571
Qed.
554572

555-
Lemma bigmaxr_mem (x0 : R) lr :
556-
(0 < size lr)%N -> bigmaxr x0 lr \in lr.
557-
Proof. by move/(bigmaxr_index x0); rewrite index_mem. Qed.
558-
559573
Lemma bigmaxr_lerP (x0 : R) lr (x : R) :
560574
(0 < size lr)%N ->
561575
reflect (forall i, (i < size lr)%N -> (nth x0 lr i) <= x) ((bigmaxr x0 lr) <= x).
@@ -603,7 +617,6 @@ rewrite nth_index //.
603617
by apply: bigmaxr_mem; apply: (leq_trans _ i_size).
604618
Qed.
605619

606-
607620
(* bigop pour le max pour des listes non vides ? *)
608621
Definition bmaxrf n (f : {ffun 'I_n.+1 -> R}) :=
609622
bigmaxr (f ord0) (codom f).

0 commit comments

Comments
 (0)