Module mathcomp.classical.mathcomp_extra
From Coq Require Import BinPos.From mathcomp Require Import all_ssreflect finmap ssralg ssrnum ssrint rat.
From mathcomp Require Import finset interval.
# MathComp extra
This files contains lemmas and definitions missing from MathComp.
```
proj i f == f i, where f : forall i, T i
dfwith f x == fun j => x if j = i, and f j otherwise
given x : T i
swap x := (x.2, x.1)
map_pair f x := (f x.1, f x.2)
monotonous A f := {in A &, {mono f : x y / x <= y}} \/
{in A &, {mono f : x y /~ x <= y}}
sigT_fun f := lifts a family of functions f into a function on
the dependent sum
prodA x := sends (X * Y) * Z to X * (Y * Z)
prodAr x := sends X * (Y * Z) to (X * Y) * Z
```
Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.
Import Order.TTheory GRing.Theory Num.Theory.
Local Open Scope ring_scope.
MathComp 2.2 additions
Notation "f \min g" := (Order.min_fun f g) : function_scope.
Notation "f \max g" := (Order.max_fun f g) : function_scope.
Lemma ler_sqrt {R : rcfType} (a b : R) :
(0 <= b -> (Num.sqrt a <= Num.sqrt b) = (a <= b))%R.
Proof.
MathComp 2.3 additions
Module Order.
Import Order.
Definition disp_t : Set.
Proof.
End Order.
Lemma invf_plt (F : numFieldType) :
{in Num.pos &, forall x y : F, (x^-1 < y)%R = (y^-1 < x)%R}.
Lemma invf_ltp (F : numFieldType) :
{in Num.pos &, forall x y : F, (x < y^-1)%R = (y < x^-1)%R}.
Lemma invf_ple (F : numFieldType) :
{in Num.pos &, forall x y : F, (x^-1 <= y)%R = (y^-1 <= x)%R}.
Lemma invf_lep (F : numFieldType) :
{in Num.pos &, forall x y : F, (x <= y^-1)%R = (y <= x^-1)%R}.
Definition proj {I} {T : I -> Type} i (f : forall i, T i) := f i.
Section DFunWith.
Variables (I : eqType) (T : I -> Type) (f : forall i, T i).
Definition dfwith i (x : T i) (j : I) : T j :=
if (i =P j) is ReflectT ij then ecast j (T j) ij x else f j.
Lemma dfwithin i x : dfwith x i = x.
Lemma dfwithout i (x : T i) j : i != j -> dfwith x j = f j.
Variant dfwith_spec i (x : T i) : forall j, T j -> Type :=
| DFunWithin : dfwith_spec x x
| DFunWithout j : i != j -> dfwith_spec x (f j).
Lemma dfwithP i (x : T i) (j : I) : dfwith_spec x (dfwith x j).
Lemma projK i (x : T i) : cancel (@dfwith i) (proj i).
End DFunWith.
Arguments dfwith {I T} f i x.
not yet backported
Definition idempotent_fun (U : Type) (f : U -> U) := f \o f =1 f.
From mathcomp Require Import poly.
Lemma deg_le2_ge0 (F : rcfType) (a b c : F) :
(forall x, 0 <= a * x ^+ 2 + b * x + c)%R -> (b ^+ 2 - 4%:R * a * c <= 0)%R.
Proof.
Section dependent_choice_Type.
Context X (R : X -> X -> Prop).
Lemma dependent_choice_Type : (forall x, {y | R x y}) ->
forall x0, {f | f 0%N = x0 /\ forall n, R (f n) (f n.+1)}.
Proof.
Section max_min.
Variable R : realFieldType.
Let nz2 : 2%:R != 0 :> R
Proof.
Lemma maxr_absE (x y : R) : Num.max x y = (x + y + `|x - y|) / 2%:R.
Proof.
Lemma minr_absE (x y : R) : Num.min x y = (x + y - `|x - y|) / 2%:R.
Proof.
End max_min.
Notation trivial := (ltac:(done)).
Section bigmax_seq.
Context d {T : orderType d} {x : T} {I : eqType}.
Variables (r : seq I) (i0 : I) (P : pred I).
Lemma le_bigmax_seq F :
i0 \in r -> P i0 -> (F i0 <= \big[Order.max/x]_(i <- r | P i) F i)%O.
Proof.
Lemma bigmax_sup_seq (m : T) (F : I -> T) :
i0 \in r -> P i0 -> (m <= F i0)%O ->
(m <= \big[Order.max/x]_(i <- r | P i) F i)%O.
Proof.
End bigmax_seq.
Arguments le_bigmax_seq {d T} x {I r} i0 P.
Lemma leq_ltn_expn m : exists n, (2 ^ n <= m.+1 < 2 ^ n.+1)%N.
Proof.
elim: m => [|m [n /andP[h1 h2]]]; first by exists O.
have [m2n|nm2] := ltnP m.+2 (2 ^ n.+1)%N.
by exists n; rewrite m2n andbT (leq_trans h1).
exists n.+1; rewrite nm2/= -addn1.
rewrite -[X in (_ <= X)%N]prednK ?expn_gt0// -[X in (_ <= X)%N]addn1 leq_add2r.
by rewrite (leq_trans h2)// -subn1 leq_subRL ?expn_gt0// add1n ltn_exp2l.
Qed.
have [m2n|nm2] := ltnP m.+2 (2 ^ n.+1)%N.
by exists n; rewrite m2n andbT (leq_trans h1).
exists n.+1; rewrite nm2/= -addn1.
rewrite -[X in (_ <= X)%N]prednK ?expn_gt0// -[X in (_ <= X)%N]addn1 leq_add2r.
by rewrite (leq_trans h2)// -subn1 leq_subRL ?expn_gt0// add1n ltn_exp2l.
Qed.
Definition monotonous d (T : porderType d) (pT : predType T) (A : pT) (f : T -> T) :=
{in A &, {mono f : x y / (x <= y)%O}} \/ {in A &, {mono f : x y /~ (x <= y)%O}}.
Section path_lt.
Context d {T : orderType d}.
Implicit Types (a b c : T) (s : seq T).
Lemma last_filterP a (P : pred T) s :
P a -> P (last a [seq x <- s | P x]).
Proof.
Lemma path_lt_filter0 a s : path <%O a s -> [seq x <- s | (x < a)%O] = [::].
Proof.
move=> /lt_path_min/allP sa; rewrite -(filter_pred0 s).
apply: eq_in_filter => x xs.
by apply/negbTE; have := sa _ xs; rewrite ltNge; apply: contra => /ltW.
Qed.
apply: eq_in_filter => x xs.
by apply/negbTE; have := sa _ xs; rewrite ltNge; apply: contra => /ltW.
Qed.
Lemma path_lt_filterT a s : path <%O a s -> [seq x <- s | (a < x)%O] = s.
Proof.
move=> /lt_path_min/allP sa; rewrite -[RHS](filter_predT s).
by apply: eq_in_filter => x xs; exact: sa.
Qed.
by apply: eq_in_filter => x xs; exact: sa.
Qed.
Lemma path_lt_head a b s : (a < b)%O -> path <%O b s -> path <%O a s.
Lemma path_lt_last_filter a b c s :
(a < c)%O -> (c < b)%O -> path <%O a s -> last a s = b ->
last c [seq x <- s | (c < x)%O] = b.
Proof.
elim/last_ind : s a b c => /= [|h t ih a b c ac cb].
move=> a b c ac cb _ ab.
by apply/eqP; rewrite eq_le (ltW cb) -ab (ltW ac).
rewrite rcons_path => /andP[ah ht]; rewrite last_rcons => tb.
by rewrite filter_rcons tb cb last_rcons.
Qed.
move=> a b c ac cb _ ab.
by apply/eqP; rewrite eq_le (ltW cb) -ab (ltW ac).
rewrite rcons_path => /andP[ah ht]; rewrite last_rcons => tb.
by rewrite filter_rcons tb cb last_rcons.
Qed.
Lemma path_lt_le_last a s : path <%O a s -> (a <= last a s)%O.
Proof.
End path_lt.
Arguments last_filterP {d T a} P s.
Lemma sumr_le0 (R : numDomainType) I (r : seq I) (P : pred I) (F : I -> R) :
(forall i, P i -> F i <= 0)%R -> (\sum_(i <- r | P i) F i <= 0)%R.
Inductive boxed T := Box of T.
Reserved Notation "`1- r" (format "`1- r", at level 2).
Reserved Notation "f \^-1" (at level 35, format "f \^-1").
Lemma fset_nat_maximum (X : choiceType) (A : {fset X})
(f : X -> nat) : A != fset0 ->
(exists i, i \in A /\ forall j, j \in A -> f j <= f i)%nat.
Proof.
Lemma image_nat_maximum n (f : nat -> nat) :
(exists i, i <= n /\ forall j, j <= n -> f j <= f i)%N.
Proof.
Lemma card_fset_sum1 (T : choiceType) (A : {fset T}) :
#|` A| = (\sum_(i <- A) 1)%N.
Proof.
Arguments big_rmcond {R idx op I r} P.
Arguments big_rmcond_in {R idx op I r} P.
Reserved Notation "`1- x" (format "`1- x", at level 2).
Section onem.
Variable R : numDomainType.
Implicit Types r : R.
Definition onem r := 1 - r.
Local Notation "`1- r" := (onem r).
Lemma onem0 : `1-0 = 1
Lemma onem1 : `1-1 = 0
Lemma onemK r : `1-(`1-r) = r
Proof.
Lemma add_onemK r : r + `1- r = 1.
Lemma onem_gt0 r : r < 1 -> 0 < `1-r
Proof.
Lemma onem_ge0 r : r <= 1 -> 0 <= `1-r.
Lemma onem_le1 r : 0 <= r -> `1-r <= 1.
Lemma onem_lt1 r : 0 < r -> `1-r < 1.
Lemma onemX_ge0 r n : 0 <= r -> r <= 1 -> 0 <= `1-(r ^+ n).
Proof.
Lemma onemX_lt1 r n : 0 < r -> `1-(r ^+ n) < 1.
Lemma onemD r s : `1-(r + s) = `1-r - s.
Lemma onemMr r s : s * `1-r = s - s * r.
Lemma onemM r s : `1-(r * s) = `1-r + `1-s - `1-r * `1-s.
Proof.
End onem.
Notation "`1- r" := (onem r) : ring_scope.
Lemma onemV (F : numFieldType) (x : F) : x != 0 -> `1-(x^-1) = (x - 1) / x.
Lemma lez_abs2 (a b : int) : 0 <= a -> a <= b -> (`|a| <= `|b|)%N.
Proof.
by case: a => //= n _; case: b. Qed.
Lemma ler_gtP (R : numFieldType) (x y : R) :
reflect (forall z, z > y -> x <= z) (x <= y).
Proof.
Lemma ler_ltP (R : numFieldType) (x y : R) :
reflect (forall z, z < x -> z <= y) (x <= y).
Proof.
Definition inv_fun T (R : unitRingType) (f : T -> R) x := (f x)^-1%R.
Notation "f \^-1" := (inv_fun f) : ring_scope.
Arguments inv_fun {T R} _ _ /.
Definition bound_side d (T : porderType d) (c : bool) (x : itv_bound T) :=
if x is BSide c' _ then c == c' else false.
Lemma real_ltr_distlC [R : numDomainType] [x y : R] (e : R) :
x - y \is Num.real -> (`|x - y| < e) = (x - e < y < x + e).
Proof.
Definition swap {T1 T2 : Type} (x : T1 * T2) := (x.2, x.1).
Section reassociate_products.
Context {X Y Z : Type}.
Definition prodA (xyz : (X * Y) * Z) : X * (Y * Z) :=
(xyz.1.1, (xyz.1.2, xyz.2)).
Definition prodAr (xyz : X * (Y * Z)) : (X * Y) * Z :=
((xyz.1, xyz.2.1), xyz.2.2).
Lemma prodAK : cancel prodA prodAr.
Proof.
by case; case. Qed.
Lemma prodArK : cancel prodAr prodA.
Proof.
by case => ? []. Qed.
End reassociate_products.
Lemma swapK {T1 T2 : Type} : cancel (@swap T1 T2) (@swap T2 T1).
Proof.
by case=> ? ?. Qed.
Definition map_pair {S U : Type} (f : S -> U) (x : (S * S)) : (U * U) :=
(f x.1, f x.2).
Section order_min.
Variables (d : Order.disp_t) (T : orderType d).
Lemma lt_min_lt (x y z : T) : (Order.min x z < Order.min y z)%O -> (x < y)%O.
Proof.
End order_min.
Section positive.
Lemma Pos_to_natE p : Pos.to_nat p = nat_of_pos p.
Proof.
End positive.
Lemma intrD1 {R : ringType} (i : int) : i%:~R + 1 = (i + 1)%:~R :> R.
Proof.
Lemma intr1D {R : ringType} (i : int) : 1 + i%:~R = (1 + i)%:~R :> R.
Proof.
From mathcomp Require Import archimedean.
Section floor_ceil.
Context {R : archiDomainType}.
Implicit Type x : R.
Lemma ge_floor x : (Num.floor x)%:~R <= x.
Lemma floor_ge_int x (z : int) : (z%:~R <= x) = (z <= Num.floor x).
Proof.
Lemma floor_lt_int x (z : int) : (x < z%:~R) = (Num.floor x < z).
Proof.
Lemma floor_ge0 x : (0 <= Num.floor x) = (0 <= x).
Proof.
Lemma floor_le0 x : (x < 1) = (Num.floor x <= 0).
Proof.
Lemma floor_lt0 x : (x < 0) = (Num.floor x < 0).
Proof.
Lemma lt_succ_floor x : x < (Num.floor x + 1)%:~R.
Proof.
Lemma floor_eq x m : (Num.floor x == m) = (m%:~R <= x < (m + 1)%:~R).
Proof.
Lemma floor_neq0 x : (Num.floor x != 0) = (x < 0) || (x >= 1).
Proof.
#[deprecated(since="mathcomp-analysis 1.3.0", note="use `Num.Theory.le_ceil` instead")]
Lemma ceil_ge x : x <= (Num.ceil x)%:~R.
#[deprecated(since="mathcomp-analysis 1.3.0", note="use `Num.Theory.ceil_le_int`")]
Lemma ceil_ge_int x (z : int) : (x <= z%:~R) = (Num.ceil x <= z).
Proof.
Lemma ceil_gt_int x (z : int) : (z%:~R < x) = (z < Num.ceil x).
Proof.
Lemma ceilN x : Num.ceil (- x) = - Num.floor x.
Lemma floorN x : Num.floor (- x) = - Num.ceil x.
Lemma ceil_ge0 x : (- 1 < x) = (0 <= Num.ceil x).
Lemma ceil_gt0 x : (0 < x) = (0 < Num.ceil x).
Proof.
Lemma ceil_le0 x : (x <= 0) = (Num.ceil x <= 0).
Proof.
Lemma abs_ceil_ge x : `|x| <= `|Num.ceil x|.+1%:R.
Proof.
End floor_ceil.
#[deprecated(since="mathcomp-analysis 1.3.0", note="renamed to `ceil_gt_int`")]
Notation ceil_lt_int := ceil_gt_int (only parsing).
Section bijection_forall.
Lemma bij_forall A B (f : A -> B) (P : B -> Prop) :
bijective f -> (forall y, P y) <-> (forall x, P (f x)).
Proof.
End bijection_forall.
Lemma and_prop_in (T : Type) (p : mem_pred T) (P Q : T -> Prop) :
{in p, forall x, P x /\ Q x} <->
{in p, forall x, P x} /\ {in p, forall x, Q x}.
Proof.
split=> [cnd|[cnd1 cnd2] x xin]; first by split=> x xin; case: (cnd x xin).
by split; [apply: cnd1 | apply: cnd2].
Qed.
by split; [apply: cnd1 | apply: cnd2].
Qed.
Lemma mem_inc_segment d (T : porderType d) (a b : T) (f : T -> T) :
{in `[a, b] &, {mono f : x y / (x <= y)%O}} ->
{homo f : x / x \in `[a, b] >-> x \in `[f a, f b]}.
Proof.
Lemma mem_dec_segment d (T : porderType d) (a b : T) (f : T -> T) :
{in `[a, b] &, {mono f : x y /~ (x <= y)%O}} ->
{homo f : x / x \in `[a, b] >-> x \in `[f b, f a]}.
Proof.
Definition sigT_fun {I : Type} {X : I -> Type} {T : Type}
(f : forall i, X i -> T) (x : {i & X i}) : T :=
(f (projT1 x) (projT2 x)).
Section FsetPartitions.
Variables T I : choiceType.
Implicit Types (x y z : T) (A B D X : {fset T}) (P Q : {fset {fset T}}).
Implicit Types (J : pred I) (F : I -> {fset T}).
Variables (R : Type) (idx : R) (op : Monoid.com_law idx).
Let rhs_cond P K E :=
(\big[op/idx]_(A <- P) \big[op/idx]_(x <- A | K x) E x)%fset.
Let rhs P E := (\big[op/idx]_(A <- P) \big[op/idx]_(x <- A) E x)%fset.
Lemma partition_disjoint_bigfcup (f : T -> R) (F : I -> {fset T})
(K : {fset I}) :
(forall i j, i \in K -> j \in K -> i != j -> [disjoint F i & F j])%fset ->
\big[op/idx]_(i <- \big[fsetU/fset0]_(x <- K) (F x)) f i =
\big[op/idx]_(k <- K) (\big[op/idx]_(i <- F k) f i).
Proof.
move=> disjF; pose P := [fset F i | i in K & F i != fset0]%fset.
have trivP : trivIfset P.
apply/trivIfsetP => _ _ /imfsetP[i iK ->] /imfsetP[j jK ->] neqFij.
move: iK; rewrite !inE/= => /andP[iK Fi0].
move: jK; rewrite !inE/= => /andP[jK Fj0].
by apply: (disjF _ _ iK jK); apply: contraNneq neqFij => ->.
have -> : (\bigcup_(i <- K) F i)%fset = fcover P.
apply/esym; rewrite /P fcover_imfset big_mkcond /=; apply eq_bigr => i _.
by case: ifPn => // /negPn/eqP.
rewrite big_trivIfset // /rhs big_imfset => [|i j iK /andP[jK notFj0] eqFij] /=.
rewrite big_filter big_mkcond; apply eq_bigr => i _.
by case: ifPn => // /negPn /eqP ->; rewrite big_seq_fset0.
move: iK; rewrite !inE/= => /andP[iK Fi0].
by apply: contraNeq (disjF _ _ iK jK) _; rewrite -fsetI_eq0 eqFij fsetIid.
Qed.
have trivP : trivIfset P.
apply/trivIfsetP => _ _ /imfsetP[i iK ->] /imfsetP[j jK ->] neqFij.
move: iK; rewrite !inE/= => /andP[iK Fi0].
move: jK; rewrite !inE/= => /andP[jK Fj0].
by apply: (disjF _ _ iK jK); apply: contraNneq neqFij => ->.
have -> : (\bigcup_(i <- K) F i)%fset = fcover P.
apply/esym; rewrite /P fcover_imfset big_mkcond /=; apply eq_bigr => i _.
by case: ifPn => // /negPn/eqP.
rewrite big_trivIfset // /rhs big_imfset => [|i j iK /andP[jK notFj0] eqFij] /=.
rewrite big_filter big_mkcond; apply eq_bigr => i _.
by case: ifPn => // /negPn /eqP ->; rewrite big_seq_fset0.
move: iK; rewrite !inE/= => /andP[iK Fi0].
by apply: contraNeq (disjF _ _ iK jK) _; rewrite -fsetI_eq0 eqFij fsetIid.
Qed.
End FsetPartitions.