diff options
| author | Assia Mahboubi | 2019-12-11 16:46:11 +0100 |
|---|---|---|
| committer | GitHub | 2019-12-11 16:46:11 +0100 |
| commit | bb8f291fc40668f987c8ea5cf3941980342e46b2 (patch) | |
| tree | 1a2abb3ee70e24b31ece51f6bc5b8a2ea248d6a2 /mathcomp | |
| parent | 732dc474f09c0231e2332cdecf99a3ed045cdd04 (diff) | |
| parent | 3f6aa286677f6cb0659300afd2b612b7bce20e73 (diff) | |
Merge pull request #270 from math-comp/experiment/order
Dispatching order and norm, and anticipating normed modules.
Diffstat (limited to 'mathcomp')
| -rw-r--r-- | mathcomp/Make | 1 | ||||
| -rw-r--r-- | mathcomp/algebra/intdiv.v | 28 | ||||
| -rw-r--r-- | mathcomp/algebra/interval.v | 52 | ||||
| -rw-r--r-- | mathcomp/algebra/rat.v | 58 | ||||
| -rw-r--r-- | mathcomp/algebra/ssralg.v | 10 | ||||
| -rw-r--r-- | mathcomp/algebra/ssrint.v | 165 | ||||
| -rw-r--r-- | mathcomp/algebra/ssrnum.v | 3317 | ||||
| -rw-r--r-- | mathcomp/character/character.v | 33 | ||||
| -rw-r--r-- | mathcomp/character/classfun.v | 28 | ||||
| -rw-r--r-- | mathcomp/character/inertia.v | 24 | ||||
| -rw-r--r-- | mathcomp/character/integral_char.v | 28 | ||||
| -rw-r--r-- | mathcomp/character/vcharacter.v | 31 | ||||
| -rw-r--r-- | mathcomp/field/algC.v | 64 | ||||
| -rw-r--r-- | mathcomp/field/algebraics_fundamentals.v | 142 | ||||
| -rw-r--r-- | mathcomp/field/algnum.v | 3 | ||||
| -rw-r--r-- | mathcomp/field/finfield.v | 12 | ||||
| -rw-r--r-- | mathcomp/ssreflect/Make | 1 | ||||
| -rw-r--r-- | mathcomp/ssreflect/all_ssreflect.v | 1 | ||||
| -rw-r--r-- | mathcomp/ssreflect/fintype.v | 18 | ||||
| -rw-r--r-- | mathcomp/ssreflect/order.v | 6058 | ||||
| -rw-r--r-- | mathcomp/ssreflect/prime.v | 29 | ||||
| -rw-r--r-- | mathcomp/ssreflect/ssrnat.v | 25 | ||||
| -rw-r--r-- | mathcomp/ssreflect/tuple.v | 12 |
23 files changed, 8370 insertions, 1770 deletions
diff --git a/mathcomp/Make b/mathcomp/Make index c4391ae..0a2c4a4 100644 --- a/mathcomp/Make +++ b/mathcomp/Make @@ -77,6 +77,7 @@ ssreflect/fingraph.v ssreflect/finset.v ssreflect/fintype.v ssreflect/generic_quotient.v +ssreflect/order.v ssreflect/path.v ssreflect/prime.v ssreflect/seq.v diff --git a/mathcomp/algebra/intdiv.v b/mathcomp/algebra/intdiv.v index 7663e63..2a485fd 100644 --- a/mathcomp/algebra/intdiv.v +++ b/mathcomp/algebra/intdiv.v @@ -1,7 +1,7 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path. -From mathcomp Require Import div choice fintype tuple finfun bigop prime. +From mathcomp Require Import div choice fintype tuple finfun bigop prime order. From mathcomp Require Import ssralg poly ssrnum ssrint rat matrix. From mathcomp Require Import polydiv finalg perm zmodp mxalgebra vector. @@ -46,7 +46,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GRing.Theory Num.Theory. +Import Order.TTheory GRing.Theory Num.Theory. Local Open Scope ring_scope. Definition divz (m d : int) := @@ -79,7 +79,7 @@ Proof. by case: d => // d; rewrite /divz /= mul1r. Qed. Lemma divzN m d : (m %/ - d)%Z = - (m %/ d)%Z. Proof. by case: m => n; rewrite /divz /= sgzN abszN mulNr. Qed. -Lemma divz_abs m d : (m %/ `|d|)%Z = (-1) ^+ (d < 0)%R * (m %/ d)%Z. +Lemma divz_abs (m d : int) : (m %/ `|d|)%Z = (-1) ^+ (d < 0)%R * (m %/ d)%Z. Proof. by rewrite {3}[d]intEsign !mulr_sign; case: ifP => -> //; rewrite divzN opprK. Qed. @@ -132,7 +132,7 @@ Qed. Lemma divzMDl q m d : d != 0 -> ((q * d + m) %/ d)%Z = q + (m %/ d)%Z. Proof. -rewrite neqr_lt -oppr_gt0 => nz_d. +rewrite neq_lt -oppr_gt0 => nz_d. wlog{nz_d} d_gt0: q d / d > 0; last case: d => // d in d_gt0 *. move=> IH; case/orP: nz_d => /IH// /(_ (- q)). by rewrite mulrNN !divzN -opprD => /oppr_inj. @@ -179,8 +179,8 @@ Lemma divzMpl p m d : p > 0 -> (p * m %/ (p * d) = m %/ d)%Z. Proof. case: p => // p p_gt0; wlog d_gt0: d / d > 0; last case: d => // d in d_gt0 *. by move=> IH; case/intP: d => [|d|d]; rewrite ?mulr0 ?divz0 ?mulrN ?divzN ?IH. -rewrite {1}(divz_eq m d) mulrDr mulrCA divzMDl ?mulf_neq0 ?gtr_eqF // addrC. -rewrite divz_small ?add0r // PoszM pmulr_rge0 ?modz_ge0 ?gtr_eqF //=. +rewrite {1}(divz_eq m d) mulrDr mulrCA divzMDl ?mulf_neq0 ?gt_eqF // addrC. +rewrite divz_small ?add0r // PoszM pmulr_rge0 ?modz_ge0 ?gt_eqF //=. by rewrite ltr_pmul2l ?ltz_pmod. Qed. Arguments divzMpl [p m d]. @@ -203,19 +203,20 @@ Qed. Lemma ltz_ceil m d : d > 0 -> m < ((m %/ d)%Z + 1) * d. Proof. -by case: d => // d d_gt0; rewrite mulrDl mul1r -ltr_subl_addl ltz_mod ?gtr_eqF. +by case: d => // d d_gt0; rewrite mulrDl mul1r -ltr_subl_addl ltz_mod ?gt_eqF. Qed. Lemma ltz_divLR m n d : d > 0 -> ((m %/ d)%Z < n) = (m < n * d). Proof. move=> d_gt0; apply/idP/idP. - by rewrite -lez_addr1 -(ler_pmul2r d_gt0); apply: ltr_le_trans (ltz_ceil _ _). -rewrite -(ltr_pmul2r d_gt0 _ n) //; apply: ler_lt_trans (lez_floor _ _). -by rewrite gtr_eqF. + by rewrite -[_ < n]lez_addr1 -(ler_pmul2r d_gt0); + apply: lt_le_trans (ltz_ceil _ _). +rewrite -(ltr_pmul2r d_gt0 _ n) //; apply: le_lt_trans (lez_floor _ _). +by rewrite gt_eqF. Qed. Lemma lez_divRL m n d : d > 0 -> (m <= (n %/ d)%Z) = (m * d <= n). -Proof. by move=> d_gt0; rewrite !lerNgt ltz_divLR. Qed. +Proof. by move=> d_gt0; rewrite !leNgt ltz_divLR. Qed. Lemma divz_ge0 m d : d > 0 -> ((m %/ d)%Z >= 0) = (m >= 0). Proof. by case: d m => // d [] n d_gt0; rewrite (divz_nat, divNz_nat). Qed. @@ -225,9 +226,9 @@ Proof. case: n => // [[|n]] _; first by rewrite mul0r !divz0 div0z. wlog p_gt0: p / p > 0; last case: p => // p in p_gt0 *. by case/intP: p => [|p|p] IH; rewrite ?mulr0 ?divz0 ?mulrN ?divzN // IH. -rewrite {2}(divz_eq m (n.+1%:Z * p)) mulrA mulrAC !divzMDl // ?gtr_eqF //. +rewrite {2}(divz_eq m (n.+1%:Z * p)) mulrA mulrAC !divzMDl // ?gt_eqF //. rewrite [rhs in _ + rhs]divz_small ?addr0 // ltz_divLR // divz_ge0 //. -by rewrite mulrC ltz_pmod ?modz_ge0 ?gtr_eqF ?pmulr_lgt0. +by rewrite mulrC ltz_pmod ?modz_ge0 ?gt_eqF ?pmulr_lgt0. Qed. Lemma modz_small m d : 0 <= m < d -> (m %% d)%Z = m. @@ -1073,4 +1074,3 @@ rewrite -defS -2!mulmxA; have ->: T *m pinvmx T = 1%:M. by apply: (row_free_inj uT); rewrite mul1mx mulmxKpV. by move=> i; rewrite mulmx1 -map_mxM 2!mxE denq_int mxE. Qed. - diff --git a/mathcomp/algebra/interval.v b/mathcomp/algebra/interval.v index eb0785f..3ed2825 100644 --- a/mathcomp/algebra/interval.v +++ b/mathcomp/algebra/interval.v @@ -1,7 +1,8 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice. -From mathcomp Require Import div fintype bigop ssralg finset fingroup ssrnum. +From mathcomp Require Import div fintype bigop order ssralg finset fingroup. +From mathcomp Require Import ssrnum. (*****************************************************************************) (* This file provide support for intervals in numerical and real domains. *) @@ -39,13 +40,14 @@ Unset Strict Implicit. Unset Printing Implicit Defensive. Local Open Scope ring_scope. -Import GRing.Theory Num.Theory. +Import Order.TTheory GRing.Theory Num.Theory. Local Notation mid x y := ((x + y) / 2%:R). Section LersifPo. Variable R : numDomainType. +Implicit Types (b : bool) (x y z : R). Definition lersif (x y : R) b := if b then x < y else x <= y. @@ -65,7 +67,7 @@ Lemma lersif_trans x y z b1 b2 : x <= y ?< if b1 -> y <= z ?< if b2 -> x <= z ?< if b1 || b2. Proof. by case: b1 b2 => [] []; - apply (ltr_trans, ltr_le_trans, ler_lt_trans, ler_trans). + [apply: lt_trans | apply: lt_le_trans | apply: le_lt_trans | apply: le_trans]. Qed. Lemma lersif01 b : 0 <= 1 ?< if b. @@ -74,16 +76,16 @@ Proof. by case: b; apply lter01. Qed. Lemma lersif_anti b1 b2 x y : (x <= y ?< if b1) && (y <= x ?< if b2) = if b1 || b2 then false else x == y. -Proof. by case: b1 b2 => [] []; rewrite lter_anti. Qed. +Proof. by case: b1 b2 => [] []; rewrite lte_anti. Qed. Lemma lersifxx x b : (x <= x ?< if b) = ~~ b. -Proof. by case: b; rewrite /= lterr. Qed. +Proof. by case: b; rewrite /= ltexx. Qed. Lemma lersifNF x y b : y <= x ?< if ~~ b -> x <= y ?< if b = false. -Proof. by case: b => /= [/ler_gtF|/ltr_geF]. Qed. +Proof. by case: b => /= [/le_gtF|/lt_geF]. Qed. Lemma lersifS x y b : x < y -> x <= y ?< if b. -Proof. by case: b => //= /ltrW. Qed. +Proof. by case: b => //= /ltW. Qed. Lemma lersifT x y : x <= y ?< if true = (x < y). Proof. by []. Qed. @@ -132,25 +134,25 @@ Definition lersif_sub_addl := (lersif_subl_addl, lersif_subr_addl). Lemma lersif_andb x y : {morph lersif x y : p q / p || q >-> p && q}. Proof. -by case=> [] [] /=; rewrite ?ler_eqVlt; +by case=> [] [] /=; rewrite ?le_eqVlt; case: (_ < _)%R; rewrite ?(orbT, orbF, andbF, andbb). Qed. Lemma lersif_orb x y : {morph lersif x y : p q / p && q >-> p || q}. Proof. -by case=> [] [] /=; rewrite ?ler_eqVlt; +by case=> [] [] /=; rewrite ?le_eqVlt; case: (_ < _)%R; rewrite ?(orbT, orbF, orbb). Qed. Lemma lersif_imply b1 b2 r1 r2 : b2 ==> b1 -> r1 <= r2 ?< if b1 -> r1 <= r2 ?< if b2. -Proof. by case: b1 b2 => [] [] //= _ /ltrW. Qed. +Proof. by case: b1 b2 => [] [] //= _ /ltW. Qed. Lemma lersifW b x y : x <= y ?< if b -> x <= y. -Proof. by case: b => // /ltrW. Qed. +Proof. by case: b => // /ltW. Qed. Lemma ltrW_lersif b x y : x < y -> x <= y ?< if b. -Proof. by case: b => // /ltrW. Qed. +Proof. by case: b => // /ltW. Qed. Lemma lersif_pmul2l b x : 0 < x -> {mono *%R x : y z / y <= z ?< if b}. Proof. by case: b; apply lter_pmul2l. Qed. @@ -166,7 +168,7 @@ Proof. by case: b; apply lter_nmul2r. Qed. Lemma real_lersifN x y b : x \is Num.real -> y \is Num.real -> x <= y ?< if ~~b = ~~ (y <= x ?< if b). -Proof. by case: b => [] xR yR /=; case: real_ltrgtP. Qed. +Proof. by case: b => [] xR yR /=; case: real_ltgtP. Qed. Lemma real_lersif_norml b x y : x \is Num.real -> @@ -208,19 +210,19 @@ Proof. by case: b; apply lter_distl. Qed. Lemma lersif_minr : (x <= Num.min y z ?< if b) = (x <= y ?< if b) && (x <= z ?< if b). -Proof. by case: b; apply lter_minr. Qed. +Proof. by case: b; rewrite /= ltexI. Qed. Lemma lersif_minl : (Num.min y z <= x ?< if b) = (y <= x ?< if b) || (z <= x ?< if b). -Proof. by case: b; apply lter_minl. Qed. +Proof. by case: b; rewrite /= lteIx. Qed. Lemma lersif_maxr : (x <= Num.max y z ?< if b) = (x <= y ?< if b) || (x <= z ?< if b). -Proof. by case: b; apply lter_maxr. Qed. +Proof. by case: b; rewrite /= ltexU. Qed. Lemma lersif_maxl : (Num.max y z <= x ?< if b) = (y <= x ?< if b) && (z <= x ?< if b). -Proof. by case: b; apply lter_maxl. Qed. +Proof. by case: b; rewrite /= lteUx. Qed. End LersifOrdered. @@ -459,13 +461,9 @@ Definition bound_in_itv := (boundl_in_itv, boundr_in_itv). Lemma itvP : forall (x : R) (i : interval R), x \in i -> itv_rewrite i x. Proof. -move=> x [[[] a|] [[] b|]] /itv_dec // [? ?]; - do ?split => //; rewrite ?bound_in_itv /le_boundl /le_boundr //=; - do 1?[apply/negbTE; rewrite (ler_gtF, ltr_geF) //]; - by [ rewrite ltrW - | rewrite (@ler_trans _ x) // 1?ltrW - | rewrite (@ltr_le_trans _ x) - | rewrite (@ler_lt_trans _ x) // 1?ltrW ]. +move=> x [[[] a|] [[] b|]] /itv_dec [ha hb]; do !split; + rewrite ?bound_in_itv //=; do 1?[apply/negbTE; rewrite (le_gtF, lt_geF)]; + by [ | apply: ltW | move: (lersif_trans ha hb) => //=; exact: ltW ]. Qed. Arguments itvP [x i]. @@ -484,7 +482,7 @@ Definition itv_intersectioni1 : right_id `]-oo, +oo[ itv_intersection. Proof. by case=> [[lb lr |] [ub ur |]]. Qed. Lemma itv_intersectionii : idempotent itv_intersection. -Proof. by case=> [[[] lr |] [[] ur |]] //=; rewrite !lerr. Qed. +Proof. by case=> [[[] lr |] [[] ur |]] //=; rewrite !lexx. Qed. Definition subitv (i1 i2 : interval R) := match i1, i2 with @@ -558,7 +556,7 @@ Lemma le_boundl_total : total (@le_boundl R). Proof. by move=> [[] l |] [[] r |] //=; case: (ltrgtP l r). Qed. Lemma le_boundr_total : total (@le_boundr R). -Proof. by move=> [[] l |] [[] r |] //=; case (ltrgtP l r). Qed. +Proof. by move=> [[] l |] [[] r |] //=; case: (ltrgtP l r). Qed. Lemma itv_splitU (xc : R) bc a b : xc \in Interval a b -> forall y, y \in Interval a b = @@ -626,7 +624,7 @@ Variable R : realFieldType. Lemma mid_in_itv : forall ba bb (xa xb : R), xa <= xb ?< if ba || bb -> mid xa xb \in Interval (BOpen_if ba xa) (BOpen_if bb xb). Proof. -by move=> [] [] xa xb /= ?; apply/itv_dec=> /=; rewrite ?midf_lte // ?ltrW. +by move=> [] [] xa xb /= ?; apply/itv_dec=> /=; rewrite ?midf_lte // ?ltW. Qed. Lemma mid_in_itvoo : forall (xa xb : R), xa < xb -> mid xa xb \in `]xa, xb[. diff --git a/mathcomp/algebra/rat.v b/mathcomp/algebra/rat.v index 228a824..14b5035 100644 --- a/mathcomp/algebra/rat.v +++ b/mathcomp/algebra/rat.v @@ -1,7 +1,8 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice. -From mathcomp Require Import fintype bigop ssralg countalg div ssrnum ssrint. +From mathcomp Require Import fintype bigop order ssralg countalg div ssrnum. +From mathcomp Require Import ssrint. (******************************************************************************) (* This file defines a datatype for rational numbers and equips it with a *) @@ -18,8 +19,7 @@ From mathcomp Require Import fintype bigop ssralg countalg div ssrnum ssrint. (* ratr x == generic embedding of (r : R) into an arbitrary unitring. *) (******************************************************************************) -Import GRing.Theory. -Import Num.Theory. +Import Order.TTheory GRing.Theory Num.Theory. Set Implicit Arguments. Unset Strict Implicit. @@ -51,16 +51,16 @@ Canonical rat_subCountType := [subCountType of rat]. Definition numq x := nosimpl ((valq x).1). Definition denq x := nosimpl ((valq x).2). -Lemma denq_gt0 x : 0 < denq x. +Lemma denq_gt0 x : 0 < denq x. Proof. by rewrite /denq; case: x=> [[a b] /= /andP []]. Qed. Hint Resolve denq_gt0 : core. -Definition denq_ge0 x := ltrW (denq_gt0 x). +Definition denq_ge0 x := ltW (denq_gt0 x). -Lemma denq_lt0 x : (denq x < 0) = false. Proof. by rewrite ltr_gtF. Qed. +Lemma denq_lt0 x : (denq x < 0) = false. Proof. by rewrite lt_gtF. Qed. Lemma denq_neq0 x : denq x != 0. -Proof. by rewrite /denq gtr_eqF ?denq_gt0. Qed. +Proof. by rewrite /denq gt_eqF ?denq_gt0. Qed. Hint Resolve denq_neq0 : core. Lemma denq_eq0 x : (denq x == 0) = false. @@ -99,7 +99,7 @@ Fact valqK x : fracq (valq x) = x. Proof. move: x => [[n d] /= Pnd]; apply: val_inj=> /=. move: Pnd; rewrite /coprime /fracq /= => /andP[] hd -/eqP hnd. -by rewrite ltr_gtF ?gtr_eqF //= hnd !divn1 mulz_sign_abs abszE gtr0_norm. +by rewrite lt_gtF ?gt_eqF //= hnd !divn1 mulz_sign_abs abszE gtr0_norm. Qed. Fact scalq_key : unit. Proof. by []. Qed. @@ -343,7 +343,7 @@ Canonical rat_comUnitRing := Eval hnf in [comUnitRingType of rat]. Fact rat_field_axiom : GRing.Field.mixin_of rat_unitRing. Proof. exact. Qed. Definition RatFieldIdomainMixin := (FieldIdomainMixin rat_field_axiom). -Canonical rat_iDomain := +Canonical rat_idomainType := Eval hnf in IdomainType rat (FieldIdomainMixin rat_field_axiom). Canonical rat_fieldType := FieldType rat rat_field_axiom. @@ -396,13 +396,13 @@ Lemma rat1 : 1%:Q = 1. Proof. by []. Qed. Lemma numqN x : numq (- x) = - numq x. Proof. rewrite /numq; case: x=> [[a b] /= /andP [hb]]; rewrite /coprime=> /eqP hab. -by rewrite ltr_gtF ?gtr_eqF // {2}abszN hab divn1 mulz_sign_abs. +by rewrite lt_gtF ?gt_eqF // {2}abszN hab divn1 mulz_sign_abs. Qed. Lemma denqN x : denq (- x) = denq x. Proof. rewrite /denq; case: x=> [[a b] /= /andP [hb]]; rewrite /coprime=> /eqP hab. -by rewrite gtr_eqF // abszN hab divn1 gtz0_abs. +by rewrite gt_eqF // abszN hab divn1 gtz0_abs. Qed. (* Will be subsumed by pnatr_eq0 *) @@ -502,7 +502,7 @@ Proof. rewrite !ge_rat0 => hnx hny. have hxy: (0 <= numq x * denq y + numq y * denq x). by rewrite addr_ge0 ?mulr_ge0. -by rewrite /numq /= -!/(denq _) ?mulf_eq0 ?denq_eq0 !ler_gtF ?mulr_ge0. +by rewrite /numq /= -!/(denq _) ?mulf_eq0 ?denq_eq0 !le_gtF ?mulr_ge0. Qed. Fact le_rat0M x y : le_rat 0 x -> le_rat 0 y -> le_rat 0 (x * y). @@ -510,12 +510,12 @@ Proof. rewrite !ge_rat0 => hnx hny. have hxy: (0 <= numq x * denq y + numq y * denq x). by rewrite addr_ge0 ?mulr_ge0. -by rewrite /numq /= -!/(denq _) ?mulf_eq0 ?denq_eq0 !ler_gtF ?mulr_ge0. +by rewrite /numq /= -!/(denq _) ?mulf_eq0 ?denq_eq0 !le_gtF ?mulr_ge0. Qed. Fact le_rat0_anti x : le_rat 0 x -> le_rat x 0 -> x = 0. Proof. -by move=> hx hy; apply/eqP; rewrite -numq_eq0 eqr_le -ge_rat0 -le_rat0 hx hy. +by move=> hx hy; apply/eqP; rewrite -numq_eq0 eq_le -ge_rat0 -le_rat0 hx hy. Qed. Lemma sgr_numq_div (n d : int) : sgr (numq (n%:Q / d%:Q)) = sgr n * sgr d. @@ -530,12 +530,12 @@ Proof. symmetry; rewrite ge_rat0 /le_rat -subr_ge0. case: ratP => nx dx cndx; case: ratP => ny dy cndy. rewrite -!mulNr addf_div ?intq_eq0 // !mulNr -!rmorphM -rmorphB /=. -symmetry; rewrite !lerNgt -sgr_cp0 sgr_numq_div mulrC gtr0_sg //. +symmetry; rewrite !leNgt -sgr_cp0 sgr_numq_div mulrC gtr0_sg //. by rewrite mul1r sgr_cp0. Qed. Fact le_rat_total : total le_rat. -Proof. by move=> x y; apply: ler_total. Qed. +Proof. by move=> x y; apply: le_total. Qed. Fact numq_sign_mul (b : bool) x : numq ((-1) ^+ b * x) = (-1) ^+ b * numq x. Proof. by case: b; rewrite ?(mul1r, mulN1r) // numqN. Qed. @@ -566,14 +566,19 @@ by rewrite /normq /= normr_num_div ?ger0_norm // divq_num_den. Qed. Fact lt_rat_def x y : (lt_rat x y) = (y != x) && (le_rat x y). -Proof. by rewrite /lt_rat ltr_def rat_eq. Qed. +Proof. by rewrite /lt_rat lt_def rat_eq. Qed. -Definition ratLeMixin := RealLeMixin le_rat0D le_rat0M le_rat0_anti - subq_ge0 (@le_rat_total 0) norm_ratN ge_rat0_norm lt_rat_def. +Definition ratLeMixin : realLeMixin rat_idomainType := + RealLeMixin le_rat0D le_rat0M le_rat0_anti subq_ge0 + (@le_rat_total 0) norm_ratN ge_rat0_norm lt_rat_def. +Canonical rat_porderType := POrderType ring_display rat ratLeMixin. +Canonical rat_distrLatticeType := DistrLatticeType rat ratLeMixin. +Canonical rat_orderType := OrderType rat le_rat_total. Canonical rat_numDomainType := NumDomainType rat ratLeMixin. +Canonical rat_normedZmodType := NormedZmodType rat rat ratLeMixin. Canonical rat_numFieldType := [numFieldType of rat]. -Canonical rat_realDomainType := RealDomainType rat (@le_rat_total 0). +Canonical rat_realDomainType := [realDomainType of rat]. Canonical rat_realFieldType := [realFieldType of rat]. Lemma numq_ge0 x : (0 <= numq x) = (0 <= x). @@ -585,10 +590,10 @@ Lemma numq_le0 x : (numq x <= 0) = (x <= 0). Proof. by rewrite -oppr_ge0 -numqN numq_ge0 oppr_ge0. Qed. Lemma numq_gt0 x : (0 < numq x) = (0 < x). -Proof. by rewrite !ltrNge numq_le0. Qed. +Proof. by rewrite !ltNge numq_le0. Qed. Lemma numq_lt0 x : (numq x < 0) = (x < 0). -Proof. by rewrite !ltrNge numq_ge0. Qed. +Proof. by rewrite !ltNge numq_ge0. Qed. Lemma sgr_numq x : sgz (numq x) = sgz x. Proof. @@ -605,7 +610,7 @@ Proof. by rewrite normrEsign denq_mulr_sign. Qed. Fact rat_archimedean : Num.archimedean_axiom [numDomainType of rat]. Proof. move=> x; exists `|numq x|.+1; rewrite mulrS ltr_spaddl //. -rewrite pmulrn abszE intr_norm numqE normrM ler_pemulr ?norm_ge0 //. +rewrite pmulrn abszE intr_norm numqE normrM ler_pemulr //. by rewrite -intr_norm ler1n absz_gt0 denq_eq0. Qed. @@ -758,7 +763,7 @@ by rewrite ![_ / _ * _]mulrAC !ler_pdivr_mulr ?ltr0z // -!rmorphM /= !ler_int. Qed. Lemma ltr_rat : {mono (@ratr F) : x y / x < y}. -Proof. exact: lerW_mono ler_rat. Qed. +Proof. exact: leW_mono ler_rat. Qed. Lemma ler0q x : (0 <= ratr F x) = (0 <= x). Proof. by rewrite (_ : 0 = ratr F 0) ?ler_rat ?rmorph0. Qed. @@ -777,8 +782,7 @@ Proof. by rewrite !sgr_def fmorph_eq0 ltrq0 rmorphMn rmorph_sign. Qed. Lemma ratr_norm x : ratr F `|x| = `|ratr F x|. Proof. -rewrite {2}[x]numEsign rmorphMsign normrMsign [`|ratr F _|]ger0_norm //. -by rewrite ler0q ?normr_ge0. +by rewrite {2}[x]numEsign rmorphMsign normrMsign [`|ratr F _|]ger0_norm ?ler0q. Qed. End InPrealField. @@ -807,7 +811,7 @@ Qed. Require setoid_ring.Field_theory setoid_ring.Field_tac. -Lemma rat_field_theory : +Lemma rat_field_theory : Field_theory.field_theory 0%Q 1%Q addq mulq subq oppq divq invq eq. Proof. split => //; first exact rat_ring_theory. diff --git a/mathcomp/algebra/ssralg.v b/mathcomp/algebra/ssralg.v index 4b1d85c..69fc44e 100644 --- a/mathcomp/algebra/ssralg.v +++ b/mathcomp/algebra/ssralg.v @@ -1700,7 +1700,7 @@ Definition clone c of phant_id class c := @Pack phR T c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). -Definition pack T b0 mul0 (axT : @axiom R (@Lmodule.Pack R _ T b0) mul0) := +Definition pack b0 mul0 (axT : @axiom R (@Lmodule.Pack R _ T b0) mul0) := fun bT b & phant_id (Ring.class bT) (b : Ring.class_of T) => fun mT m & phant_id (@Lmodule.class R phR mT) (@Lmodule.Class R T b m) => fun ax & phant_id axT ax => @@ -3761,6 +3761,9 @@ Variables (subS : zmodPred S) (kS : keyed_pred subS). Lemma rpredB : {in kS &, forall u v, u - v \in kS}. Proof. by move=> u v Su Sv; rewrite /= rpredD ?rpredN. Qed. +Lemma rpredBC u v : u - v \in kS = (v - u \in kS). +Proof. by rewrite -rpredN opprB. Qed. + Lemma rpredMNn n : {in kS, forall u, u *- n \in kS}. Proof. by move=> u Su; rewrite /= rpredN rpredMn. Qed. @@ -5800,6 +5803,7 @@ Definition rpred_sum := rpred_sum. Definition rpredMn := rpredMn. Definition rpredN := rpredN. Definition rpredB := rpredB. +Definition rpredBC := rpredBC. Definition rpredMNn := rpredMNn. Definition rpredDr := rpredDr. Definition rpredDl := rpredDl. @@ -6454,6 +6458,10 @@ Canonical pair_comUnitRingType (R1 R2 : comUnitRingType) := Canonical pair_unitAlgType (R : comUnitRingType) (A1 A2 : unitAlgType R) := Eval hnf in [unitAlgType R of A1 * A2]. +Lemma pairMnE (M1 M2 : zmodType) (x : M1 * M2) n : + x *+ n = (x.1 *+ n, x.2 *+ n). +Proof. by case: x => x y; elim: n => //= n; rewrite !mulrS => ->. Qed. + (* begin hide *) (* Testing subtype hierarchy diff --git a/mathcomp/algebra/ssrint.v b/mathcomp/algebra/ssrint.v index 2a17a4a..feaa884 100644 --- a/mathcomp/algebra/ssrint.v +++ b/mathcomp/algebra/ssrint.v @@ -1,7 +1,8 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat choice seq. -From mathcomp Require Import fintype finfun bigop ssralg countalg ssrnum poly. +From mathcomp Require Import fintype finfun bigop order ssralg countalg ssrnum. +From mathcomp Require Import poly. (******************************************************************************) (* This file develops a basic theory of signed integers, defining: *) @@ -39,7 +40,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GRing.Theory Num.Theory. +Import Order.TTheory GRing.Theory Num.Theory. Delimit Scope int_scope with Z. Local Open Scope int_scope. @@ -351,7 +352,7 @@ End intUnitRing. Canonical int_unitRingType := Eval hnf in UnitRingType int intUnitRing.comMixin. Canonical int_comUnitRing := Eval hnf in [comUnitRingType of int]. -Canonical int_iDomain := +Canonical int_idomainType := Eval hnf in IdomainType int intUnitRing.idomain_axiomz. Canonical int_countZmodType := [countZmodType of int]. @@ -390,26 +391,14 @@ Definition ltz m n := | Negz m', Negz n' => (n' < m')%N end. -Fact lez_norm_add x y : lez (normz (x + y)) (normz x + normz y). -Proof. -move: x y=> [] m [] n; rewrite /= ?addnS //=; -rewrite /GRing.add /GRing.Zmodule.add /=; case: ltnP=> //=; -rewrite ?addSn ?ltnS ?leq_subLR ?(addnS, addSn) ?(leq_trans _ (leqnSn _)) //; -by rewrite 1?addnCA ?leq_addr ?addnA ?leq_addl. -Qed. - -Fact ltz_add x y : ltz 0 x -> ltz 0 y -> ltz 0 (x + y). -Proof. by move: x y => [] x [] y //= hx hy; rewrite ltn_addr. Qed. - -Fact eq0_normz x : normz x = 0 -> x = 0. Proof. by case: x. Qed. +Fact lez_add m n : lez 0 m -> lez 0 n -> lez 0 (m + n). +Proof. by case: m n => [] m [] n. Qed. -Fact lez_total x y : lez x y || lez y x. -Proof. by move: x y => [] x [] y //=; apply: leq_total. Qed. +Fact lez_mul m n : lez 0 m -> lez 0 n -> lez 0 (m * n). +Proof. by case: m n => [] m [] n. Qed. -Lemma abszN (n : nat) : absz (- n%:Z) = n. Proof. by case: n. Qed. - -Fact normzM : {morph (fun n => normz n) : x y / x * y}. -Proof. by move=> [] x [] y; rewrite // abszN // mulnC. Qed. +Fact lez_anti m : lez 0 m -> lez m 0 -> m = 0. +Proof. by case: m; first case. Qed. Lemma subz_ge0 m n : lez 0 (n - m) = lez m n. Proof. @@ -420,23 +409,33 @@ by [ rewrite subzn // move: hmn; rewrite -subn_gt0; case: (_ - _)%N]. Qed. -Fact lez_def x y : (lez x y) = (normz (y - x) == y - x). -Proof. by rewrite -subz_ge0; move: (_ - _) => [] n //=; rewrite eqxx. Qed. +Fact lez_total m n : lez m n || lez n m. +Proof. by move: m n => [] m [] n //=; apply: leq_total. Qed. + +Fact normzN m : normz (- m) = normz m. +Proof. by case: m => // -[]. Qed. + +Fact gez0_norm m : lez 0 m -> normz m = m. +Proof. by case: m. Qed. -Fact ltz_def x y : (ltz x y) = (y != x) && (lez x y). +Fact ltz_def m n : (ltz m n) = (n != m) && (lez m n). Proof. -by move: x y=> [] x [] y //=; rewrite (ltn_neqAle, leq_eqVlt) // eq_sym. +by move: m n => [] m [] n //=; rewrite (ltn_neqAle, leq_eqVlt) // eq_sym. Qed. -Definition Mixin := - NumMixin lez_norm_add ltz_add eq0_normz (in2W lez_total) normzM - lez_def ltz_def. +Definition Mixin : realLeMixin int_idomainType := + RealLeMixin + lez_add lez_mul lez_anti subz_ge0 (lez_total 0) normzN gez0_norm ltz_def. End intOrdered. End intOrdered. +Canonical int_porderType := POrderType ring_display int intOrdered.Mixin. +Canonical int_distrLatticeType := DistrLatticeType int intOrdered.Mixin. +Canonical int_orderType := OrderType int intOrdered.lez_total. Canonical int_numDomainType := NumDomainType int intOrdered.Mixin. -Canonical int_realDomainType := RealDomainType int (intOrdered.lez_total 0). +Canonical int_normedZmodType := NormedZmodType int int intOrdered.Mixin. +Canonical int_realDomainType := [realDomainType of int]. Section intOrderedTheory. @@ -448,7 +447,7 @@ Lemma lez_nat m n : (m <= n :> int) = (m <= n)%N. Proof. by []. Qed. Lemma ltz_nat m n : (m < n :> int) = (m < n)%N. -Proof. by rewrite ltnNge ltrNge lez_nat. Qed. +Proof. by rewrite ltnNge ltNge lez_nat. Qed. Definition ltez_nat := (lez_nat, ltz_nat). @@ -805,7 +804,7 @@ Lemma ler_pmulz2r n (hn : 0 < n) : {mono *~%R^~ n :x y / x <= y :> R}. Proof. by move=> x y; case: n hn=> [[]|] // n _; rewrite ler_pmuln2r. Qed. Lemma ltr_pmulz2r n (hn : 0 < n) : {mono *~%R^~ n : x y / x < y :> R}. -Proof. exact: lerW_mono (ler_pmulz2r _). Qed. +Proof. exact: leW_mono (ler_pmulz2r _). Qed. Lemma ler_nmulz2r n (hn : n < 0) : {mono *~%R^~ n : x y /~ x <= y :> R}. Proof. @@ -814,7 +813,7 @@ by rewrite ler_pmulz2r (oppr_cp0, ler_opp2). Qed. Lemma ltr_nmulz2r n (hn : n < 0) : {mono *~%R^~ n : x y /~ x < y :> R}. -Proof. exact: lerW_nmono (ler_nmulz2r _). Qed. +Proof. exact: leW_nmono (ler_nmulz2r _). Qed. Lemma ler_wpmulz2r n (hn : 0 <= n) : {homo *~%R^~ n : x y / x <= y :> R}. Proof. by move=> x y xy; case: n hn=> [] // n _; rewrite ler_wmuln2r. Qed. @@ -883,10 +882,10 @@ by move=> m n /= hmn; rewrite -subr_gt0 -mulrzBr nmulrz_lgt0 // subr_lt0. Qed. Lemma ltr_pmulz2l x (hx : 0 < x) : {mono *~%R x : x y / x < y}. -Proof. exact: lerW_mono (ler_pmulz2l _). Qed. +Proof. exact: leW_mono (ler_pmulz2l _). Qed. Lemma ltr_nmulz2l x (hx : x < 0) : {mono *~%R x : x y /~ x < y}. -Proof. exact: lerW_nmono (ler_nmulz2l _). Qed. +Proof. exact: leW_nmono (ler_nmulz2l _). Qed. Lemma pmulrz_rgt0 x n (x0 : 0 < x) : 0 < x *~ n = (0 < n). Proof. by rewrite -(mulr0z x) ltr_pmulz2l. Qed. @@ -915,7 +914,7 @@ Proof. by rewrite -(mulr0z x) ler_nmulz2l. Qed. Lemma mulrIz x (hx : x != 0) : injective ( *~%R x). Proof. move=> y z; rewrite -![x *~ _]mulrzr => /(mulfI hx). -by apply: incr_inj y z; apply: ler_pmulz2l. +by apply: inc_inj y z; apply: ler_pmulz2l. Qed. Lemma ler_int m n : (m%:~R <= n%:~R :> R) = (m <= n). @@ -961,7 +960,7 @@ Lemma mulrz_neq0 x n : x *~ n != 0 = ((n != 0) && (x != 0)). Proof. by rewrite mulrz_eq0 negb_or. Qed. Lemma realz n : (n%:~R : R) \in Num.real. -Proof. by rewrite -topredE /Num.real /= ler0z lerz0 ler_total. Qed. +Proof. by rewrite -topredE /Num.real /= ler0z lerz0 le_total. Qed. Hint Resolve realz : core. Definition intr_inj := @mulrIz 1 (oner_neq0 R). @@ -1070,7 +1069,7 @@ Qed. Lemma exprzDr x (ux : x \is a GRing.unit) m n : x ^ (m + n) = x ^ m * x ^ n. Proof. -move: n m; apply: wlog_ler=> n m hnm. +move: n m; apply: wlog_le=> n m hnm. by rewrite addrC hnm commrXz //; apply: commr_sym; apply: commrXz. case: (intP m) hnm=> {m} [|m|m]; rewrite ?mul1r ?add0r //; case: (intP n)=> {n} [|n|n _]; rewrite ?mulr1 ?addr0 //; @@ -1244,17 +1243,17 @@ Fact ler_wneexpz2l x (x1 : 1 <= x) : {in <= 0 &, {homo (exprz x) : x y / x <= y}}. Proof. move=> m n hm hn /= hmn. -rewrite -lef_pinv -?topredE /= ?exprz_gt0 ?(ltr_le_trans ltr01) //. +rewrite -lef_pinv -?topredE /= ?exprz_gt0 ?(lt_le_trans ltr01) //. by rewrite !invr_expz ler_wpeexpz2l ?ler_opp2 -?topredE //= oppr_cp0. Qed. Lemma ler_weexpz2l x (x1 : 1 <= x) : {homo (exprz x) : x y / x <= y}. Proof. -move=> m n /= hmn; case: (lerP 0 m)=> [|/ltrW] hm. - by rewrite ler_wpeexpz2l // [_ \in _](ler_trans hm). -case: (lerP n 0)=> [|/ltrW] hn. - by rewrite ler_wneexpz2l // [_ \in _](ler_trans hmn). -apply: (@ler_trans _ (x ^ 0)); first by rewrite ler_wneexpz2l. +move=> m n /= hmn; case: (lerP 0 m)=> [|/ltW] hm. + by rewrite ler_wpeexpz2l // [_ \in _](le_trans hm). +case: (lerP n 0)=> [|/ltW] hn. + by rewrite ler_wneexpz2l // [_ \in _](le_trans hmn). +apply: (@le_trans _ _ (x ^ 0)); first by rewrite ler_wneexpz2l. by rewrite ler_wpeexpz2l. Qed. @@ -1266,45 +1265,45 @@ Qed. Lemma ieexprIz x (x0 : 0 < x) (nx1 : x != 1) : injective (exprz x). Proof. -apply: wlog_ltr=> // m n hmn; first by move=> hmn'; rewrite hmn. +apply: wlog_lt=> // m n hmn; first by move=> hmn'; rewrite hmn. move=> /(f_equal ( *%R^~ (x ^ (- n)))). -rewrite -!expfzDr ?gtr_eqF // subrr expr0z=> /eqP. -by rewrite pexprz_eq1 ?(ltrW x0) // (negPf nx1) subr_eq0 orbF=> /eqP. +rewrite -!expfzDr ?gt_eqF // subrr expr0z=> /eqP. +by rewrite pexprz_eq1 ?(ltW x0) // (negPf nx1) subr_eq0 orbF=> /eqP. Qed. Lemma ler_piexpz2l x (x0 : 0 < x) (x1 : x < 1) : {in >= 0 &, {mono (exprz x) : x y /~ x <= y}}. Proof. -apply: (ler_nmono_in (inj_nhomo_ltr_in _ _)). - by move=> n m hn hm /=; apply: ieexprIz; rewrite // ltr_eqF. -by apply: ler_wpiexpz2l; rewrite ?ltrW. +apply: (le_nmono_in (inj_nhomo_lt_in _ _)). + by move=> n m hn hm /=; apply: ieexprIz; rewrite // lt_eqF. +by apply: ler_wpiexpz2l; rewrite ?ltW. Qed. Lemma ltr_piexpz2l x (x0 : 0 < x) (x1 : x < 1) : {in >= 0 &, {mono (exprz x) : x y /~ x < y}}. -Proof. exact: (lerW_nmono_in (ler_piexpz2l _ _)). Qed. +Proof. exact: (leW_nmono_in (ler_piexpz2l _ _)). Qed. Lemma ler_niexpz2l x (x0 : 0 < x) (x1 : x < 1) : {in < 0 &, {mono (exprz x) : x y /~ x <= y}}. Proof. -apply: (ler_nmono_in (inj_nhomo_ltr_in _ _)). - by move=> n m hn hm /=; apply: ieexprIz; rewrite // ltr_eqF. -by apply: ler_wniexpz2l; rewrite ?ltrW. +apply: (le_nmono_in (inj_nhomo_lt_in _ _)). + by move=> n m hn hm /=; apply: ieexprIz; rewrite // lt_eqF. +by apply: ler_wniexpz2l; rewrite ?ltW. Qed. Lemma ltr_niexpz2l x (x0 : 0 < x) (x1 : x < 1) : {in < 0 &, {mono (exprz x) : x y /~ x < y}}. -Proof. exact: (lerW_nmono_in (ler_niexpz2l _ _)). Qed. +Proof. exact: (leW_nmono_in (ler_niexpz2l _ _)). Qed. Lemma ler_eexpz2l x (x1 : 1 < x) : {mono (exprz x) : x y / x <= y}. Proof. -apply: (ler_mono (inj_homo_ltr _ _)). - by apply: ieexprIz; rewrite ?(ltr_trans ltr01) // gtr_eqF. -by apply: ler_weexpz2l; rewrite ?ltrW. +apply: (le_mono (inj_homo_lt _ _)). + by apply: ieexprIz; rewrite ?(lt_trans ltr01) // gt_eqF. +by apply: ler_weexpz2l; rewrite ?ltW. Qed. Lemma ltr_eexpz2l x (x1 : 1 < x) : {mono (exprz x) : x y / x < y}. -Proof. exact: (lerW_mono (ler_eexpz2l _)). Qed. +Proof. exact: (leW_mono (ler_eexpz2l _)). Qed. Lemma ler_wpexpz2r n (hn : 0 <= n) : {in >= 0 & , {homo ((@exprz R)^~ n) : x y / x <= y}}. @@ -1314,7 +1313,7 @@ Lemma ler_wnexpz2r n (hn : n <= 0) : {in > 0 & , {homo ((@exprz R)^~ n) : x y /~ x <= y}}. Proof. move=> x y /= hx hy hxy; rewrite -lef_pinv ?[_ \in _]exprz_gt0 //. -by rewrite !invr_expz ler_wpexpz2r ?[_ \in _]ltrW // oppr_cp0. +by rewrite !invr_expz ler_wpexpz2r ?[_ \in _]ltW // oppr_cp0. Qed. Lemma pexpIrz n (n0 : n != 0) : {in >= 0 &, injective ((@exprz R)^~ n)}. @@ -1324,46 +1323,46 @@ move=> x y; rewrite ![_ \in _]le0r=> /orP [/eqP-> _ /eqP|hx]. case/orP=> [/eqP-> /eqP|hy]. by rewrite exp0rz ?(negPf n0) expfz_eq0=> /andP [_ /eqP]. move=> /(f_equal ( *%R^~ (y ^ (- n)))) /eqP. -rewrite -expfzDr ?(gtr_eqF hy) // subrr expr0z -exprz_inv -expfzMl. -rewrite pexprz_eq1 ?(negPf n0) /= ?mulr_ge0 ?invr_ge0 ?ltrW //. -by rewrite (can2_eq (mulrVK _) (mulrK _)) ?unitfE ?(gtr_eqF hy) // mul1r=> /eqP. +rewrite -expfzDr ?(gt_eqF hy) // subrr expr0z -exprz_inv -expfzMl. +rewrite pexprz_eq1 ?(negPf n0) /= ?mulr_ge0 ?invr_ge0 ?ltW //. +by rewrite (can2_eq (mulrVK _) (mulrK _)) ?unitfE ?(gt_eqF hy) // mul1r=> /eqP. Qed. Lemma nexpIrz n (n0 : n != 0) : {in <= 0 &, injective ((@exprz R)^~ n)}. Proof. -move=> x y; rewrite ![_ \in _]ler_eqVlt => /orP [/eqP -> _ /eqP|hx]. +move=> x y; rewrite ![_ \in _]le_eqVlt => /orP [/eqP -> _ /eqP|hx]. by rewrite exp0rz ?(negPf n0) eq_sym expfz_eq0=> /andP [_ /eqP->]. case/orP=> [/eqP -> /eqP|hy]. by rewrite exp0rz ?(negPf n0) expfz_eq0=> /andP [_ /eqP]. move=> /(f_equal ( *%R^~ (y ^ (- n)))) /eqP. -rewrite -expfzDr ?(ltr_eqF hy) // subrr expr0z -exprz_inv -expfzMl. -rewrite pexprz_eq1 ?(negPf n0) /= ?mulr_le0 ?invr_le0 ?ltrW //. -by rewrite (can2_eq (mulrVK _) (mulrK _)) ?unitfE ?(ltr_eqF hy) // mul1r=> /eqP. +rewrite -expfzDr ?(lt_eqF hy) // subrr expr0z -exprz_inv -expfzMl. +rewrite pexprz_eq1 ?(negPf n0) /= ?mulr_le0 ?invr_le0 ?ltW //. +by rewrite (can2_eq (mulrVK _) (mulrK _)) ?unitfE ?(lt_eqF hy) // mul1r=> /eqP. Qed. Lemma ler_pexpz2r n (hn : 0 < n) : {in >= 0 & , {mono ((@exprz R)^~ n) : x y / x <= y}}. Proof. -apply: ler_mono_in (inj_homo_ltr_in _ _). - by move=> x y hx hy /=; apply: pexpIrz; rewrite // gtr_eqF. -by apply: ler_wpexpz2r; rewrite ltrW. +apply: le_mono_in (inj_homo_lt_in _ _). + by move=> x y hx hy /=; apply: pexpIrz; rewrite // gt_eqF. +by apply: ler_wpexpz2r; rewrite ltW. Qed. Lemma ltr_pexpz2r n (hn : 0 < n) : {in >= 0 & , {mono ((@exprz R)^~ n) : x y / x < y}}. -Proof. exact: lerW_mono_in (ler_pexpz2r _). Qed. +Proof. exact: leW_mono_in (ler_pexpz2r _). Qed. Lemma ler_nexpz2r n (hn : n < 0) : {in > 0 & , {mono ((@exprz R)^~ n) : x y /~ x <= y}}. Proof. -apply: ler_nmono_in (inj_nhomo_ltr_in _ _); last first. - by apply: ler_wnexpz2r; rewrite ltrW. -by move=> x y hx hy /=; apply: pexpIrz; rewrite ?[_ \in _]ltrW ?ltr_eqF. +apply: le_nmono_in (inj_nhomo_lt_in _ _); last first. + by apply: ler_wnexpz2r; rewrite ltW. +by move=> x y hx hy /=; apply: pexpIrz; rewrite ?[_ \in _]ltW ?lt_eqF. Qed. Lemma ltr_nexpz2r n (hn : n < 0) : {in > 0 & , {mono ((@exprz R)^~ n) : x y /~ x < y}}. -Proof. exact: lerW_nmono_in (ler_nexpz2r _). Qed. +Proof. exact: leW_nmono_in (ler_nexpz2r _). Qed. Lemma eqr_expz2 n x y : n != 0 -> 0 <= x -> 0 <= y -> (x ^ n == y ^ n) = (x == y). @@ -1388,10 +1387,10 @@ Proof. by rewrite /sgz; case: (_ == _); case: (_ < _). Qed. Lemma sgrEz x : sgr x = (sgz x)%:~R. Proof. by rewrite !(fun_if intr). Qed. Lemma gtr0_sgz x : 0 < x -> sgz x = 1. -Proof. by move=> x_gt0; rewrite /sgz ltr_neqAle andbC eqr_le ltr_geF //. Qed. +Proof. by move=> x_gt0; rewrite /sgz lt_neqAle andbC eq_le lt_geF. Qed. Lemma ltr0_sgz x : x < 0 -> sgz x = -1. -Proof. by move=> x_lt0; rewrite /sgz eq_sym eqr_le x_lt0 ltr_geF. Qed. +Proof. by move=> x_lt0; rewrite /sgz eq_sym eq_le x_lt0 lt_geF. Qed. Lemma sgz0 : sgz (0 : R) = 0. Proof. by rewrite /sgz eqxx. Qed. Lemma sgz1 : sgz (1 : R) = 1. Proof. by rewrite gtr0_sgz // ltr01. Qed. @@ -1459,7 +1458,7 @@ Lemma sgzP x : (sgz x == 0) (sgz x == -1) (sgz x == 1) `|x| (sgr x) (sgz x). Proof. rewrite ![_ == sgz _]eq_sym ![_ == sgr _]eq_sym !sgr_cp0 !sgz_cp0. -by rewrite /sgr /sgz !lerNgt; case: ltrgt0P; constructor. +by rewrite /sgr /sgz !leNgt; case: ltrgt0P; constructor. Qed. Lemma sgzN x : sgz (- x) = - sgz x. @@ -1661,14 +1660,14 @@ Lemma leqif_add_distz m1 m2 m3 : `|m1 - m3| <= `|m1 - m2| + `|m2 - m3| ?= iff (m1 <= m2 <= m3)%R || (m3 <= m2 <= m1)%R. Proof. -apply/leqifP; rewrite -ltz_nat -eqz_nat PoszD !abszE; apply/lerifP. +apply/leqifP; rewrite -ltz_nat -eqz_nat PoszD !abszE; apply/leifP. wlog le_m31 : m1 m3 / (m3 <= m1)%R. - move=> IH; case/orP: (ler_total m1 m3) => /IH //. + move=> IH; case/orP: (le_total m1 m3) => /IH //. by rewrite (addrC `|_|)%R orbC !(distrC m1) !(distrC m3). rewrite ger0_norm ?subr_ge0 // orb_idl => [|/andP[le_m12 le_m23]]; last first. - by have /eqP->: m2 == m3; rewrite ?lerr // eqr_le le_m23 (ler_trans le_m31). -rewrite -{1}(subrK m2 m1) -addrA -subr_ge0 andbC -subr_ge0. -by apply: lerif_add; apply/real_lerif_norm/num_real. + by have /eqP->: m2 == m3; rewrite ?lexx // eq_le le_m23 (le_trans le_m31). +rewrite -{1}(subrK m2 m1) -addrA -subr_ge0 andbC -[X in X && _]subr_ge0. +by apply: leif_add; apply/real_leif_norm/num_real. Qed. Lemma leqif_add_dist n1 n2 n3 : @@ -1692,7 +1691,7 @@ Section NormInt. Variable R : numDomainType. -Lemma intr_norm m : `|m|%:~R = `|m%:~R| :> R. +Lemma intr_norm m : `|m|%:~R = `|m%:~R : R|. Proof. by rewrite {2}[m]intEsign rmorphMsign normrMsign abszE normr_nat. Qed. Lemma normrMz m (x : R) : `|x *~ m| = `|x| *~ `|m|. diff --git a/mathcomp/algebra/ssrnum.v b/mathcomp/algebra/ssrnum.v index 495ce18..05ac8ed 100644 --- a/mathcomp/algebra/ssrnum.v +++ b/mathcomp/algebra/ssrnum.v @@ -1,95 +1,100 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice. -From mathcomp Require Import div fintype path bigop finset fingroup. +From mathcomp Require Import div fintype path bigop order finset fingroup. From mathcomp Require Import ssralg poly. (******************************************************************************) -(* *) (* This file defines some classes to manipulate number structures, i.e *) -(* structures with an order and a norm *) +(* structures with an order and a norm. To use this file, insert *) +(* "Import Num.Theory." before your scripts. You can also "Import Num.Def." *) +(* to enjoy shorter notations (e.g., minr instead of Num.min, lerif instead *) +(* of Num.leif, etc.). *) (* *) -(* * NumDomain (Integral domain with an order and a norm) *) -(* NumMixin == the mixin that provides an order and a norm over *) -(* a ring and their characteristic properties. *) +(* * NumDomain (Integral domain with an order and a norm) *) (* numDomainType == interface for a num integral domain. *) -(* NumDomainType T m *) -(* == packs the num mixin into a numberDomainType. The *) -(* carrier T must have a integral domain structure. *) -(* [numDomainType of T for S ] *) -(* == T-clone of the numDomainType structure S. *) -(* [numDomainType of T] *) -(* == clone of a canonical numDomainType structure on T. *) +(* NumDomainType T m *) +(* == packs the num mixin into a numDomainType. The carrier *) +(* T must have an integral domain and a partial order *) +(* structures. *) +(* [numDomainType of T for S] *) +(* == T-clone of the numDomainType structure S. *) +(* [numDomainType of T] *) +(* == clone of a canonical numDomainType structure on T. *) +(* *) +(* * NormedZmodule (Zmodule with a norm) *) +(* normedZmodType R *) +(* == interface for a normed Zmodule structure indexed by *) +(* numDomainType R. *) +(* NormedZmodType R T m *) +(* == pack the normed Zmodule mixin into a normedZmodType. *) +(* The carrier T must have an integral domain structure. *) +(* [normedZmodType R of T for S] *) +(* == T-clone of the normedZmodType R structure S. *) +(* [normedZmodType R of T] *) +(* == clone of a canonical normedZmodType R structure on T. *) (* *) (* * NumField (Field with an order and a norm) *) -(* numFieldType == interface for a num field. *) -(* [numFieldType of T] *) -(* == clone of a canonical numFieldType structure on T *) +(* numFieldType == interface for a num field. *) +(* [numFieldType of T] *) +(* == clone of a canonical numFieldType structure on T. *) (* *) -(* * NumClosedField (Closed Field with an order and a norm) *) -(* numClosedFieldType *) -(* == interface for a num closed field. *) -(* [numClosedFieldType of T] *) -(* == clone of a canonical numClosedFieldType structure on T *) +(* * NumClosedField (Partially ordered Closed Field with conjugation) *) +(* numClosedFieldType *) +(* == interface for a closed field with conj. *) +(* NumClosedFieldType T r *) +(* == packs the real closed axiom r into a *) +(* numClosedFieldType. The carrier T must have a closed *) +(* field type structure. *) +(* [numClosedFieldType of T] *) +(* == clone of a canonical numClosedFieldType structure on T.*) +(* [numClosedFieldType of T for S] *) +(* == T-clone of the numClosedFieldType structure S. *) (* *) -(* * RealDomain (Num domain where all elements are positive or negative) *) +(* * RealDomain (Num domain where all elements are positive or negative) *) (* realDomainType == interface for a real integral domain. *) -(* RealDomainType T r *) -(* == packs the real axiom r into a realDomainType. The *) -(* carrier T must have a num domain structure. *) -(* [realDomainType of T for S ] *) -(* == T-clone of the realDomainType structure S. *) -(* [realDomainType of T] *) +(* [realDomainType of T] *) (* == clone of a canonical realDomainType structure on T. *) (* *) (* * RealField (Num Field where all elements are positive or negative) *) (* realFieldType == interface for a real field. *) -(* [realFieldType of T] *) -(* == clone of a canonical realFieldType structure on T *) +(* [realFieldType of T] *) +(* == clone of a canonical realFieldType structure on T. *) (* *) (* * ArchiField (A Real Field with the archimedean axiom) *) -(* archiFieldType == interface for an archimedean field. *) +(* archiFieldType == interface for an archimedean field. *) (* ArchiFieldType T r *) -(* == packs the archimeadean axiom r into an archiFieldType. *) -(* The carrier T must have a real field type structure. *) -(* [archiFieldType of T for S ] *) -(* == T-clone of the archiFieldType structure S. *) -(* [archiFieldType of T] *) -(* == clone of a canonical archiFieldType structure on T *) +(* == packs the archimedean axiom r into an archiFieldType. *) +(* The carrier T must have a real field type structure. *) +(* [archiFieldType of T for S] *) +(* == T-clone of the archiFieldType structure S. *) +(* [archiFieldType of T] *) +(* == clone of a canonical archiFieldType structure on T. *) (* *) (* * RealClosedField (Real Field with the real closed axiom) *) -(* rcfType == interface for a real closed field. *) -(* RcfType T r == packs the real closed axiom r into a *) -(* rcfType. The carrier T must have a real *) -(* field type structure. *) -(* [rcfType of T] == clone of a canonical realClosedFieldType structure on *) +(* rcfType == interface for a real closed field. *) +(* RcfType T r == packs the real closed axiom r into a rcfType. *) +(* The carrier T must have a real field type structure. *) +(* [rcfType of T] == clone of a canonical realClosedFieldType structure on *) (* T. *) -(* [rcfType of T for S ] *) -(* == T-clone of the realClosedFieldType structure S. *) +(* [rcfType of T for S] *) +(* == T-clone of the realClosedFieldType structure S. *) (* *) -(* * NumClosedField (Partially ordered Closed Field with conjugation) *) -(* numClosedFieldType == interface for a closed field with conj. *) -(* NumClosedFieldType T r == packs the real closed axiom r into a *) -(* numClosedFieldType. The carrier T must have a closed *) -(* field type structure. *) -(* [numClosedFieldType of T] == clone of a canonical numClosedFieldType *) -(* structure on T *) -(* [numClosedFieldType of T for S ] *) -(* == T-clone of the realClosedFieldType structure S. *) +(* The ordering symbols and notations (<, <=, >, >=, _ <= _ ?= iff _, >=<, *) +(* and ><) and lattice operations (meet and join) defined in order.v are *) +(* redefined for the ring_display in the ring_scope (%R). 0-ary ordering *) +(* symbols for the ring_display have the suffix "%R", e.g., <%R. All the *) +(* other ordering notations are the same as order.v. The meet and join *) +(* operators for the ring_display are Num.min and Num.max. *) (* *) (* Over these structures, we have the following operations *) -(* `|x| == norm of x. *) -(* x <= y <=> x is less than or equal to y (:= '|y - x| == y - x). *) -(* x < y <=> x is less than y (:= (x <= y) && (x != y)). *) -(* x <= y ?= iff C <-> x is less than y, or equal iff C is true. *) -(* Num.sg x == sign of x: equal to 0 iff x = 0, to 1 iff x > 0, and *) -(* to -1 in all other cases (including x < 0). *) +(* `|x| == norm of x. *) +(* Num.sg x == sign of x: equal to 0 iff x = 0, to 1 iff x > 0, and *) +(* to -1 in all other cases (including x < 0). *) (* x \is a Num.pos <=> x is positive (:= x > 0). *) (* x \is a Num.neg <=> x is negative (:= x < 0). *) (* x \is a Num.nneg <=> x is positive or 0 (:= x >= 0). *) (* x \is a Num.real <=> x is real (:= x >= 0 or x < 0). *) -(* Num.min x y == minimum of x y *) -(* Num.max x y == maximum of x y *) (* Num.bound x == in archimedean fields, and upper bound for x, i.e., *) (* and n such that `|x| < n%:R. *) (* Num.sqrt x == in a real-closed field, a positive square root of x if *) @@ -107,17 +112,6 @@ From mathcomp Require Import ssralg poly. (* an thus not equal to -1 for n odd > 1 (this will be shown in *) (* file cyclotomic.v). *) (* *) -(* There are now three distinct uses of the symbols <, <=, > and >=: *) -(* 0-ary, unary (prefix) and binary (infix). *) -(* 0. <%R, <=%R, >%R, >=%R stand respectively for lt, le, gt and ge. *) -(* 1. (< x), (<= x), (> x), (>= x) stand respectively for *) -(* (gt x), (ge x), (lt x), (le x). *) -(* So (< x) is a predicate characterizing elements smaller than x. *) -(* 2. (x < y), (x <= y), ... mean what they are expected to. *) -(* These convention are compatible with haskell's, *) -(* where ((< y) x) = (x < y) = ((<) x y), *) -(* except that we write <%R instead of (<). *) -(* *) (* - list of prefixes : *) (* p : positive *) (* n : negative *) @@ -126,24 +120,15 @@ From mathcomp Require Import ssralg poly. (* i : interior = in [0, 1] or ]0, 1[ *) (* e : exterior = in [1, +oo[ or ]1; +oo[ *) (* w : non strict (weak) monotony *) -(* *) -(* [arg minr_(i < i0 | P) M] == a value i : T minimizing M : R, subject *) -(* to the condition P (i may appear in P and M), and *) -(* provided P holds for i0. *) -(* [arg maxr_(i > i0 | P) M] == a value i maximizing M subject to P and *) -(* provided P holds for i0. *) -(* [arg minr_(i < i0 in A) M] == an i \in A minimizing M if i0 \in A. *) -(* [arg maxr_(i > i0 in A) M] == an i \in A maximizing M if i0 \in A. *) -(* [arg minr_(i < i0) M] == an i : T minimizing M, given i0 : T. *) -(* [arg maxr_(i > i0) M] == an i : T maximizing M, given i0 : T. *) (******************************************************************************) Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. +Local Open Scope order_scope. Local Open Scope ring_scope. -Import GRing.Theory. +Import Order.TTheory GRing.Theory. Reserved Notation "<= y" (at level 35). Reserved Notation ">= y" (at level 35). @@ -154,45 +139,62 @@ Reserved Notation ">= y :> T" (at level 35, y at next level). Reserved Notation "< y :> T" (at level 35, y at next level). Reserved Notation "> y :> T" (at level 35, y at next level). +Fact ring_display : unit. Proof. exact: tt. Qed. + Module Num. -(* Principal mixin; further classes add axioms rather than operations. *) -Record mixin_of (R : ringType) := Mixin { - norm_op : R -> R; - le_op : rel R; - lt_op : rel R; +Record normed_mixin_of (R T : zmodType) (Rorder : lePOrderMixin R) + (le_op := Order.POrder.le Rorder) + := NormedMixin { + norm_op : T -> R; _ : forall x y, le_op (norm_op (x + y)) (norm_op x + norm_op y); - _ : forall x y, lt_op 0 x -> lt_op 0 y -> lt_op 0 (x + y); _ : forall x, norm_op x = 0 -> x = 0; + _ : forall x n, norm_op (x *+ n) = norm_op x *+ n; + _ : forall x, norm_op (- x) = norm_op x; +}. + +Record mixin_of (R : ringType) (Rorder : lePOrderMixin R) + (le_op := Order.POrder.le Rorder) (lt_op := Order.POrder.lt Rorder) + (normed : @normed_mixin_of R R Rorder) (norm_op := norm_op normed) + := Mixin { + _ : forall x y, lt_op 0 x -> lt_op 0 y -> lt_op 0 (x + y); _ : forall x y, le_op 0 x -> le_op 0 y -> le_op x y || le_op y x; _ : {morph norm_op : x y / x * y}; _ : forall x y, (le_op x y) = (norm_op (y - x) == y - x); - _ : forall x y, (lt_op x y) = (y != x) && (le_op x y) }. Local Notation ring_for T b := (@GRing.Ring.Pack T b). -(* Base interface. *) Module NumDomain. Section ClassDef. - Record class_of T := Class { base : GRing.IntegralDomain.class_of T; - mixin : mixin_of (ring_for T base) + order_mixin : lePOrderMixin (ring_for T base); + normed_mixin : normed_mixin_of (ring_for T base) order_mixin; + mixin : mixin_of normed_mixin; }. + Local Coercion base : class_of >-> GRing.IntegralDomain.class_of. +Local Coercion order_base T (class_of_T : class_of T) := + @Order.POrder.Class _ class_of_T (order_mixin class_of_T). + Structure type := Pack {sort; _ : class_of sort}. Local Coercion sort : type >-> Sortclass. Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - Definition clone c of phant_id class c := @Pack T c. -Definition pack b0 (m0 : mixin_of (ring_for T b0)) := - fun bT b & phant_id (GRing.IntegralDomain.class bT) b => - fun m & phant_id m0 m => Pack (@Class T b m). +Definition pack (b0 : GRing.IntegralDomain.class_of _) om0 + (nm0 : @normed_mixin_of (ring_for T b0) (ring_for T b0) om0) + (m0 : @mixin_of (ring_for T b0) om0 nm0) := + fun bT (b : GRing.IntegralDomain.class_of T) + & phant_id (@GRing.IntegralDomain.class bT) b => + fun om & phant_id om0 om => + fun nm & phant_id nm0 nm => + fun m & phant_id m0 m => + @Pack T (@Class T b om nm m). Definition eqType := @Equality.Pack cT xclass. Definition choiceType := @Choice.Pack cT xclass. @@ -202,14 +204,22 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. +Definition porder_zmodType := @GRing.Zmodule.Pack porderType xclass. +Definition porder_ringType := @GRing.Ring.Pack porderType xclass. +Definition porder_comRingType := @GRing.ComRing.Pack porderType xclass. +Definition porder_unitRingType := @GRing.UnitRing.Pack porderType xclass. +Definition porder_comUnitRingType := @GRing.ComUnitRing.Pack porderType xclass. +Definition porder_idomainType := @GRing.IntegralDomain.Pack porderType xclass. End ClassDef. Module Exports. -Coercion base : class_of >-> GRing.IntegralDomain.class_of. -Coercion mixin : class_of >-> mixin_of. Coercion sort : type >-> Sortclass. Bind Scope ring_scope with sort. +Coercion base : class_of >-> GRing.IntegralDomain.class_of. +Coercion order_base : class_of >-> Order.POrder.class_of. +Coercion normed_mixin : class_of >-> normed_mixin_of. Coercion eqType : type >-> Equality.type. Canonical eqType. Coercion choiceType : type >-> Choice.type. @@ -226,9 +236,16 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. +Canonical porder_zmodType. +Canonical porder_ringType. +Canonical porder_comRingType. +Canonical porder_unitRingType. +Canonical porder_comUnitRingType. +Canonical porder_idomainType. Notation numDomainType := type. -Notation NumMixin := Mixin. -Notation NumDomainType T m := (@pack T _ m _ _ id _ id). +Notation NumDomainType T m := (@pack T _ _ _ m _ _ id _ id _ id _ id). Notation "[ 'numDomainType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) (at level 0, format "[ 'numDomainType' 'of' T 'for' cT ]") : form_scope. Notation "[ 'numDomainType' 'of' T ]" := (@clone T _ _ id) @@ -238,36 +255,167 @@ End Exports. End NumDomain. Import NumDomain.Exports. -Module Import Def. Section Def. +Local Notation num_for T b := (@NumDomain.Pack T b). + +Module NormedZmodule. + +Section ClassDef. + +Variable R : numDomainType. + +Record class_of (T : Type) := Class { + base : GRing.Zmodule.class_of T; + mixin : @normed_mixin_of R (@GRing.Zmodule.Pack T base) (NumDomain.class R); +}. + +Local Coercion base : class_of >-> GRing.Zmodule.class_of. +Local Coercion mixin : class_of >-> normed_mixin_of. + +Structure type (phR : phant R) := + Pack { sort; _ : class_of sort }. +Local Coercion sort : type >-> Sortclass. + +Variables (phR : phant R) (T : Type) (cT : type phR). + +Definition class := let: Pack _ c := cT return class_of cT in c. +Definition clone c of phant_id class c := @Pack phR T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). +Definition pack b0 (m0 : @normed_mixin_of R (@GRing.Zmodule.Pack T b0) + (NumDomain.class R)) := + Pack phR (@Class T b0 m0). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition zmodType := @GRing.Zmodule.Pack cT xclass. + +End ClassDef. + +(* TODO: Ideally,`numDomain_normedZmodType` should be located in *) +(* `NumDomain_joins`. Currently, it's located here to make `hierarchy.ml` can *) +(* recognize that `numDomainType` inherits `normedZmodType`. *) +Definition numDomain_normedZmodType (R : numDomainType) : type (Phant R) := + @Pack R (Phant R) R (Class (NumDomain.normed_mixin (NumDomain.class R))). + +Module Exports. +Coercion base : class_of >-> GRing.Zmodule.class_of. +Coercion mixin : class_of >-> normed_mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical eqType. +Coercion choiceType : type >-> Choice.type. +Canonical choiceType. +Coercion zmodType : type >-> GRing.Zmodule.type. +Canonical zmodType. +Coercion numDomain_normedZmodType : NumDomain.type >-> type. +Canonical numDomain_normedZmodType. +Notation normedZmodType R := (type (Phant R)). +Notation NormedZmodType R T m := (@pack _ (Phant R) T _ m). +Notation NormedZmodMixin := Mixin. +Notation "[ 'normedZmodType' R 'of' T 'for' cT ]" := + (@clone _ (Phant R) T cT _ idfun) + (at level 0, format "[ 'normedZmodType' R 'of' T 'for' cT ]") : + form_scope. +Notation "[ 'normedZmodType' R 'of' T ]" := (@clone _ (Phant R) T _ _ id) + (at level 0, format "[ 'normedZmodType' R 'of' T ]") : form_scope. +End Exports. + +End NormedZmodule. +Import NormedZmodule.Exports. + +Module NumDomain_joins. Import NumDomain. -Context {R : type}. -Implicit Types (x y : R) (C : bool). - -Definition normr : R -> R := norm_op (class R). -Definition ler : rel R := le_op (class R). -Definition ltr : rel R := lt_op (class R). -Local Notation "x <= y" := (ler x y) : ring_scope. -Local Notation "x < y" := (ltr x y) : ring_scope. - -Definition ger : simpl_rel R := [rel x y | y <= x]. -Definition gtr : simpl_rel R := [rel x y | y < x]. -Definition lerif x y C : Prop := ((x <= y) * ((x == y) = C))%type. -Definition sgr x : R := if x == 0 then 0 else if x < 0 then -1 else 1. -Definition minr x y : R := if x <= y then x else y. -Definition maxr x y : R := if y <= x then x else y. +Section NumDomain_joins. + +Variables (T : Type) (cT : type). + +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class cT : class_of xT). + +(* Definition normedZmodType : normedZmodType cT := *) +(* @NormedZmodule.Pack *) +(* cT (Phant cT) cT *) +(* (NormedZmodule.Class (NumDomain.normed_mixin xclass)). *) +Notation normedZmodType := (NormedZmodule.numDomain_normedZmodType cT). +Definition normedZmod_ringType := + @GRing.Ring.Pack normedZmodType xclass. +Definition normedZmod_comRingType := + @GRing.ComRing.Pack normedZmodType xclass. +Definition normedZmod_unitRingType := + @GRing.UnitRing.Pack normedZmodType xclass. +Definition normedZmod_comUnitRingType := + @GRing.ComUnitRing.Pack normedZmodType xclass. +Definition normedZmod_idomainType := + @GRing.IntegralDomain.Pack normedZmodType xclass. +Definition normedZmod_porderType := + @Order.POrder.Pack ring_display normedZmodType xclass. + +End NumDomain_joins. +Module Exports. +(* Coercion normedZmodType : type >-> NormedZmodule.type. *) +(* Canonical normedZmodType. *) +Canonical normedZmod_ringType. +Canonical normedZmod_comRingType. +Canonical normedZmod_unitRingType. +Canonical normedZmod_comUnitRingType. +Canonical normedZmod_idomainType. +Canonical normedZmod_porderType. +End Exports. +End NumDomain_joins. +Export NumDomain_joins.Exports. + +Module Import Def. + +Definition normr (R : numDomainType) (T : normedZmodType R) : T -> R := + nosimpl (norm_op (NormedZmodule.class T)). +Arguments normr {R T} x. + +Notation ler := (@Order.le ring_display _) (only parsing). +Notation "@ 'ler' R" := (@Order.le ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation ltr := (@Order.lt ring_display _) (only parsing). +Notation "@ 'ltr' R" := (@Order.lt ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation ger := (@Order.ge ring_display _) (only parsing). +Notation "@ 'ger' R" := (@Order.ge ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation gtr := (@Order.gt ring_display _) (only parsing). +Notation "@ 'gtr' R" := (@Order.gt ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation lerif := (@Order.leif ring_display _) (only parsing). +Notation "@ 'lerif' R" := (@Order.leif ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation comparabler := (@Order.comparable ring_display _) (only parsing). +Notation "@ 'comparabler' R" := (@Order.comparable ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation maxr := (@Order.join ring_display _). +Notation "@ 'maxr' R" := (@Order.join ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. +Notation minr := (@Order.meet ring_display _). +Notation "@ 'minr' R" := (@Order.meet ring_display R) + (at level 10, R at level 8, only parsing) : fun_scope. + +Section Def. +Context {R : numDomainType}. +Implicit Types (x : R). + +Definition sgr x : R := if x == 0 then 0 else if x < 0 then -1 else 1. Definition Rpos : qualifier 0 R := [qualify x : R | 0 < x]. Definition Rneg : qualifier 0 R := [qualify x : R | x < 0]. Definition Rnneg : qualifier 0 R := [qualify x : R | 0 <= x]. Definition Rreal : qualifier 0 R := [qualify x : R | (0 <= x) || (x <= 0)]. + End Def. End Def. (* Shorter qualified names, when Num.Def is not imported. *) -Notation norm := normr. -Notation le := ler. -Notation lt := ltr. -Notation ge := ger. -Notation gt := gtr. +Notation norm := normr (only parsing). +Notation le := ler (only parsing). +Notation lt := ltr (only parsing). +Notation ge := ger (only parsing). +Notation gt := gtr (only parsing). +Notation leif := lerif (only parsing). +Notation comparable := comparabler (only parsing). Notation sg := sgr. Notation max := maxr. Notation min := minr. @@ -286,7 +434,6 @@ Fact Rnneg_key : pred_key (@nneg R). Proof. by []. Qed. Definition Rnneg_keyed := KeyedQualifier Rnneg_key. Fact Rreal_key : pred_key (@real R). Proof. by []. Qed. Definition Rreal_keyed := KeyedQualifier Rreal_key. -Definition ler_of_leif x y C (le_xy : @lerif R x y C) := le_xy.1 : le x y. End Keys. End Keys. (* (Exported) symbolic syntax. *) @@ -295,32 +442,37 @@ Import Def Keys. Notation "`| x |" := (norm x) : ring_scope. -Notation "<%R" := lt : ring_scope. -Notation ">%R" := gt : ring_scope. -Notation "<=%R" := le : ring_scope. -Notation ">=%R" := ge : ring_scope. -Notation "<?=%R" := lerif : ring_scope. - -Notation "< y" := (gt y) : ring_scope. -Notation "< y :> T" := (< (y : T)) : ring_scope. -Notation "> y" := (lt y) : ring_scope. -Notation "> y :> T" := (> (y : T)) : ring_scope. +Notation "<=%R" := le : fun_scope. +Notation ">=%R" := ge : fun_scope. +Notation "<%R" := lt : fun_scope. +Notation ">%R" := gt : fun_scope. +Notation "<?=%R" := leif : fun_scope. +Notation ">=<%R" := comparable : fun_scope. +Notation "><%R" := (fun x y => ~~ (comparable x y)) : fun_scope. Notation "<= y" := (ge y) : ring_scope. -Notation "<= y :> T" := (<= (y : T)) : ring_scope. +Notation "<= y :> T" := (<= (y : T)) (only parsing) : ring_scope. Notation ">= y" := (le y) : ring_scope. -Notation ">= y :> T" := (>= (y : T)) : ring_scope. +Notation ">= y :> T" := (>= (y : T)) (only parsing) : ring_scope. -Notation "x < y" := (lt x y) : ring_scope. -Notation "x < y :> T" := ((x : T) < (y : T)) : ring_scope. -Notation "x > y" := (y < x) (only parsing) : ring_scope. -Notation "x > y :> T" := ((x : T) > (y : T)) (only parsing) : ring_scope. +Notation "< y" := (gt y) : ring_scope. +Notation "< y :> T" := (< (y : T)) (only parsing) : ring_scope. +Notation "> y" := (lt y) : ring_scope. +Notation "> y :> T" := (> (y : T)) (only parsing) : ring_scope. + +Notation ">=< y" := (comparable y) : ring_scope. +Notation ">=< y :> T" := (>=< (y : T)) (only parsing) : ring_scope. Notation "x <= y" := (le x y) : ring_scope. -Notation "x <= y :> T" := ((x : T) <= (y : T)) : ring_scope. +Notation "x <= y :> T" := ((x : T) <= (y : T)) (only parsing) : ring_scope. Notation "x >= y" := (y <= x) (only parsing) : ring_scope. Notation "x >= y :> T" := ((x : T) >= (y : T)) (only parsing) : ring_scope. +Notation "x < y" := (lt x y) : ring_scope. +Notation "x < y :> T" := ((x : T) < (y : T)) (only parsing) : ring_scope. +Notation "x > y" := (y < x) (only parsing) : ring_scope. +Notation "x > y :> T" := ((x : T) > (y : T)) (only parsing) : ring_scope. + Notation "x <= y <= z" := ((x <= y) && (y <= z)) : ring_scope. Notation "x < y <= z" := ((x < y) && (y <= z)) : ring_scope. Notation "x <= y < z" := ((x <= y) && (y < z)) : ring_scope. @@ -330,13 +482,21 @@ Notation "x <= y ?= 'iff' C" := (lerif x y C) : ring_scope. Notation "x <= y ?= 'iff' C :> R" := ((x : R) <= (y : R) ?= iff C) (only parsing) : ring_scope. -Coercion ler_of_leif : lerif >-> is_true. +Notation ">=< x" := (comparable x) : ring_scope. +Notation ">=< x :> T" := (>=< (x : T)) (only parsing) : ring_scope. +Notation "x >=< y" := (comparable x y) : ring_scope. + +Notation ">< x" := (fun y => ~~ (comparable x y)) : ring_scope. +Notation ">< x :> T" := (>< (x : T)) (only parsing) : ring_scope. +Notation "x >< y" := (~~ (comparable x y)) : ring_scope. Canonical Rpos_keyed. Canonical Rneg_keyed. Canonical Rnneg_keyed. Canonical Rreal_keyed. +Export Order.POCoercions. + End Syntax. Section ExtensionAxioms. @@ -353,18 +513,18 @@ Definition real_closed_axiom : Prop := End ExtensionAxioms. -Local Notation num_for T b := (@NumDomain.Pack T b). - (* The rest of the numbers interface hierarchy. *) Module NumField. Section ClassDef. -Record class_of R := - Class { base : GRing.Field.class_of R; mixin : mixin_of (ring_for R base) }. -Definition base2 R (c : class_of R) := NumDomain.Class (mixin c). -Local Coercion base : class_of >-> GRing.Field.class_of. -Local Coercion base2 : class_of >-> NumDomain.class_of. +Record class_of R := Class { + base : NumDomain.class_of R; + mixin : GRing.Field.mixin_of (num_for R base); +}. +Local Coercion base : class_of >-> NumDomain.class_of. +Local Coercion base2 R (c : class_of R) : GRing.Field.class_of _ := + GRing.Field.Class (@mixin _ c). Structure type := Pack {sort; _ : class_of sort}. Local Coercion sort : type >-> Sortclass. @@ -372,10 +532,9 @@ Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - Definition pack := - fun bT b & phant_id (GRing.Field.class bT) (b : GRing.Field.class_of T) => - fun mT m & phant_id (NumDomain.class mT) (@NumDomain.Class T b m) => + fun bT b & phant_id (NumDomain.class bT) (b : NumDomain.class_of T) => + fun mT m & phant_id (GRing.Field.mixin (GRing.Field.class mT)) m => Pack (@Class T b m). Definition eqType := @Equality.Pack cT xclass. @@ -386,15 +545,20 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. Definition numDomainType := @NumDomain.Pack cT xclass. Definition fieldType := @GRing.Field.Pack cT xclass. -Definition join_numDomainType := @NumDomain.Pack fieldType xclass. +Definition normedZmodType := NormedZmodType numDomainType cT xclass. +Definition porder_fieldType := @GRing.Field.Pack porderType xclass. +Definition normedZmod_fieldType := + @GRing.Field.Pack normedZmodType xclass. +Definition numDomain_fieldType := @GRing.Field.Pack numDomainType xclass. End ClassDef. Module Exports. -Coercion base : class_of >-> GRing.Field.class_of. -Coercion base2 : class_of >-> NumDomain.class_of. +Coercion base : class_of >-> NumDomain.class_of. +Coercion base2 : class_of >-> GRing.Field.class_of. Coercion sort : type >-> Sortclass. Bind Scope ring_scope with sort. Coercion eqType : type >-> Equality.type. @@ -413,11 +577,17 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. Coercion numDomainType : type >-> NumDomain.type. Canonical numDomainType. Coercion fieldType : type >-> GRing.Field.type. Canonical fieldType. -Canonical join_numDomainType. +Coercion normedZmodType : type >-> NormedZmodule.type. +Canonical normedZmodType. +Canonical porder_fieldType. +Canonical normedZmod_fieldType. +Canonical numDomain_fieldType. Notation numFieldType := type. Notation "[ 'numFieldType' 'of' T ]" := (@pack T _ _ id _ _ id) (at level 0, format "[ 'numFieldType' 'of' T ]") : form_scope. @@ -438,13 +608,16 @@ Record imaginary_mixin_of (R : numDomainType) := ImaginaryMixin { }. Record class_of R := Class { - base : GRing.ClosedField.class_of R; - mixin : mixin_of (ring_for R base); - conj_mixin : imaginary_mixin_of (num_for R (NumDomain.Class mixin)) + base : NumField.class_of R; + decField_mixin : GRing.DecidableField.mixin_of (num_for R base); + closedField_axiom : GRing.ClosedField.axiom (num_for R base); + conj_mixin : imaginary_mixin_of (num_for R base); }. -Definition base2 R (c : class_of R) := NumField.Class (mixin c). -Local Coercion base : class_of >-> GRing.ClosedField.class_of. -Local Coercion base2 : class_of >-> NumField.class_of. +Local Coercion base : class_of >-> NumField.class_of. +Local Coercion base2 R (c : class_of R) : GRing.ClosedField.class_of R := + @GRing.ClosedField.Class + R (@GRing.DecidableField.Class R (base c) (@decField_mixin _ c)) + (@closedField_axiom _ c). Structure type := Pack {sort; _ : class_of sort}. Local Coercion sort : type >-> Sortclass. @@ -452,13 +625,14 @@ Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - -Definition pack := - fun bT b & phant_id (GRing.ClosedField.class bT) - (b : GRing.ClosedField.class_of T) => - fun mT m & phant_id (NumField.class mT) (@NumField.Class T b m) => - fun mc => Pack (@Class T b m mc). Definition clone := fun b & phant_id class (b : class_of T) => Pack b. +Definition pack := + fun bT b & phant_id (NumField.class bT) (b : NumField.class_of T) => + fun mT dec closed + & phant_id (GRing.ClosedField.class mT) + (@GRing.ClosedField.Class + _ (@GRing.DecidableField.Class _ b dec) closed) => + fun mc => Pack (@Class T b dec closed mc). Definition eqType := @Equality.Pack cT xclass. Definition choiceType := @Choice.Pack cT xclass. @@ -468,21 +642,33 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. Definition numDomainType := @NumDomain.Pack cT xclass. Definition fieldType := @GRing.Field.Pack cT xclass. Definition numFieldType := @NumField.Pack cT xclass. Definition decFieldType := @GRing.DecidableField.Pack cT xclass. Definition closedFieldType := @GRing.ClosedField.Pack cT xclass. -Definition join_dec_numDomainType := @NumDomain.Pack decFieldType xclass. -Definition join_dec_numFieldType := @NumField.Pack decFieldType xclass. -Definition join_numDomainType := @NumDomain.Pack closedFieldType xclass. -Definition join_numFieldType := @NumField.Pack closedFieldType xclass. +Definition normedZmodType := NormedZmodType numDomainType cT xclass. +Definition porder_decFieldType := @GRing.DecidableField.Pack porderType xclass. +Definition normedZmod_decFieldType := + @GRing.DecidableField.Pack normedZmodType xclass. +Definition numDomain_decFieldType := + @GRing.DecidableField.Pack numDomainType xclass. +Definition numField_decFieldType := + @GRing.DecidableField.Pack numFieldType xclass. +Definition porder_closedFieldType := @GRing.ClosedField.Pack porderType xclass. +Definition normedZmod_closedFieldType := + @GRing.ClosedField.Pack normedZmodType xclass. +Definition numDomain_closedFieldType := + @GRing.ClosedField.Pack numDomainType xclass. +Definition numField_closedFieldType := + @GRing.ClosedField.Pack numFieldType xclass. End ClassDef. Module Exports. -Coercion base : class_of >-> GRing.ClosedField.class_of. -Coercion base2 : class_of >-> NumField.class_of. +Coercion base : class_of >-> NumField.class_of. +Coercion base2 : class_of >-> GRing.ClosedField.class_of. Coercion sort : type >-> Sortclass. Bind Scope ring_scope with sort. Coercion eqType : type >-> Equality.type. @@ -501,6 +687,8 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. Coercion numDomainType : type >-> NumDomain.type. Canonical numDomainType. Coercion fieldType : type >-> GRing.Field.type. @@ -511,12 +699,18 @@ Coercion numFieldType : type >-> NumField.type. Canonical numFieldType. Coercion closedFieldType : type >-> GRing.ClosedField.type. Canonical closedFieldType. -Canonical join_dec_numDomainType. -Canonical join_dec_numFieldType. -Canonical join_numDomainType. -Canonical join_numFieldType. +Coercion normedZmodType : type >-> NormedZmodule.type. +Canonical normedZmodType. +Canonical porder_decFieldType. +Canonical normedZmod_decFieldType. +Canonical numDomain_decFieldType. +Canonical numField_decFieldType. +Canonical porder_closedFieldType. +Canonical normedZmod_closedFieldType. +Canonical numDomain_closedFieldType. +Canonical numField_closedFieldType. Notation numClosedFieldType := type. -Notation NumClosedFieldType T m := (@pack T _ _ id _ _ id m). +Notation NumClosedFieldType T m := (@pack T _ _ id _ _ _ id m). Notation "[ 'numClosedFieldType' 'of' T 'for' cT ]" := (@clone T cT _ id) (at level 0, format "[ 'numClosedFieldType' 'of' T 'for' cT ]") : form_scope. @@ -531,9 +725,18 @@ Module RealDomain. Section ClassDef. -Record class_of R := - Class {base : NumDomain.class_of R; _ : @real_axiom (num_for R base)}. +Record class_of R := Class { + base : NumDomain.class_of R; + lmixin_disp : unit; + lmixin : Order.DistrLattice.mixin_of (Order.POrder.Pack lmixin_disp base); + tmixin_disp : unit; + tmixin : Order.Total.mixin_of + (Order.DistrLattice.Pack + tmixin_disp (Order.DistrLattice.Class lmixin)); +}. Local Coercion base : class_of >-> NumDomain.class_of. +Local Coercion base2 T (c : class_of T) : Order.Total.class_of T := + @Order.Total.Class _ (Order.DistrLattice.Class (@lmixin _ c)) _ (@tmixin _ c). Structure type := Pack {sort; _ : class_of sort}. Local Coercion sort : type >-> Sortclass. @@ -541,11 +744,13 @@ Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - -Definition clone c of phant_id class c := @Pack T c. -Definition pack b0 (m0 : real_axiom (num_for T b0)) := - fun bT b & phant_id (NumDomain.class bT) b => - fun m & phant_id m0 m => Pack (@Class T b m). +Definition pack := + fun bT b & phant_id (NumDomain.class bT) (b : NumDomain.class_of T) => + fun mT ldisp l mdisp m & + phant_id (@Order.Total.class ring_display mT) + (@Order.Total.Class + T (@Order.DistrLattice.Class T b ldisp l) mdisp m) => + Pack (@Class T b ldisp l mdisp m). Definition eqType := @Equality.Pack cT xclass. Definition choiceType := @Choice.Pack cT xclass. @@ -555,12 +760,47 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. +Definition distrLatticeType := @Order.DistrLattice.Pack ring_display cT xclass. +Definition orderType := @Order.Total.Pack ring_display cT xclass. Definition numDomainType := @NumDomain.Pack cT xclass. +Definition normedZmodType := NormedZmodType numDomainType cT xclass. +Definition zmod_distrLatticeType := + @Order.DistrLattice.Pack ring_display zmodType xclass. +Definition ring_distrLatticeType := + @Order.DistrLattice.Pack ring_display ringType xclass. +Definition comRing_distrLatticeType := + @Order.DistrLattice.Pack ring_display comRingType xclass. +Definition unitRing_distrLatticeType := + @Order.DistrLattice.Pack ring_display unitRingType xclass. +Definition comUnitRing_distrLatticeType := + @Order.DistrLattice.Pack ring_display comUnitRingType xclass. +Definition idomain_distrLatticeType := + @Order.DistrLattice.Pack ring_display idomainType xclass. +Definition normedZmod_distrLatticeType := + @Order.DistrLattice.Pack ring_display normedZmodType xclass. +Definition numDomain_distrLatticeType := + @Order.DistrLattice.Pack ring_display numDomainType xclass. +Definition zmod_orderType := @Order.Total.Pack ring_display zmodType xclass. +Definition ring_orderType := @Order.Total.Pack ring_display ringType xclass. +Definition comRing_orderType := + @Order.Total.Pack ring_display comRingType xclass. +Definition unitRing_orderType := + @Order.Total.Pack ring_display unitRingType xclass. +Definition comUnitRing_orderType := + @Order.Total.Pack ring_display comUnitRingType xclass. +Definition idomain_orderType := + @Order.Total.Pack ring_display idomainType xclass. +Definition normedZmod_orderType := + @Order.Total.Pack ring_display normedZmodType xclass. +Definition numDomain_orderType := + @Order.Total.Pack ring_display numDomainType xclass. End ClassDef. Module Exports. Coercion base : class_of >-> NumDomain.class_of. +Coercion base2 : class_of >-> Order.Total.class_of. Coercion sort : type >-> Sortclass. Bind Scope ring_scope with sort. Coercion eqType : type >-> Equality.type. @@ -579,13 +819,34 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. Coercion numDomainType : type >-> NumDomain.type. Canonical numDomainType. +Coercion distrLatticeType : type >-> Order.DistrLattice.type. +Canonical distrLatticeType. +Coercion orderType : type >-> Order.Total.type. +Canonical orderType. +Coercion normedZmodType : type >-> NormedZmodule.type. +Canonical normedZmodType. +Canonical zmod_distrLatticeType. +Canonical ring_distrLatticeType. +Canonical comRing_distrLatticeType. +Canonical unitRing_distrLatticeType. +Canonical comUnitRing_distrLatticeType. +Canonical idomain_distrLatticeType. +Canonical normedZmod_distrLatticeType. +Canonical numDomain_distrLatticeType. +Canonical zmod_orderType. +Canonical ring_orderType. +Canonical comRing_orderType. +Canonical unitRing_orderType. +Canonical comUnitRing_orderType. +Canonical idomain_orderType. +Canonical normedZmod_orderType. +Canonical numDomain_orderType. Notation realDomainType := type. -Notation RealDomainType T m := (@pack T _ m _ _ id _ id). -Notation "[ 'realDomainType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) - (at level 0, format "[ 'realDomainType' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'realDomainType' 'of' T ]" := (@clone T _ _ id) +Notation "[ 'realDomainType' 'of' T ]" := (@pack T _ _ id _ _ _ _ _ id) (at level 0, format "[ 'realDomainType' 'of' T ]") : form_scope. End Exports. @@ -596,11 +857,18 @@ Module RealField. Section ClassDef. -Record class_of R := - Class { base : NumField.class_of R; mixin : real_axiom (num_for R base) }. -Definition base2 R (c : class_of R) := RealDomain.Class (@mixin R c). +Record class_of R := Class { + base : NumField.class_of R; + lmixin_disp : unit; + lmixin : Order.DistrLattice.mixin_of (@Order.POrder.Pack lmixin_disp R base); + tmixin_disp : unit; + tmixin : Order.Total.mixin_of + (Order.DistrLattice.Pack + tmixin_disp (Order.DistrLattice.Class lmixin)); +}. Local Coercion base : class_of >-> NumField.class_of. -Local Coercion base2 : class_of >-> RealDomain.class_of. +Local Coercion base2 R (c : class_of R) : RealDomain.class_of R := + RealDomain.Class (@tmixin R c). Structure type := Pack {sort; _ : class_of sort}. Local Coercion sort : type >-> Sortclass. @@ -608,11 +876,11 @@ Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - Definition pack := - fun bT b & phant_id (NumField.class bT) (b : NumField.class_of T) => - fun mT m & phant_id (RealDomain.class mT) (@RealDomain.Class T b m) => - Pack (@Class T b m). + fun bT (b : NumField.class_of T) & phant_id (NumField.class bT) b => + fun mT ldisp l tdisp t & phant_id (RealDomain.class mT) + (@RealDomain.Class T b ldisp l tdisp t) => + Pack (@Class T b ldisp l tdisp t). Definition eqType := @Equality.Pack cT xclass. Definition choiceType := @Choice.Pack cT xclass. @@ -622,12 +890,23 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. Definition numDomainType := @NumDomain.Pack cT xclass. +Definition distrLatticeType := @Order.DistrLattice.Pack ring_display cT xclass. +Definition orderType := @Order.Total.Pack ring_display cT xclass. Definition realDomainType := @RealDomain.Pack cT xclass. Definition fieldType := @GRing.Field.Pack cT xclass. Definition numFieldType := @NumField.Pack cT xclass. -Definition join_fieldType := @GRing.Field.Pack realDomainType xclass. -Definition join_numFieldType := @NumField.Pack realDomainType xclass. +Definition normedZmodType := NormedZmodType numDomainType cT xclass. +Definition field_distrLatticeType := + @Order.DistrLattice.Pack ring_display fieldType xclass. +Definition field_orderType := @Order.Total.Pack ring_display fieldType xclass. +Definition field_realDomainType := @RealDomain.Pack fieldType xclass. +Definition numField_distrLatticeType := + @Order.DistrLattice.Pack ring_display numFieldType xclass. +Definition numField_orderType := + @Order.Total.Pack ring_display numFieldType xclass. +Definition numField_realDomainType := @RealDomain.Pack numFieldType xclass. End ClassDef. @@ -652,18 +931,30 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. Coercion numDomainType : type >-> NumDomain.type. Canonical numDomainType. +Coercion distrLatticeType : type >-> Order.DistrLattice.type. +Canonical distrLatticeType. +Coercion orderType : type >-> Order.Total.type. +Canonical orderType. Coercion realDomainType : type >-> RealDomain.type. Canonical realDomainType. Coercion fieldType : type >-> GRing.Field.type. Canonical fieldType. Coercion numFieldType : type >-> NumField.type. Canonical numFieldType. -Canonical join_fieldType. -Canonical join_numFieldType. +Coercion normedZmodType : type >-> NormedZmodule.type. +Canonical normedZmodType. +Canonical field_distrLatticeType. +Canonical field_orderType. +Canonical field_realDomainType. +Canonical numField_distrLatticeType. +Canonical numField_orderType. +Canonical numField_realDomainType. Notation realFieldType := type. -Notation "[ 'realFieldType' 'of' T ]" := (@pack T _ _ id _ _ id) +Notation "[ 'realFieldType' 'of' T ]" := (@pack T _ _ id _ _ _ _ _ id) (at level 0, format "[ 'realFieldType' 'of' T ]") : form_scope. End Exports. @@ -684,7 +975,6 @@ Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - Definition clone c of phant_id class c := @Pack T c. Definition pack b0 (m0 : archimedean_axiom (num_for T b0)) := fun bT b & phant_id (RealField.class bT) b => @@ -698,11 +988,15 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. Definition numDomainType := @NumDomain.Pack cT xclass. +Definition distrLatticeType := @Order.DistrLattice.Pack ring_display cT xclass. +Definition orderType := @Order.Total.Pack ring_display cT xclass. Definition realDomainType := @RealDomain.Pack cT xclass. Definition fieldType := @GRing.Field.Pack cT xclass. Definition numFieldType := @NumField.Pack cT xclass. Definition realFieldType := @RealField.Pack cT xclass. +Definition normedZmodType := NormedZmodType numDomainType cT xclass. End ClassDef. @@ -726,8 +1020,14 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. Coercion numDomainType : type >-> NumDomain.type. Canonical numDomainType. +Coercion distrLatticeType : type >-> Order.DistrLattice.type. +Canonical distrLatticeType. +Coercion orderType : type >-> Order.Total.type. +Canonical orderType. Coercion realDomainType : type >-> RealDomain.type. Canonical realDomainType. Coercion fieldType : type >-> GRing.Field.type. @@ -736,6 +1036,8 @@ Coercion numFieldType : type >-> NumField.type. Canonical numFieldType. Coercion realFieldType : type >-> RealField.type. Canonical realFieldType. +Coercion normedZmodType : type >-> NormedZmodule.type. +Canonical normedZmodType. Notation archiFieldType := type. Notation ArchiFieldType T m := (@pack T _ m _ _ id _ id). Notation "[ 'archiFieldType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) @@ -761,7 +1063,6 @@ Variables (T : Type) (cT : type). Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. Let xT := let: Pack T _ := cT in T. Notation xclass := (class : class_of xT). - Definition clone c of phant_id class c := @Pack T c. Definition pack b0 (m0 : real_closed_axiom (num_for T b0)) := fun bT b & phant_id (RealField.class bT) b => @@ -775,11 +1076,15 @@ Definition comRingType := @GRing.ComRing.Pack cT xclass. Definition unitRingType := @GRing.UnitRing.Pack cT xclass. Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass. Definition idomainType := @GRing.IntegralDomain.Pack cT xclass. +Definition porderType := @Order.POrder.Pack ring_display cT xclass. Definition numDomainType := @NumDomain.Pack cT xclass. +Definition distrLatticeType := @Order.DistrLattice.Pack ring_display cT xclass. +Definition orderType := @Order.Total.Pack ring_display cT xclass. Definition realDomainType := @RealDomain.Pack cT xclass. Definition fieldType := @GRing.Field.Pack cT xclass. Definition numFieldType := @NumField.Pack cT xclass. Definition realFieldType := @RealField.Pack cT xclass. +Definition normedZmodType := NormedZmodType numDomainType cT xclass. End ClassDef. @@ -803,16 +1108,24 @@ Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. Canonical comUnitRingType. Coercion idomainType : type >-> GRing.IntegralDomain.type. Canonical idomainType. +Coercion porderType : type >-> Order.POrder.type. +Canonical porderType. Coercion numDomainType : type >-> NumDomain.type. Canonical numDomainType. Coercion realDomainType : type >-> RealDomain.type. Canonical realDomainType. +Coercion distrLatticeType : type >-> Order.DistrLattice.type. +Canonical distrLatticeType. +Coercion orderType : type >-> Order.Total.type. +Canonical orderType. Coercion fieldType : type >-> GRing.Field.type. Canonical fieldType. Coercion numFieldType : type >-> NumField.type. Canonical numFieldType. Coercion realFieldType : type >-> RealField.type. Canonical realFieldType. +Coercion normedZmodType : type >-> NormedZmodule.type. +Canonical normedZmodType. Notation rcfType := Num.RealClosedField.type. Notation RcfType T m := (@pack T _ m _ _ id _ id). Notation "[ 'rcfType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) @@ -828,32 +1141,41 @@ Import RealClosedField.Exports. (* operations for the extensions described above. *) Module Import Internals. -Section Domain. -Variable R : numDomainType. -Implicit Types x y : R. +Section NormedZmodule. +Variables (R : numDomainType) (V : normedZmodType R). +Implicit Types (l : R) (x y : V). -(* Lemmas from the signature *) +Lemma ler_norm_add x y : `|x + y| <= `|x| + `|y|. +Proof. by case: V x y => ? [? []]. Qed. Lemma normr0_eq0 x : `|x| = 0 -> x = 0. -Proof. by case: R x => ? [? []]. Qed. +Proof. by case: V x => ? [? []]. Qed. -Lemma ler_norm_add x y : `|x + y| <= `|x| + `|y|. -Proof. by case: R x y => ? [? []]. Qed. +Lemma normrMn x n : `|x *+ n| = `|x| *+ n. +Proof. by case: V x => ? [? []]. Qed. + +Lemma normrN x : `|- x| = `|x|. +Proof. by case: V x => ? [? []]. Qed. + +End NormedZmodule. + +Section NumDomain. +Variable R : numDomainType. +Implicit Types x y : R. + +(* Lemmas from the signature *) Lemma addr_gt0 x y : 0 < x -> 0 < y -> 0 < x + y. -Proof. by case: R x y => ? [? []]. Qed. +Proof. by case: R x y => ? [? ? ? []]. Qed. Lemma ger_leVge x y : 0 <= x -> 0 <= y -> (x <= y) || (y <= x). -Proof. by case: R x y => ? [? []]. Qed. +Proof. by case: R x y => ? [? ? ? []]. Qed. -Lemma normrM : {morph norm : x y / x * y : R}. -Proof. by case: R => ? [? []]. Qed. +Lemma normrM : {morph norm : x y / (x : R) * y}. +Proof. by case: R => ? [? ? ? []]. Qed. Lemma ler_def x y : (x <= y) = (`|y - x| == y - x). -Proof. by case: R x y => ? [? []]. Qed. - -Lemma ltr_def x y : (x < y) = (y != x) && (x <= y). -Proof. by case: R x y => ? [? []]. Qed. +Proof. by case: R x y => ? [? ? ? []]. Qed. (* Basic consequences (just enough to get predicate closure properties). *) @@ -868,33 +1190,24 @@ Proof. by rewrite -sub0r subr_ge0. Qed. Lemma ler01 : 0 <= 1 :> R. Proof. -have n1_nz: `|1| != 0 :> R by apply: contraNneq (@oner_neq0 R) => /normr0_eq0->. +have n1_nz: `|1 : R| != 0 by apply: contraNneq (@oner_neq0 R) => /normr0_eq0->. by rewrite ger0_def -(inj_eq (mulfI n1_nz)) -normrM !mulr1. Qed. -Lemma ltr01 : 0 < 1 :> R. Proof. by rewrite ltr_def oner_neq0 ler01. Qed. - -Lemma ltrW x y : x < y -> x <= y. Proof. by rewrite ltr_def => /andP[]. Qed. - -Lemma lerr x : x <= x. -Proof. -have n2: `|2%:R| == 2%:R :> R by rewrite -ger0_def ltrW ?addr_gt0 ?ltr01. -rewrite ler_def subrr -(inj_eq (addrI `|0|)) addr0 -mulr2n -mulr_natr. -by rewrite -(eqP n2) -normrM mul0r. -Qed. +Lemma ltr01 : 0 < 1 :> R. Proof. by rewrite lt_def oner_neq0 ler01. Qed. Lemma le0r x : (0 <= x) = (x == 0) || (0 < x). -Proof. by rewrite ltr_def; case: eqP => // ->; rewrite lerr. Qed. +Proof. by rewrite lt_def; case: eqP => // ->; rewrite lexx. Qed. Lemma addr_ge0 x y : 0 <= x -> 0 <= y -> 0 <= x + y. Proof. rewrite le0r; case/predU1P=> [-> | x_pos]; rewrite ?add0r // le0r. -by case/predU1P=> [-> | y_pos]; rewrite ltrW ?addr0 ?addr_gt0. +by case/predU1P=> [-> | y_pos]; rewrite ltW ?addr0 ?addr_gt0. Qed. Lemma pmulr_rgt0 x y : 0 < x -> (0 < x * y) = (0 < y). Proof. -rewrite !ltr_def !ger0_def normrM mulf_eq0 negb_or => /andP[x_neq0 /eqP->]. +rewrite !lt_def !ger0_def normrM mulf_eq0 negb_or => /andP[x_neq0 /eqP->]. by rewrite x_neq0 (inj_eq (mulfI x_neq0)). Qed. @@ -923,7 +1236,7 @@ Canonical nneg_mulrPred := MulrPred nneg_divr_closed. Canonical nneg_divrPred := DivrPred nneg_divr_closed. Fact nneg_addr_closed : addr_closed (@nneg R). -Proof. by split; [apply: lerr | apply: addr_ge0]. Qed. +Proof. by split; [apply: lexx | apply: addr_ge0]. Qed. Canonical nneg_addrPred := AddrPred nneg_addr_closed. Canonical nneg_semiringPred := SemiringPred nneg_divr_closed. @@ -933,7 +1246,7 @@ Canonical real_opprPred := OpprPred real_oppr_closed. Fact real_addr_closed : addr_closed (@real R). Proof. -split=> [|x y Rx Ry]; first by rewrite realE lerr. +split=> [|x y Rx Ry]; first by rewrite realE lexx. without loss{Rx} x_ge0: x y Ry / 0 <= x. case/orP: Rx => [? | x_le0]; first exact. by rewrite -rpredN opprD; apply; rewrite ?rpredN ?oppr_ge0. @@ -961,10 +1274,10 @@ Canonical real_semiringPred := SemiringPred real_divr_closed. Canonical real_subringPred := SubringPred real_divr_closed. Canonical real_divringPred := DivringPred real_divr_closed. -End Domain. +End NumDomain. Lemma num_real (R : realDomainType) (x : R) : x \is real. -Proof. by case: R x => T []. Qed. +Proof. exact: le_total. Qed. Fact archi_bound_subproof (R : archiFieldType) : archimedean_axiom R. Proof. by case: R => ? []. Qed. @@ -1023,28 +1336,28 @@ End ExtraDef. Notation bound := archi_bound. Notation sqrt := sqrtr. -Module Theory. +Module Import Theory. Section NumIntegralDomainTheory. Variable R : numDomainType. -Implicit Types x y z t : R. +Implicit Types (V : normedZmodType R) (x y z t : R). (* Lemmas from the signature (reexported from internals). *) -Definition ler_norm_add x y : `|x + y| <= `|x| + `|y| := ler_norm_add x y. +Definition ler_norm_add V (x y : V) : `|x + y| <= `|x| + `|y| := + ler_norm_add x y. Definition addr_gt0 x y : 0 < x -> 0 < y -> 0 < x + y := @addr_gt0 R x y. -Definition normr0_eq0 x : `|x| = 0 -> x = 0 := @normr0_eq0 R x. +Definition normr0_eq0 V (x : V) : `|x| = 0 -> x = 0 := @normr0_eq0 R V x. Definition ger_leVge x y : 0 <= x -> 0 <= y -> (x <= y) || (y <= x) := @ger_leVge R x y. -Definition normrM : {morph normr : x y / x * y : R} := @normrM R. -Definition ler_def x y : (x <= y) = (`|y - x| == y - x) := @ler_def R x y. -Definition ltr_def x y : (x < y) = (y != x) && (x <= y) := @ltr_def R x y. +Definition normrM : {morph norm : x y / (x : R) * y} := @normrM R. +Definition ler_def x y : (x <= y) = (`|y - x| == y - x) := ler_def x y. +Definition normrMn V (x : V) n : `|x *+ n| = `|x| *+ n := normrMn x n. +Definition normrN V (x : V) : `|- x| = `|x| := normrN x. -(* Predicate and relation definitions. *) +(* Predicate definitions. *) -Lemma gerE x y : ge x y = (y <= x). Proof. by []. Qed. -Lemma gtrE x y : gt x y = (y < x). Proof. by []. Qed. Lemma posrE x : (x \is pos) = (0 < x). Proof. by []. Qed. Lemma negrE x : (x \is neg) = (x < 0). Proof. by []. Qed. Lemma nnegrE x : (x \is nneg) = (0 <= x). Proof. by []. Qed. @@ -1052,31 +1365,14 @@ Lemma realE x : (x \is real) = (0 <= x) || (x <= 0). Proof. by []. Qed. (* General properties of <= and < *) -Lemma lerr x : x <= x. Proof. exact: lerr. Qed. -Lemma ltrr x : x < x = false. Proof. by rewrite ltr_def eqxx. Qed. -Lemma ltrW x y : x < y -> x <= y. Proof. exact: ltrW. Qed. -Hint Resolve lerr ltrr ltrW : core. - -Lemma ltr_neqAle x y : (x < y) = (x != y) && (x <= y). -Proof. by rewrite ltr_def eq_sym. Qed. - -Lemma ler_eqVlt x y : (x <= y) = (x == y) || (x < y). -Proof. by rewrite ltr_neqAle; case: eqP => // ->; rewrite lerr. Qed. - -Lemma lt0r x : (0 < x) = (x != 0) && (0 <= x). Proof. by rewrite ltr_def. Qed. +Lemma lt0r x : (0 < x) = (x != 0) && (0 <= x). Proof. by rewrite lt_def. Qed. Lemma le0r x : (0 <= x) = (x == 0) || (0 < x). Proof. exact: le0r. Qed. -Lemma lt0r_neq0 (x : R) : 0 < x -> x != 0. +Lemma lt0r_neq0 (x : R) : 0 < x -> x != 0. Proof. by rewrite lt0r; case/andP. Qed. -Lemma ltr0_neq0 (x : R) : x < 0 -> x != 0. -Proof. by rewrite ltr_neqAle; case/andP. Qed. - -Lemma gtr_eqF x y : y < x -> x == y = false. -Proof. by rewrite ltr_def; case/andP; move/negPf=> ->. Qed. - -Lemma ltr_eqF x y : x < y -> x == y = false. -Proof. by move=> hyx; rewrite eq_sym gtr_eqF. Qed. +Lemma ltr0_neq0 (x : R) : x < 0 -> x != 0. +Proof. by rewrite lt_neqAle; case/andP. Qed. Lemma pmulr_rgt0 x y : 0 < x -> (0 < x * y) = (0 < y). Proof. exact: pmulr_rgt0. Qed. @@ -1098,7 +1394,7 @@ Proof. by case: n => //= n; apply: ltr0Sn. Qed. Hint Resolve ltr0Sn : core. Lemma pnatr_eq0 n : (n%:R == 0 :> R) = (n == 0)%N. -Proof. by case: n => [|n]; rewrite ?mulr0n ?eqxx // gtr_eqF. Qed. +Proof. by case: n => [|n]; rewrite ?mulr0n ?eqxx // gt_eqF. Qed. Lemma char_num : [char R] =i pred0. Proof. by case=> // p /=; rewrite !inE pnatr_eq0 andbF. Qed. @@ -1109,12 +1405,8 @@ Lemma ger0_def x : (0 <= x) = (`|x| == x). Proof. exact: ger0_def. Qed. Lemma normr_idP {x} : reflect (`|x| = x) (0 <= x). Proof. by rewrite ger0_def; apply: eqP. Qed. Lemma ger0_norm x : 0 <= x -> `|x| = x. Proof. exact: normr_idP. Qed. - -Lemma normr0 : `|0| = 0 :> R. Proof. exact: ger0_norm. Qed. -Lemma normr1 : `|1| = 1 :> R. Proof. exact: ger0_norm. Qed. -Lemma normr_nat n : `|n%:R| = n%:R :> R. Proof. exact: ger0_norm. Qed. -Lemma normrMn x n : `|x *+ n| = `|x| *+ n. -Proof. by rewrite -mulr_natl normrM normr_nat mulr_natl. Qed. +Lemma normr1 : `|1 : R| = 1. Proof. exact: ger0_norm. Qed. +Lemma normr_nat n : `|n%:R : R| = n%:R. Proof. exact: ger0_norm. Qed. Lemma normr_prod I r (P : pred I) (F : I -> R) : `|\prod_(i <- r | P i) F i| = \prod_(i <- r | P i) `|F i|. @@ -1123,60 +1415,78 @@ Proof. exact: (big_morph norm normrM normr1). Qed. Lemma normrX n x : `|x ^+ n| = `|x| ^+ n. Proof. by rewrite -(card_ord n) -!prodr_const normr_prod. Qed. -Lemma normr_unit : {homo (@norm R) : x / x \is a GRing.unit}. +Lemma normr_unit : {homo (@norm R R) : x / x \is a GRing.unit}. Proof. move=> x /= /unitrP [y [yx xy]]; apply/unitrP; exists `|y|. by rewrite -!normrM xy yx normr1. Qed. -Lemma normrV : {in GRing.unit, {morph (@normr R) : x / x ^-1}}. +Lemma normrV : {in GRing.unit, {morph (@norm R R) : x / x ^-1}}. Proof. move=> x ux; apply: (mulrI (normr_unit ux)). by rewrite -normrM !divrr ?normr1 ?normr_unit. Qed. -Lemma normr0P {x} : reflect (`|x| = 0) (x == 0). -Proof. by apply: (iffP eqP)=> [->|/normr0_eq0 //]; apply: normr0. Qed. - -Definition normr_eq0 x := sameP (`|x| =P 0) normr0P. - -Lemma normrN1 : `|-1| = 1 :> R. +Lemma normrN1 : `|-1 : R| = 1. Proof. -have: `|-1| ^+ 2 == 1 :> R by rewrite -normrX -signr_odd normr1. +have: `|-1 : R| ^+ 2 == 1 by rewrite -normrX -signr_odd normr1. rewrite sqrf_eq1 => /orP[/eqP //|]; rewrite -ger0_def le0r oppr_eq0 oner_eq0. -by move/(addr_gt0 ltr01); rewrite subrr ltrr. +by move/(addr_gt0 ltr01); rewrite subrr ltxx. Qed. -Lemma normrN x : `|- x| = `|x|. -Proof. by rewrite -mulN1r normrM normrN1 mul1r. Qed. +Section NormedZmoduleTheory. -Lemma distrC x y : `|x - y| = `|y - x|. -Proof. by rewrite -opprB normrN. Qed. +Variable V : normedZmodType R. +Implicit Types (v w : V). -Lemma ler0_def x : (x <= 0) = (`|x| == - x). -Proof. by rewrite ler_def sub0r normrN. Qed. +Lemma normr0 : `|0 : V| = 0. +Proof. by rewrite -(mulr0n 0) normrMn mulr0n. Qed. + +Lemma normr0P v : reflect (`|v| = 0) (v == 0). +Proof. by apply: (iffP eqP)=> [->|/normr0_eq0 //]; apply: normr0. Qed. -Lemma normr_id x : `|`|x| | = `|x|. +Definition normr_eq0 v := sameP (`|v| =P 0) (normr0P v). + +Lemma distrC v w : `|v - w| = `|w - v|. +Proof. by rewrite -opprB normrN. Qed. + +Lemma normr_id v : `| `|v| | = `|v|. Proof. have nz2: 2%:R != 0 :> R by rewrite pnatr_eq0. apply: (mulfI nz2); rewrite -{1}normr_nat -normrM mulr_natl mulr2n ger0_norm //. -by rewrite -{2}normrN -normr0 -(subrr x) ler_norm_add. +by rewrite -{2}normrN -normr0 -(subrr v) ler_norm_add. Qed. -Lemma normr_ge0 x : 0 <= `|x|. Proof. by rewrite ger0_def normr_id. Qed. -Hint Resolve normr_ge0 : core. +Lemma normr_ge0 v : 0 <= `|v|. Proof. by rewrite ger0_def normr_id. Qed. + +Lemma normr_le0 v : `|v| <= 0 = (v == 0). +Proof. by rewrite -normr_eq0 eq_le normr_ge0 andbT. Qed. + +Lemma normr_lt0 v : `|v| < 0 = false. +Proof. by rewrite lt_neqAle normr_le0 normr_eq0 andNb. Qed. + +Lemma normr_gt0 v : `|v| > 0 = (v != 0). +Proof. by rewrite lt_def normr_eq0 normr_ge0 andbT. Qed. + +Definition normrE := (normr_id, normr0, normr1, normrN1, normr_ge0, normr_eq0, + normr_lt0, normr_le0, normr_gt0, normrN). + +End NormedZmoduleTheory. + +Lemma ler0_def x : (x <= 0) = (`|x| == - x). +Proof. by rewrite ler_def sub0r normrN. Qed. Lemma ler0_norm x : x <= 0 -> `|x| = - x. Proof. by move=> x_le0; rewrite -[r in _ = r]ger0_norm ?normrN ?oppr_ge0. Qed. -Definition gtr0_norm x (hx : 0 < x) := ger0_norm (ltrW hx). -Definition ltr0_norm x (hx : x < 0) := ler0_norm (ltrW hx). +Definition gtr0_norm x (hx : 0 < x) := ger0_norm (ltW hx). +Definition ltr0_norm x (hx : x < 0) := ler0_norm (ltW hx). (* Comparision to 0 of a difference *) Lemma subr_ge0 x y : (0 <= y - x) = (x <= y). Proof. exact: subr_ge0. Qed. Lemma subr_gt0 x y : (0 < y - x) = (x < y). -Proof. by rewrite !ltr_def subr_eq0 subr_ge0. Qed. +Proof. by rewrite !lt_def subr_eq0 subr_ge0. Qed. Lemma subr_le0 x y : (y - x <= 0) = (y <= x). Proof. by rewrite -subr_ge0 opprB add0r subr_ge0. Qed. Lemma subr_lt0 x y : (y - x < 0) = (y < x). @@ -1186,360 +1496,38 @@ Definition subr_lte0 := (subr_le0, subr_lt0). Definition subr_gte0 := (subr_ge0, subr_gt0). Definition subr_cp0 := (subr_lte0, subr_gte0). -(* Ordered ring properties. *) - -Lemma ler_asym : antisymmetric (<=%R : rel R). -Proof. -move=> x y; rewrite !ler_def distrC -opprB -addr_eq0 => /andP[/eqP->]. -by rewrite -mulr2n -mulr_natl mulf_eq0 subr_eq0 pnatr_eq0 => /eqP. -Qed. - -Lemma eqr_le x y : (x == y) = (x <= y <= x). -Proof. by apply/eqP/idP=> [->|/ler_asym]; rewrite ?lerr. Qed. +(* Comparability in a numDomain *) -Lemma ltr_trans : transitive (@ltr R). -Proof. -move=> y x z le_xy le_yz. -by rewrite -subr_gt0 -(subrK y z) -addrA addr_gt0 ?subr_gt0. -Qed. +Lemma comparabler0 x : (x >=< 0)%R = (x \is Num.real). +Proof. by rewrite comparable_sym. Qed. -Lemma ler_lt_trans y x z : x <= y -> y < z -> x < z. -Proof. by rewrite !ler_eqVlt => /orP[/eqP -> //|/ltr_trans]; apply. Qed. +Lemma subr_comparable0 x y : (x - y >=< 0)%R = (x >=< y)%R. +Proof. by rewrite /comparable subr_ge0 subr_le0. Qed. -Lemma ltr_le_trans y x z : x < y -> y <= z -> x < z. -Proof. by rewrite !ler_eqVlt => lxy /orP[/eqP <- //|/(ltr_trans lxy)]. Qed. +Lemma comparablerE x y : (x >=< y)%R = (x - y \is Num.real). +Proof. by rewrite -comparabler0 subr_comparable0. Qed. -Lemma ler_trans : transitive (@ler R). +Lemma comparabler_trans : transitive (comparable : rel R). Proof. -move=> y x z; rewrite !ler_eqVlt => /orP [/eqP -> //|lxy]. -by move=> /orP [/eqP <-|/(ltr_trans lxy) ->]; rewrite ?lxy orbT. +move=> y x z; rewrite !comparablerE => xBy_real yBz_real. +by have := rpredD xBy_real yBz_real; rewrite addrA addrNK. Qed. +(* Ordered ring properties. *) + Definition lter01 := (ler01, ltr01). -Definition lterr := (lerr, ltrr). Lemma addr_ge0 x y : 0 <= x -> 0 <= y -> 0 <= x + y. Proof. exact: addr_ge0. Qed. -Lemma lerifP x y C : reflect (x <= y ?= iff C) (if C then x == y else x < y). -Proof. -rewrite /lerif ler_eqVlt; apply: (iffP idP)=> [|[]]. - by case: C => [/eqP->|lxy]; rewrite ?eqxx // lxy ltr_eqF. -by move=> /orP[/eqP->|lxy] <-; rewrite ?eqxx // ltr_eqF. -Qed. - -Lemma ltr_asym x y : x < y < x = false. -Proof. by apply/negP=> /andP [/ltr_trans hyx /hyx]; rewrite ltrr. Qed. - -Lemma ler_anti : antisymmetric (@ler R). -Proof. by move=> x y; rewrite -eqr_le=> /eqP. Qed. - -Lemma ltr_le_asym x y : x < y <= x = false. -Proof. by rewrite ltr_neqAle -andbA -eqr_le eq_sym; case: (_ == _). Qed. - -Lemma ler_lt_asym x y : x <= y < x = false. -Proof. by rewrite andbC ltr_le_asym. Qed. - -Definition lter_anti := (=^~ eqr_le, ltr_asym, ltr_le_asym, ler_lt_asym). - -Lemma ltr_geF x y : x < y -> (y <= x = false). -Proof. -by move=> xy; apply: contraTF isT=> /(ltr_le_trans xy); rewrite ltrr. -Qed. - -Lemma ler_gtF x y : x <= y -> (y < x = false). -Proof. by apply: contraTF=> /ltr_geF->. Qed. - -Definition ltr_gtF x y hxy := ler_gtF (@ltrW x y hxy). - -(* Norm and order properties. *) - -Lemma normr_le0 x : (`|x| <= 0) = (x == 0). -Proof. by rewrite -normr_eq0 eqr_le normr_ge0 andbT. Qed. - -Lemma normr_lt0 x : `|x| < 0 = false. -Proof. by rewrite ltr_neqAle normr_le0 normr_eq0 andNb. Qed. - -Lemma normr_gt0 x : (`|x| > 0) = (x != 0). -Proof. by rewrite ltr_def normr_eq0 normr_ge0 andbT. Qed. - -Definition normrE x := (normr_id, normr0, normr1, normrN1, normr_ge0, normr_eq0, - normr_lt0, normr_le0, normr_gt0, normrN). - End NumIntegralDomainTheory. Arguments ler01 {R}. Arguments ltr01 {R}. Arguments normr_idP {R x}. -Arguments normr0P {R x}. -Arguments lerifP {R x y C}. -Hint Resolve @ler01 @ltr01 lerr ltrr ltrW ltr_eqF ltr0Sn ler0n normr_ge0 : core. - -Section NumIntegralDomainMonotonyTheory. - -Variables R R' : numDomainType. -Implicit Types m n p : nat. -Implicit Types x y z : R. -Implicit Types u v w : R'. - -(****************************************************************************) -(* This listing of "Let"s factor out the required premices for the *) -(* subsequent lemmas, putting them in the context so that "done" solves the *) -(* goals quickly *) -(****************************************************************************) - -Let leqnn := leqnn. -Let ltnE := ltn_neqAle. -Let ltrE := @ltr_neqAle R. -Let ltr'E := @ltr_neqAle R'. -Let gtnE (m n : nat) : (m > n)%N = (m != n) && (m >= n)%N. -Proof. by rewrite ltn_neqAle eq_sym. Qed. -Let gtrE (x y : R) : (x > y) = (x != y) && (x >= y). -Proof. by rewrite ltr_neqAle eq_sym. Qed. -Let gtr'E (x y : R') : (x > y) = (x != y) && (x >= y). -Proof. by rewrite ltr_neqAle eq_sym. Qed. -Let leq_anti : antisymmetric leq. -Proof. by move=> m n; rewrite -eqn_leq => /eqP. Qed. -Let geq_anti : antisymmetric geq. -Proof. by move=> m n; rewrite -eqn_leq => /eqP. Qed. -Let ler_antiR := @ler_anti R. -Let ler_antiR' := @ler_anti R'. -Let ger_antiR : antisymmetric (>=%R : rel R). -Proof. by move=> ??; rewrite andbC; apply: ler_anti. Qed. -Let ger_antiR' : antisymmetric (>=%R : rel R'). -Proof. by move=> ??; rewrite andbC; apply: ler_anti. Qed. -Let leq_total := leq_total. -Let geq_total : total geq. -Proof. by move=> m n; apply: leq_total. Qed. - -Section AcrossTypes. - -Variables (D D' : {pred R}) (f : R -> R'). - -Lemma ltrW_homo : {homo f : x y / x < y} -> {homo f : x y / x <= y}. -Proof. exact: homoW. Qed. - -Lemma ltrW_nhomo : {homo f : x y /~ x < y} -> {homo f : x y /~ x <= y}. -Proof. exact: homoW. Qed. - -Lemma inj_homo_ltr : - injective f -> {homo f : x y / x <= y} -> {homo f : x y / x < y}. -Proof. exact: inj_homo. Qed. - -Lemma inj_nhomo_ltr : - injective f -> {homo f : x y /~ x <= y} -> {homo f : x y /~ x < y}. -Proof. exact: inj_homo. Qed. - -Lemma incr_inj : {mono f : x y / x <= y} -> injective f. -Proof. exact: mono_inj. Qed. - -Lemma decr_inj : {mono f : x y /~ x <= y} -> injective f. -Proof. exact: mono_inj. Qed. - -Lemma lerW_mono : {mono f : x y / x <= y} -> {mono f : x y / x < y}. -Proof. exact: anti_mono. Qed. - -Lemma lerW_nmono : {mono f : x y /~ x <= y} -> {mono f : x y /~ x < y}. -Proof. exact: anti_mono. Qed. - -(* Monotony in D D' *) -Lemma ltrW_homo_in : - {in D & D', {homo f : x y / x < y}} -> {in D & D', {homo f : x y / x <= y}}. -Proof. exact: homoW_in. Qed. - -Lemma ltrW_nhomo_in : - {in D & D', {homo f : x y /~ x < y}} -> {in D & D', {homo f : x y /~ x <= y}}. -Proof. exact: homoW_in. Qed. - -Lemma inj_homo_ltr_in : - {in D & D', injective f} -> {in D & D', {homo f : x y / x <= y}} -> - {in D & D', {homo f : x y / x < y}}. -Proof. exact: inj_homo_in. Qed. - -Lemma inj_nhomo_ltr_in : - {in D & D', injective f} -> {in D & D', {homo f : x y /~ x <= y}} -> - {in D & D', {homo f : x y /~ x < y}}. -Proof. exact: inj_homo_in. Qed. - -Lemma incr_inj_in : {in D &, {mono f : x y / x <= y}} -> - {in D &, injective f}. -Proof. exact: mono_inj_in. Qed. - -Lemma decr_inj_in : - {in D &, {mono f : x y /~ x <= y}} -> {in D &, injective f}. -Proof. exact: mono_inj_in. Qed. - -Lemma lerW_mono_in : - {in D &, {mono f : x y / x <= y}} -> {in D &, {mono f : x y / x < y}}. -Proof. exact: anti_mono_in. Qed. - -Lemma lerW_nmono_in : - {in D &, {mono f : x y /~ x <= y}} -> {in D &, {mono f : x y /~ x < y}}. -Proof. exact: anti_mono_in. Qed. - -End AcrossTypes. - -Section NatToR. - -Variables (D D' : {pred nat}) (f : nat -> R). - -Lemma ltnrW_homo : {homo f : m n / (m < n)%N >-> m < n} -> - {homo f : m n / (m <= n)%N >-> m <= n}. -Proof. exact: homoW. Qed. - -Lemma ltnrW_nhomo : {homo f : m n / (n < m)%N >-> m < n} -> - {homo f : m n / (n <= m)%N >-> m <= n}. -Proof. exact: homoW. Qed. - -Lemma inj_homo_ltnr : injective f -> - {homo f : m n / (m <= n)%N >-> m <= n} -> - {homo f : m n / (m < n)%N >-> m < n}. -Proof. exact: inj_homo. Qed. - -Lemma inj_nhomo_ltnr : injective f -> - {homo f : m n / (n <= m)%N >-> m <= n} -> - {homo f : m n / (n < m)%N >-> m < n}. -Proof. exact: inj_homo. Qed. - -Lemma incnr_inj : {mono f : m n / (m <= n)%N >-> m <= n} -> injective f. -Proof. exact: mono_inj. Qed. - -Lemma decnr_inj_inj : {mono f : m n / (n <= m)%N >-> m <= n} -> injective f. -Proof. exact: mono_inj. Qed. - -Lemma lenrW_mono : {mono f : m n / (m <= n)%N >-> m <= n} -> - {mono f : m n / (m < n)%N >-> m < n}. -Proof. exact: anti_mono. Qed. - -Lemma lenrW_nmono : {mono f : m n / (n <= m)%N >-> m <= n} -> - {mono f : m n / (n < m)%N >-> m < n}. -Proof. exact: anti_mono. Qed. - -Lemma lenr_mono : {homo f : m n / (m < n)%N >-> m < n} -> - {mono f : m n / (m <= n)%N >-> m <= n}. -Proof. exact: total_homo_mono. Qed. - -Lemma lenr_nmono : {homo f : m n / (n < m)%N >-> m < n} -> - {mono f : m n / (n <= m)%N >-> m <= n}. -Proof. exact: total_homo_mono. Qed. - -Lemma ltnrW_homo_in : {in D & D', {homo f : m n / (m < n)%N >-> m < n}} -> - {in D & D', {homo f : m n / (m <= n)%N >-> m <= n}}. -Proof. exact: homoW_in. Qed. - -Lemma ltnrW_nhomo_in : {in D & D', {homo f : m n / (n < m)%N >-> m < n}} -> - {in D & D', {homo f : m n / (n <= m)%N >-> m <= n}}. -Proof. exact: homoW_in. Qed. - -Lemma inj_homo_ltnr_in : {in D & D', injective f} -> - {in D & D', {homo f : m n / (m <= n)%N >-> m <= n}} -> - {in D & D', {homo f : m n / (m < n)%N >-> m < n}}. -Proof. exact: inj_homo_in. Qed. - -Lemma inj_nhomo_ltnr_in : {in D & D', injective f} -> - {in D & D', {homo f : m n / (n <= m)%N >-> m <= n}} -> - {in D & D', {homo f : m n / (n < m)%N >-> m < n}}. -Proof. exact: inj_homo_in. Qed. - -Lemma incnr_inj_in : {in D &, {mono f : m n / (m <= n)%N >-> m <= n}} -> - {in D &, injective f}. -Proof. exact: mono_inj_in. Qed. - -Lemma decnr_inj_inj_in : {in D &, {mono f : m n / (n <= m)%N >-> m <= n}} -> - {in D &, injective f}. -Proof. exact: mono_inj_in. Qed. - -Lemma lenrW_mono_in : {in D &, {mono f : m n / (m <= n)%N >-> m <= n}} -> - {in D &, {mono f : m n / (m < n)%N >-> m < n}}. -Proof. exact: anti_mono_in. Qed. - -Lemma lenrW_nmono_in : {in D &, {mono f : m n / (n <= m)%N >-> m <= n}} -> - {in D &, {mono f : m n / (n < m)%N >-> m < n}}. -Proof. exact: anti_mono_in. Qed. - -Lemma lenr_mono_in : {in D &, {homo f : m n / (m < n)%N >-> m < n}} -> - {in D &, {mono f : m n / (m <= n)%N >-> m <= n}}. -Proof. exact: total_homo_mono_in. Qed. - -Lemma lenr_nmono_in : {in D &, {homo f : m n / (n < m)%N >-> m < n}} -> - {in D &, {mono f : m n / (n <= m)%N >-> m <= n}}. -Proof. exact: total_homo_mono_in. Qed. - -End NatToR. - -Section RToNat. - -Variables (D D' : {pred R}) (f : R -> nat). - -Lemma ltrnW_homo : {homo f : m n / m < n >-> (m < n)%N} -> - {homo f : m n / m <= n >-> (m <= n)%N}. -Proof. exact: homoW. Qed. - -Lemma ltrnW_nhomo : {homo f : m n / n < m >-> (m < n)%N} -> - {homo f : m n / n <= m >-> (m <= n)%N}. -Proof. exact: homoW. Qed. - -Lemma inj_homo_ltrn : injective f -> - {homo f : m n / m <= n >-> (m <= n)%N} -> - {homo f : m n / m < n >-> (m < n)%N}. -Proof. exact: inj_homo. Qed. - -Lemma inj_nhomo_ltrn : injective f -> - {homo f : m n / n <= m >-> (m <= n)%N} -> - {homo f : m n / n < m >-> (m < n)%N}. -Proof. exact: inj_homo. Qed. - -Lemma incrn_inj : {mono f : m n / m <= n >-> (m <= n)%N} -> injective f. -Proof. exact: mono_inj. Qed. - -Lemma decrn_inj : {mono f : m n / n <= m >-> (m <= n)%N} -> injective f. -Proof. exact: mono_inj. Qed. - -Lemma lernW_mono : {mono f : m n / m <= n >-> (m <= n)%N} -> - {mono f : m n / m < n >-> (m < n)%N}. -Proof. exact: anti_mono. Qed. - -Lemma lernW_nmono : {mono f : m n / n <= m >-> (m <= n)%N} -> - {mono f : m n / n < m >-> (m < n)%N}. -Proof. exact: anti_mono. Qed. - -Lemma ltrnW_homo_in : {in D & D', {homo f : m n / m < n >-> (m < n)%N}} -> - {in D & D', {homo f : m n / m <= n >-> (m <= n)%N}}. -Proof. exact: homoW_in. Qed. - -Lemma ltrnW_nhomo_in : {in D & D', {homo f : m n / n < m >-> (m < n)%N}} -> - {in D & D', {homo f : m n / n <= m >-> (m <= n)%N}}. -Proof. exact: homoW_in. Qed. - -Lemma inj_homo_ltrn_in : {in D & D', injective f} -> - {in D & D', {homo f : m n / m <= n >-> (m <= n)%N}} -> - {in D & D', {homo f : m n / m < n >-> (m < n)%N}}. -Proof. exact: inj_homo_in. Qed. - -Lemma inj_nhomo_ltrn_in : {in D & D', injective f} -> - {in D & D', {homo f : m n / n <= m >-> (m <= n)%N}} -> - {in D & D', {homo f : m n / n < m >-> (m < n)%N}}. -Proof. exact: inj_homo_in. Qed. - -Lemma incrn_inj_in : {in D &, {mono f : m n / m <= n >-> (m <= n)%N}} -> - {in D &, injective f}. -Proof. exact: mono_inj_in. Qed. - -Lemma decrn_inj_in : {in D &, {mono f : m n / n <= m >-> (m <= n)%N}} -> - {in D &, injective f}. -Proof. exact: mono_inj_in. Qed. - -Lemma lernW_mono_in : {in D &, {mono f : m n / m <= n >-> (m <= n)%N}} -> - {in D &, {mono f : m n / m < n >-> (m < n)%N}}. -Proof. exact: anti_mono_in. Qed. - -Lemma lernW_nmono_in : {in D &, {mono f : m n / n <= m >-> (m <= n)%N}} -> - {in D &, {mono f : m n / n < m >-> (m < n)%N}}. -Proof. exact: anti_mono_in. Qed. - -End RToNat. - -End NumIntegralDomainMonotonyTheory. +Arguments normr0P {R V v}. +Hint Resolve @ler01 @ltr01 ltr0Sn ler0n : core. +Hint Extern 0 (is_true (0 <= norm _)) => exact: normr_ge0 : core. Section NumDomainOperationTheory. @@ -1552,7 +1540,7 @@ Lemma ler_opp2 : {mono -%R : x y /~ x <= y :> R}. Proof. by move=> x y /=; rewrite -subr_ge0 opprK addrC subr_ge0. Qed. Hint Resolve ler_opp2 : core. Lemma ltr_opp2 : {mono -%R : x y /~ x < y :> R}. -Proof. by move=> x y /=; rewrite lerW_nmono. Qed. +Proof. by move=> x y /=; rewrite leW_nmono. Qed. Hint Resolve ltr_opp2 : core. Definition lter_opp2 := (ler_opp2, ltr_opp2). @@ -1560,7 +1548,7 @@ Lemma ler_oppr x y : (x <= - y) = (y <= - x). Proof. by rewrite (monoRL opprK ler_opp2). Qed. Lemma ltr_oppr x y : (x < - y) = (y < - x). -Proof. by rewrite (monoRL opprK (lerW_nmono _)). Qed. +Proof. by rewrite (monoRL opprK (leW_nmono _)). Qed. Definition lter_oppr := (ler_oppr, ltr_oppr). @@ -1568,7 +1556,7 @@ Lemma ler_oppl x y : (- x <= y) = (- y <= x). Proof. by rewrite (monoLR opprK ler_opp2). Qed. Lemma ltr_oppl x y : (- x < y) = (- y < x). -Proof. by rewrite (monoLR opprK (lerW_nmono _)). Qed. +Proof. by rewrite (monoLR opprK (leW_nmono _)). Qed. Definition lter_oppl := (ler_oppl, ltr_oppl). @@ -1591,23 +1579,23 @@ Definition oppr_cp0 := (oppr_gte0, oppr_lte0). Definition lter_oppE := (oppr_cp0, lter_opp2). Lemma ge0_cp x : 0 <= x -> (- x <= 0) * (- x <= x). -Proof. by move=> hx; rewrite oppr_cp0 hx (@ler_trans _ 0) ?oppr_cp0. Qed. +Proof. by move=> hx; rewrite oppr_cp0 hx (@le_trans _ _ 0) ?oppr_cp0. Qed. Lemma gt0_cp x : 0 < x -> (0 <= x) * (- x <= 0) * (- x <= x) * (- x < 0) * (- x < x). Proof. -move=> hx; move: (ltrW hx) => hx'; rewrite !ge0_cp hx' //. -by rewrite oppr_cp0 hx // (@ltr_trans _ 0) ?oppr_cp0. +move=> hx; move: (ltW hx) => hx'; rewrite !ge0_cp hx' //. +by rewrite oppr_cp0 hx // (@lt_trans _ _ 0) ?oppr_cp0. Qed. Lemma le0_cp x : x <= 0 -> (0 <= - x) * (x <= - x). -Proof. by move=> hx; rewrite oppr_cp0 hx (@ler_trans _ 0) ?oppr_cp0. Qed. +Proof. by move=> hx; rewrite oppr_cp0 hx (@le_trans _ _ 0) ?oppr_cp0. Qed. Lemma lt0_cp x : x < 0 -> (x <= 0) * (0 <= - x) * (x <= - x) * (0 < - x) * (x < - x). Proof. -move=> hx; move: (ltrW hx) => hx'; rewrite !le0_cp // hx'. -by rewrite oppr_cp0 hx // (@ltr_trans _ 0) ?oppr_cp0. +move=> hx; move: (ltW hx) => hx'; rewrite !le0_cp // hx'. +by rewrite oppr_cp0 hx // (@lt_trans _ _ 0) ?oppr_cp0. Qed. (* Properties of the real subset. *) @@ -1619,10 +1607,10 @@ Lemma ler0_real x : x <= 0 -> x \is real. Proof. by rewrite realE orbC => ->. Qed. Lemma gtr0_real x : 0 < x -> x \is real. -Proof. by move=> /ltrW/ger0_real. Qed. +Proof. by move=> /ltW/ger0_real. Qed. Lemma ltr0_real x : x < 0 -> x \is real. -Proof. by move=> /ltrW/ler0_real. Qed. +Proof. by move=> /ltW/ler0_real. Qed. Lemma real0 : 0 \is @real R. Proof. by rewrite ger0_real. Qed. Hint Resolve real0 : core. @@ -1638,19 +1626,21 @@ Proof. by rewrite -!oppr_ge0 => /(ger_leVge _) h /h; rewrite !ler_opp2. Qed. Lemma real_leVge x y : x \is real -> y \is real -> (x <= y) || (y <= x). Proof. rewrite !realE; have [x_ge0 _|x_nge0 /= x_le0] := boolP (_ <= _); last first. - by have [/(ler_trans x_le0)->|_ /(ler_leVge x_le0) //] := boolP (0 <= _). -by have [/(ger_leVge x_ge0)|_ /ler_trans->] := boolP (0 <= _); rewrite ?orbT. + by have [/(le_trans x_le0)->|_ /(ler_leVge x_le0) //] := boolP (0 <= _). +by have [/(ger_leVge x_ge0)|_ /le_trans->] := boolP (0 <= _); rewrite ?orbT. Qed. +Lemma real_comparable x y : x \is real -> y \is real -> x >=< y. +Proof. exact: real_leVge. Qed. + Lemma realB : {in real &, forall x y, x - y \is real}. Proof. exact: rpredB. Qed. Lemma realN : {mono (@GRing.opp R) : x / x \is real}. Proof. exact: rpredN. Qed. -(* :TODO: add a rpredBC in ssralg *) Lemma realBC x y : (x - y \is real) = (y - x \is real). -Proof. by rewrite -realN opprB. Qed. +Proof. exact: rpredBC. Qed. Lemma realD : {in real &, forall x y, x + y \is real}. Proof. exact: rpredD. Qed. @@ -1668,46 +1658,40 @@ Variant ltr_xor_ge (x y : R) : R -> R -> bool -> bool -> Set := Variant comparer x y : R -> R -> bool -> bool -> bool -> bool -> bool -> bool -> Set := | ComparerLt of x < y : comparer x y (y - x) (y - x) - false false true false true false - | ComparerGt of x > y : comparer x y (x - y) (x - y) false false false true false true + | ComparerGt of x > y : comparer x y (x - y) (x - y) + false false true false true false | ComparerEq of x = y : comparer x y 0 0 true true true true false false. -Lemma real_lerP x y : +Lemma real_leP x y : x \is real -> y \is real -> ler_xor_gt x y `|x - y| `|y - x| (x <= y) (y < x). Proof. -move=> xR /(real_leVge xR); have [le_xy _|Nle_xy /= le_yx] := boolP (_ <= _). - have [/(ler_lt_trans le_xy)|] := boolP (_ < _); first by rewrite ltrr. - by rewrite ler0_norm ?ger0_norm ?subr_cp0 ?opprB //; constructor. -have [lt_yx|] := boolP (_ < _). - by rewrite ger0_norm ?ler0_norm ?subr_cp0 ?opprB //; constructor. -by rewrite ltr_def le_yx andbT negbK=> /eqP exy; rewrite exy lerr in Nle_xy. +move=> xR yR; case: (comparable_leP (real_leVge xR yR)) => xy. +- by rewrite [`|x - y|]distrC !ger0_norm ?subr_cp0 //; constructor. +- by rewrite [`|y - x|]distrC !gtr0_norm ?subr_cp0 //; constructor. Qed. -Lemma real_ltrP x y : +Lemma real_ltP x y : x \is real -> y \is real -> ltr_xor_ge x y `|x - y| `|y - x| (y <= x) (x < y). -Proof. by move=> xR yR; case: real_lerP=> //; constructor. Qed. +Proof. by move=> xR yR; case: real_leP=> //; constructor. Qed. -Lemma real_ltrNge : {in real &, forall x y, (x < y) = ~~ (y <= x)}. -Proof. by move=> x y xR yR /=; case: real_lerP. Qed. +Lemma real_ltNge : {in real &, forall x y, (x < y) = ~~ (y <= x)}. +Proof. by move=> x y xR yR /=; case: real_leP. Qed. -Lemma real_lerNgt : {in real &, forall x y, (x <= y) = ~~ (y < x)}. -Proof. by move=> x y xR yR /=; case: real_lerP. Qed. +Lemma real_leNgt : {in real &, forall x y, (x <= y) = ~~ (y < x)}. +Proof. by move=> x y xR yR /=; case: real_leP. Qed. -Lemma real_ltrgtP x y : - x \is real -> y \is real -> +Lemma real_ltgtP x y : x \is real -> y \is real -> comparer x y `|x - y| `|y - x| - (y == x) (x == y) (x <= y) (y <= x) (x < y) (x > y). + (y == x) (x == y) (x >= y) (x <= y) (x > y) (x < y). Proof. -move=> xR yR; case: real_lerP => // [le_yx|lt_xy]; last first. - by rewrite gtr_eqF // ltr_eqF // ler_gtF ?ltrW //; constructor. -case: real_lerP => // [le_xy|lt_yx]; last first. - by rewrite ltr_eqF // gtr_eqF //; constructor. -have /eqP ->: x == y by rewrite eqr_le le_yx le_xy. -by rewrite subrr eqxx; constructor. +move=> xR yR; case: (comparable_ltgtP (real_leVge xR yR)) => [?|?|->]. +- by rewrite [`|x - y|]distrC !gtr0_norm ?subr_gt0//; constructor. +- by rewrite [`|y - x|]distrC !gtr0_norm ?subr_gt0//; constructor. +- by rewrite subrr normr0; constructor. Qed. Variant ger0_xor_lt0 (x : R) : R -> bool -> bool -> Set := @@ -1724,28 +1708,28 @@ Variant comparer0 x : | ComparerLt0 of x < 0 : comparer0 x (- x) false false true false true false | ComparerEq0 of x = 0 : comparer0 x 0 true true true true false false. -Lemma real_ger0P x : x \is real -> ger0_xor_lt0 x `|x| (x < 0) (0 <= x). +Lemma real_ge0P x : x \is real -> ger0_xor_lt0 x `|x| (x < 0) (0 <= x). Proof. -move=> hx; rewrite -{2}[x]subr0; case: real_ltrP; +move=> hx; rewrite -{2}[x]subr0; case: real_ltP; by rewrite ?subr0 ?sub0r //; constructor. Qed. -Lemma real_ler0P x : x \is real -> ler0_xor_gt0 x `|x| (0 < x) (x <= 0). +Lemma real_le0P x : x \is real -> ler0_xor_gt0 x `|x| (0 < x) (x <= 0). Proof. -move=> hx; rewrite -{2}[x]subr0; case: real_ltrP; +move=> hx; rewrite -{2}[x]subr0; case: real_ltP; by rewrite ?subr0 ?sub0r //; constructor. Qed. -Lemma real_ltrgt0P x : +Lemma real_ltgt0P x : x \is real -> comparer0 x `|x| (0 == x) (x == 0) (x <= 0) (0 <= x) (x < 0) (x > 0). Proof. -move=> hx; rewrite -{2}[x]subr0; case: real_ltrgtP; +move=> hx; rewrite -{2}[x]subr0; case: real_ltgtP; by rewrite ?subr0 ?sub0r //; constructor. Qed. Lemma real_neqr_lt : {in real &, forall x y, (x != y) = (x < y) || (y < x)}. -Proof. by move=> * /=; case: real_ltrgtP. Qed. +Proof. by move=> * /=; case: real_ltgtP. Qed. Lemma ler_sub_real x y : x <= y -> y - x \is real. Proof. by move=> le_xy; rewrite ger0_real // subr_ge0. Qed. @@ -1769,10 +1753,10 @@ Lemma Nreal_geF x y : y \is real -> x \notin real -> (y <= x) = false. Proof. by move=> yR; apply: contraNF=> /ger_real->. Qed. Lemma Nreal_ltF x y : y \is real -> x \notin real -> (x < y) = false. -Proof. by move=> yR xNR; rewrite ltr_def Nreal_leF ?andbF. Qed. +Proof. by move=> yR xNR; rewrite lt_def Nreal_leF ?andbF. Qed. Lemma Nreal_gtF x y : y \is real -> x \notin real -> (y < x) = false. -Proof. by move=> yR xNR; rewrite ltr_def Nreal_geF ?andbF. Qed. +Proof. by move=> yR xNR; rewrite lt_def Nreal_geF ?andbF. Qed. (* real wlog *) @@ -1781,7 +1765,7 @@ Lemma real_wlog_ler P : forall a b : R, a \is real -> b \is real -> P a b. Proof. move=> sP hP a b ha hb; wlog: a b ha hb / a <= b => [hwlog|]; last exact: hP. -by case: (real_lerP ha hb)=> [/hP //|/ltrW hba]; apply: sP; apply: hP. +by case: (real_leP ha hb)=> [/hP //|/ltW hba]; apply/sP/hP. Qed. Lemma real_wlog_ltr P : @@ -1790,7 +1774,7 @@ Lemma real_wlog_ltr P : forall a b : R, a \is real -> b \is real -> P a b. Proof. move=> rP sP hP; apply: real_wlog_ler=> // a b. -by rewrite ler_eqVlt; case: (altP (_ =P _))=> [->|] //= _ lab; apply: hP. +rewrite le_eqVlt; case: eqVneq => [->|] //= _ lab; exact: hP. Qed. (* Monotony of addition *) @@ -1803,10 +1787,10 @@ Lemma ler_add2r x : {mono +%R^~ x : y z / y <= z}. Proof. by move=> y z /=; rewrite ![_ + x]addrC ler_add2l. Qed. Lemma ltr_add2l x : {mono +%R x : y z / y < z}. -Proof. by move=> y z /=; rewrite (lerW_mono (ler_add2l _)). Qed. +Proof. by move=> y z /=; rewrite (leW_mono (ler_add2l _)). Qed. Lemma ltr_add2r x : {mono +%R^~ x : y z / y < z}. -Proof. by move=> y z /=; rewrite (lerW_mono (ler_add2r _)). Qed. +Proof. by move=> y z /=; rewrite (leW_mono (ler_add2r _)). Qed. Definition ler_add2 := (ler_add2l, ler_add2r). Definition ltr_add2 := (ltr_add2l, ltr_add2r). @@ -1814,16 +1798,16 @@ Definition lter_add2 := (ler_add2, ltr_add2). (* Addition, subtraction and transitivity *) Lemma ler_add x y z t : x <= y -> z <= t -> x + z <= y + t. -Proof. by move=> lxy lzt; rewrite (@ler_trans _ (y + z)) ?lter_add2. Qed. +Proof. by move=> lxy lzt; rewrite (@le_trans _ _ (y + z)) ?lter_add2. Qed. Lemma ler_lt_add x y z t : x <= y -> z < t -> x + z < y + t. -Proof. by move=> lxy lzt; rewrite (@ler_lt_trans _ (y + z)) ?lter_add2. Qed. +Proof. by move=> lxy lzt; rewrite (@le_lt_trans _ _ (y + z)) ?lter_add2. Qed. Lemma ltr_le_add x y z t : x < y -> z <= t -> x + z < y + t. -Proof. by move=> lxy lzt; rewrite (@ltr_le_trans _ (y + z)) ?lter_add2. Qed. +Proof. by move=> lxy lzt; rewrite (@lt_le_trans _ _ (y + z)) ?lter_add2. Qed. Lemma ltr_add x y z t : x < y -> z < t -> x + z < y + t. -Proof. by move=> lxy lzt; rewrite ltr_le_add // ltrW. Qed. +Proof. by move=> lxy lzt; rewrite ltr_le_add // ltW. Qed. Lemma ler_sub x y z t : x <= y -> t <= z -> x - z <= y - t. Proof. by move=> lxy ltz; rewrite ler_add // lter_opp2. Qed. @@ -1951,7 +1935,7 @@ Lemma paddr_eq0 (x y : R) : 0 <= x -> 0 <= y -> (x + y == 0) = (x == 0) && (y == 0). Proof. rewrite le0r; case/orP=> [/eqP->|hx]; first by rewrite add0r eqxx. -by rewrite (gtr_eqF hx) /= => hy; rewrite gtr_eqF // ltr_spaddl. +by rewrite (gt_eqF hx) /= => hy; rewrite gt_eqF // ltr_spaddl. Qed. Lemma naddr_eq0 (x y : R) : @@ -1973,7 +1957,7 @@ Proof. exact: (big_ind _ _ (@ler_paddl 0)). Qed. Lemma ler_sum I (r : seq I) (P : pred I) (F G : I -> R) : (forall i, P i -> F i <= G i) -> \sum_(i <- r | P i) F i <= \sum_(i <- r | P i) G i. -Proof. exact: (big_ind2 _ (lerr _) ler_add). Qed. +Proof. exact: (big_ind2 _ (lexx _) ler_add). Qed. Lemma psumr_eq0 (I : eqType) (r : seq I) (P : pred I) (F : I -> R) : (forall i, P i -> 0 <= F i) -> @@ -2000,7 +1984,7 @@ by move=> x_gt0 y z /=; rewrite -subr_ge0 -mulrBr pmulr_rge0 // subr_ge0. Qed. Lemma ltr_pmul2l x : 0 < x -> {mono *%R x : x y / x < y}. -Proof. by move=> x_gt0; apply: lerW_mono (ler_pmul2l _). Qed. +Proof. by move=> x_gt0; apply: leW_mono (ler_pmul2l _). Qed. Definition lter_pmul2l := (ler_pmul2l, ltr_pmul2l). @@ -2008,7 +1992,7 @@ Lemma ler_pmul2r x : 0 < x -> {mono *%R^~ x : x y / x <= y}. Proof. by move=> x_gt0 y z /=; rewrite ![_ * x]mulrC ler_pmul2l. Qed. Lemma ltr_pmul2r x : 0 < x -> {mono *%R^~ x : x y / x < y}. -Proof. by move=> x_gt0; apply: lerW_mono (ler_pmul2r _). Qed. +Proof. by move=> x_gt0; apply: leW_mono (ler_pmul2r _). Qed. Definition lter_pmul2r := (ler_pmul2r, ltr_pmul2r). @@ -2018,7 +2002,7 @@ by move=> x_lt0 y z /=; rewrite -ler_opp2 -!mulNr ler_pmul2l ?oppr_gt0. Qed. Lemma ltr_nmul2l x : x < 0 -> {mono *%R x : x y /~ x < y}. -Proof. by move=> x_lt0; apply: lerW_nmono (ler_nmul2l _). Qed. +Proof. by move=> x_lt0; apply: leW_nmono (ler_nmul2l _). Qed. Definition lter_nmul2l := (ler_nmul2l, ltr_nmul2l). @@ -2026,7 +2010,7 @@ Lemma ler_nmul2r x : x < 0 -> {mono *%R^~ x : x y /~ x <= y}. Proof. by move=> x_lt0 y z /=; rewrite ![_ * x]mulrC ler_nmul2l. Qed. Lemma ltr_nmul2r x : x < 0 -> {mono *%R^~ x : x y /~ x < y}. -Proof. by move=> x_lt0; apply: lerW_nmono (ler_nmul2r _). Qed. +Proof. by move=> x_lt0; apply: leW_nmono (ler_nmul2r _). Qed. Definition lter_nmul2r := (ler_nmul2r, ltr_nmul2r). @@ -2053,15 +2037,15 @@ Qed. Lemma ler_pmul x1 y1 x2 y2 : 0 <= x1 -> 0 <= x2 -> x1 <= y1 -> x2 <= y2 -> x1 * x2 <= y1 * y2. Proof. -move=> x1ge0 x2ge0 le_xy1 le_xy2; have y1ge0 := ler_trans x1ge0 le_xy1. -exact: ler_trans (ler_wpmul2r x2ge0 le_xy1) (ler_wpmul2l y1ge0 le_xy2). +move=> x1ge0 x2ge0 le_xy1 le_xy2; have y1ge0 := le_trans x1ge0 le_xy1. +exact: le_trans (ler_wpmul2r x2ge0 le_xy1) (ler_wpmul2l y1ge0 le_xy2). Qed. Lemma ltr_pmul x1 y1 x2 y2 : 0 <= x1 -> 0 <= x2 -> x1 < y1 -> x2 < y2 -> x1 * x2 < y1 * y2. Proof. -move=> x1ge0 x2ge0 lt_xy1 lt_xy2; have y1gt0 := ler_lt_trans x1ge0 lt_xy1. -by rewrite (ler_lt_trans (ler_wpmul2r x2ge0 (ltrW lt_xy1))) ?ltr_pmul2l. +move=> x1ge0 x2ge0 lt_xy1 lt_xy2; have y1gt0 := le_lt_trans x1ge0 lt_xy1. +by rewrite (le_lt_trans (ler_wpmul2r x2ge0 (ltW lt_xy1))) ?ltr_pmul2l. Qed. (* complement for x *+ n and <= or < *) @@ -2072,10 +2056,10 @@ by case: n => // n _ x y /=; rewrite -mulr_natl -[y *+ _]mulr_natl ler_pmul2l. Qed. Lemma ltr_pmuln2r n : (0 < n)%N -> {mono (@GRing.natmul R)^~ n : x y / x < y}. -Proof. by move/ler_pmuln2r/lerW_mono. Qed. +Proof. by move/ler_pmuln2r/leW_mono. Qed. Lemma pmulrnI n : (0 < n)%N -> injective ((@GRing.natmul R)^~ n). -Proof. by move/ler_pmuln2r/incr_inj. Qed. +Proof. by move/ler_pmuln2r/inc_inj. Qed. Lemma eqr_pmuln2r n : (0 < n)%N -> {mono (@GRing.natmul R)^~ n : x y / x == y}. Proof. by move/pmulrnI/inj_eq. Qed. @@ -2108,13 +2092,13 @@ Lemma mulrn_wle0 x n : x <= 0 -> x *+ n <= 0. Proof. by move=> /(ler_wmuln2r n); rewrite mul0rn. Qed. Lemma ler_muln2r n x y : (x *+ n <= y *+ n) = ((n == 0%N) || (x <= y)). -Proof. by case: n => [|n]; rewrite ?lerr ?eqxx // ler_pmuln2r. Qed. +Proof. by case: n => [|n]; rewrite ?lexx ?eqxx // ler_pmuln2r. Qed. Lemma ltr_muln2r n x y : (x *+ n < y *+ n) = ((0 < n)%N && (x < y)). -Proof. by case: n => [|n]; rewrite ?lerr ?eqxx // ltr_pmuln2r. Qed. +Proof. by case: n => [|n]; rewrite ?lexx ?eqxx // ltr_pmuln2r. Qed. Lemma eqr_muln2r n x y : (x *+ n == y *+ n) = (n == 0)%N || (x == y). -Proof. by rewrite !eqr_le !ler_muln2r -orb_andr. Qed. +Proof. by rewrite !(@eq_le _ R) !ler_muln2r -orb_andr. Qed. (* More characteristic zero properties. *) @@ -2147,14 +2131,14 @@ Proof. by case: n => // n hx; rewrite pmulrn_llt0. Qed. Lemma ler_pmuln2l x : 0 < x -> {mono (@GRing.natmul R x) : m n / (m <= n)%N >-> m <= n}. Proof. -move=> x_gt0 m n /=; case: leqP => hmn; first by rewrite ler_wpmuln2l // ltrW. -rewrite -(subnK (ltnW hmn)) mulrnDr ger_addr ltr_geF //. +move=> x_gt0 m n /=; case: leqP => hmn; first by rewrite ler_wpmuln2l // ltW. +rewrite -(subnK (ltnW hmn)) mulrnDr ger_addr lt_geF //. by rewrite mulrn_wgt0 // subn_gt0. Qed. Lemma ltr_pmuln2l x : 0 < x -> {mono (@GRing.natmul R x) : m n / (m < n)%N >-> m < n}. -Proof. by move=> x_gt0; apply: lenrW_mono (ler_pmuln2l _). Qed. +Proof. by move=> x_gt0; apply: leW_mono (ler_pmuln2l _). Qed. Lemma ler_nmuln2l x : x < 0 -> {mono (@GRing.natmul R x) : m n / (n <= m)%N >-> m <= n}. @@ -2164,7 +2148,7 @@ Qed. Lemma ltr_nmuln2l x : x < 0 -> {mono (@GRing.natmul R x) : m n / (n < m)%N >-> m < n}. -Proof. by move=> x_lt0; apply: lenrW_nmono (ler_nmuln2l _). Qed. +Proof. by move=> x_lt0; apply: leW_nmono (ler_nmuln2l _). Qed. Lemma ler_nat m n : (m%:R <= n%:R :> R) = (m <= n)%N. Proof. by rewrite ler_pmuln2l. Qed. @@ -2191,10 +2175,10 @@ Lemma ltrn1 n : n%:R < 1 :> R = (n < 1)%N. Proof. by rewrite -ltr_nat. Qed. Lemma ltrN10 : -1 < 0 :> R. Proof. by rewrite oppr_lt0. Qed. Lemma lerN10 : -1 <= 0 :> R. Proof. by rewrite oppr_le0. Qed. -Lemma ltr10 : 1 < 0 :> R = false. Proof. by rewrite ler_gtF. Qed. -Lemma ler10 : 1 <= 0 :> R = false. Proof. by rewrite ltr_geF. Qed. -Lemma ltr0N1 : 0 < -1 :> R = false. Proof. by rewrite ler_gtF // lerN10. Qed. -Lemma ler0N1 : 0 <= -1 :> R = false. Proof. by rewrite ltr_geF // ltrN10. Qed. +Lemma ltr10 : 1 < 0 :> R = false. Proof. by rewrite le_gtF. Qed. +Lemma ler10 : 1 <= 0 :> R = false. Proof. by rewrite lt_geF. Qed. +Lemma ltr0N1 : 0 < -1 :> R = false. Proof. by rewrite le_gtF // lerN10. Qed. +Lemma ler0N1 : 0 <= -1 :> R = false. Proof. by rewrite lt_geF // ltrN10. Qed. Lemma pmulrn_rgt0 x n : 0 < x -> 0 < x *+ n = (0 < n)%N. Proof. by move=> x_gt0; rewrite -(mulr0n x) ltr_pmuln2l. Qed. @@ -2309,9 +2293,9 @@ Lemma ltr_prod I r (P : pred I) (E1 E2 : I -> R) : Proof. elim: r => //= i r IHr; rewrite !big_cons; case: ifP => {IHr}// Pi _ ltE12. have /andP[le0E1i ltE12i] := ltE12 i Pi; set E2r := \prod_(j <- r | P j) E2 j. -apply: ler_lt_trans (_ : E1 i * E2r < E2 i * E2r). - by rewrite ler_wpmul2l ?ler_prod // => j /ltE12/andP[-> /ltrW]. -by rewrite ltr_pmul2r ?prodr_gt0 // => j /ltE12/andP[le0E1j /ler_lt_trans->]. +apply: le_lt_trans (_ : E1 i * E2r < E2 i * E2r). + by rewrite ler_wpmul2l ?ler_prod // => j /ltE12/andP[-> /ltW]. +by rewrite ltr_pmul2r ?prodr_gt0 // => j /ltE12/andP[le0E1j /le_lt_trans->]. Qed. Lemma ltr_prod_nat (E1 E2 : nat -> R) (n m : nat) : @@ -2326,7 +2310,7 @@ Qed. Lemma realMr x y : x != 0 -> x \is real -> (x * y \is real) = (y \is real). Proof. -move=> x_neq0 xR; case: real_ltrgtP x_neq0 => // hx _; rewrite !realE. +move=> x_neq0 xR; case: real_ltgtP x_neq0 => // hx _; rewrite !realE. by rewrite nmulr_rge0 // nmulr_rle0 // orbC. by rewrite pmulr_rge0 // pmulr_rle0 // orbC. Qed. @@ -2418,21 +2402,21 @@ Lemma ler_nimulr x y : y <= 0 -> x <= 1 -> y <= y * x. Proof. by move=> hx hy; rewrite -{1}[y]mulr1 ler_wnmul2l. Qed. Lemma mulr_ile1 x y : 0 <= x -> 0 <= y -> x <= 1 -> y <= 1 -> x * y <= 1. -Proof. by move=> *; rewrite (@ler_trans _ y) ?ler_pimull. Qed. +Proof. by move=> *; rewrite (@le_trans _ _ y) ?ler_pimull. Qed. Lemma mulr_ilt1 x y : 0 <= x -> 0 <= y -> x < 1 -> y < 1 -> x * y < 1. -Proof. by move=> *; rewrite (@ler_lt_trans _ y) ?ler_pimull // ltrW. Qed. +Proof. by move=> *; rewrite (@le_lt_trans _ _ y) ?ler_pimull // ltW. Qed. Definition mulr_ilte1 := (mulr_ile1, mulr_ilt1). Lemma mulr_ege1 x y : 1 <= x -> 1 <= y -> 1 <= x * y. Proof. -by move=> le1x le1y; rewrite (@ler_trans _ y) ?ler_pemull // (ler_trans ler01). +by move=> le1x le1y; rewrite (@le_trans _ _ y) ?ler_pemull // (le_trans ler01). Qed. Lemma mulr_egt1 x y : 1 < x -> 1 < y -> 1 < x * y. Proof. -by move=> le1x lt1y; rewrite (@ltr_trans _ y) // ltr_pmull // (ltr_trans ltr01). +by move=> le1x lt1y; rewrite (@lt_trans _ _ y) // ltr_pmull // (lt_trans ltr01). Qed. Definition mulr_egte1 := (mulr_ege1, mulr_egt1). Definition mulr_cp1 := (mulr_ilte1, mulr_egte1). @@ -2489,7 +2473,7 @@ Qed. Lemma exprn_ilt1 n x : 0 <= x -> x < 1 -> x ^+ n < 1 = (n != 0%N). Proof. move=> xge0 xlt1. -case: n; [by rewrite eqxx ltrr | elim=> [|n ihn]; first by rewrite expr1]. +case: n; [by rewrite eqxx ltxx | elim=> [|n ihn]; first by rewrite expr1]. by rewrite exprS mulr_ilt1 // exprn_ge0. Qed. @@ -2502,7 +2486,7 @@ Qed. Lemma exprn_egt1 n x : 1 < x -> 1 < x ^+ n = (n != 0%N). Proof. -move=> xgt1; case: n; first by rewrite eqxx ltrr. +move=> xgt1; case: n; first by rewrite eqxx ltxx. elim=> [|n ihn]; first by rewrite expr1. by rewrite exprS mulr_egt1 // exprn_ge0. Qed. @@ -2515,8 +2499,8 @@ Proof. by case: n => n // *; rewrite exprS ler_pimulr // exprn_ile1. Qed. Lemma ltr_iexpr x n : 0 < x -> x < 1 -> (x ^+ n < x) = (1 < n)%N. Proof. -case: n=> [|[|n]] //; first by rewrite expr0 => _ /ltr_gtF ->. -by move=> x0 x1; rewrite exprS gtr_pmulr // ?exprn_ilt1 // ltrW. +case: n=> [|[|n]] //; first by rewrite expr0 => _ /lt_gtF ->. +by move=> x0 x1; rewrite exprS gtr_pmulr // ?exprn_ilt1 // ltW. Qed. Definition lter_iexpr := (ler_iexpr, ltr_iexpr). @@ -2524,13 +2508,13 @@ Definition lter_iexpr := (ler_iexpr, ltr_iexpr). Lemma ler_eexpr x n : (0 < n)%N -> 1 <= x -> x <= x ^+ n. Proof. case: n => // n _ x_ge1. -by rewrite exprS ler_pemulr ?(ler_trans _ x_ge1) // exprn_ege1. +by rewrite exprS ler_pemulr ?(le_trans _ x_ge1) // exprn_ege1. Qed. Lemma ltr_eexpr x n : 1 < x -> (x < x ^+ n) = (1 < n)%N. Proof. -move=> x_ge1; case: n=> [|[|n]] //; first by rewrite expr0 ltr_gtF. -by rewrite exprS ltr_pmulr ?(ltr_trans _ x_ge1) ?exprn_egt1. +move=> x_ge1; case: n=> [|[|n]] //; first by rewrite expr0 lt_gtF. +by rewrite exprS ltr_pmulr ?(lt_trans _ x_ge1) ?exprn_egt1. Qed. Definition lter_eexpr := (ler_eexpr, ltr_eexpr). @@ -2547,15 +2531,15 @@ Lemma ler_weexpn2l x : 1 <= x -> {homo (GRing.exp x) : m n / (m <= n)%N >-> m <= n}. Proof. move=> xge1 m n /= hmn; rewrite -(subnK hmn) exprD. -by rewrite ler_pemull ?(exprn_ge0, exprn_ege1) // (ler_trans _ xge1) ?ler01. +by rewrite ler_pemull ?(exprn_ge0, exprn_ege1) // (le_trans _ xge1) ?ler01. Qed. Lemma ieexprn_weq1 x n : 0 <= x -> (x ^+ n == 1) = ((n == 0%N) || (x == 1)). Proof. move=> xle0; case: n => [|n]; first by rewrite expr0 eqxx. -case: (@real_ltrgtP x 1); do ?by rewrite ?ger0_real. -+ by move=> x_lt1; rewrite ?ltr_eqF // exprn_ilt1. -+ by move=> x_lt1; rewrite ?gtr_eqF // exprn_egt1. +case: (@real_ltgtP x 1); do ?by rewrite ?ger0_real. ++ by move=> x_lt1; rewrite 1?lt_eqF // exprn_ilt1. ++ by move=> x_lt1; rewrite 1?gt_eqF // exprn_egt1. by move->; rewrite expr1n eqxx. Qed. @@ -2564,50 +2548,50 @@ Proof. move=> x_gt0 x_neq1 m n; without loss /subnK <-: m n / (n <= m)%N. by move=> IH eq_xmn; case/orP: (leq_total m n) => /IH->. case: {m}(m - n)%N => // m /eqP/idPn[]; rewrite -[x ^+ n]mul1r exprD. -by rewrite (inj_eq (mulIf _)) ?ieexprn_weq1 ?ltrW // expf_neq0 ?gtr_eqF. +by rewrite (inj_eq (mulIf _)) ?ieexprn_weq1 ?ltW // expf_neq0 ?gt_eqF. Qed. Lemma ler_iexpn2l x : 0 < x -> x < 1 -> {mono (GRing.exp x) : m n / (n <= m)%N >-> m <= n}. Proof. -move=> xgt0 xlt1; apply: (lenr_nmono (inj_nhomo_ltnr _ _)); last first. - by apply: ler_wiexpn2l; rewrite ltrW. -by apply: ieexprIn; rewrite ?ltr_eqF ?ltr_cpable. +move=> xgt0 xlt1; apply: (le_nmono (inj_nhomo_lt _ _)); last first. + by apply: ler_wiexpn2l; rewrite ltW. +by apply: ieexprIn; rewrite ?lt_eqF ?ltr_cpable. Qed. Lemma ltr_iexpn2l x : 0 < x -> x < 1 -> {mono (GRing.exp x) : m n / (n < m)%N >-> m < n}. -Proof. by move=> xgt0 xlt1; apply: (lenrW_nmono (ler_iexpn2l _ _)). Qed. +Proof. by move=> xgt0 xlt1; apply: (leW_nmono (ler_iexpn2l _ _)). Qed. Definition lter_iexpn2l := (ler_iexpn2l, ltr_iexpn2l). Lemma ler_eexpn2l x : 1 < x -> {mono (GRing.exp x) : m n / (m <= n)%N >-> m <= n}. Proof. -move=> xgt1; apply: (lenr_mono (inj_homo_ltnr _ _)); last first. - by apply: ler_weexpn2l; rewrite ltrW. -by apply: ieexprIn; rewrite ?gtr_eqF ?gtr_cpable //; apply: ltr_trans xgt1. +move=> xgt1; apply: (le_mono (inj_homo_lt _ _)); last first. + by apply: ler_weexpn2l; rewrite ltW. +by apply: ieexprIn; rewrite ?gt_eqF ?gtr_cpable //; apply: lt_trans xgt1. Qed. Lemma ltr_eexpn2l x : 1 < x -> {mono (GRing.exp x) : m n / (m < n)%N >-> m < n}. -Proof. by move=> xgt1; apply: (lenrW_mono (ler_eexpn2l _)). Qed. +Proof. by move=> xgt1; apply: (leW_mono (ler_eexpn2l _)). Qed. Definition lter_eexpn2l := (ler_eexpn2l, ltr_eexpn2l). Lemma ltr_expn2r n x y : 0 <= x -> x < y -> x ^+ n < y ^+ n = (n != 0%N). Proof. -move=> xge0 xlty; case: n; first by rewrite ltrr. +move=> xge0 xlty; case: n; first by rewrite ltxx. elim=> [|n IHn]; rewrite ?[_ ^+ _.+2]exprS //. -rewrite (@ler_lt_trans _ (x * y ^+ n.+1)) ?ler_wpmul2l ?ltr_pmul2r ?IHn //. - by rewrite ltrW // ihn. -by rewrite exprn_gt0 // (ler_lt_trans xge0). +rewrite (@le_lt_trans _ _ (x * y ^+ n.+1)) ?ler_wpmul2l ?ltr_pmul2r ?IHn //. + by rewrite ltW // ihn. +by rewrite exprn_gt0 // (le_lt_trans xge0). Qed. Lemma ler_expn2r n : {in nneg & , {homo ((@GRing.exp R)^~ n) : x y / x <= y}}. Proof. move=> x y /= x0 y0 xy; elim: n => [|n IHn]; rewrite !(expr0, exprS) //. -by rewrite (@ler_trans _ (x * y ^+ n)) ?ler_wpmul2l ?ler_wpmul2r ?exprn_ge0. +by rewrite (@le_trans _ _ (x * y ^+ n)) ?ler_wpmul2l ?ler_wpmul2r ?exprn_ge0. Qed. Definition lter_expn2r := (ler_expn2r, ltr_expn2r). @@ -2629,13 +2613,13 @@ Qed. Lemma ltr_pexpn2r n : (0 < n)%N -> {in nneg & , {mono ((@GRing.exp R)^~ n) : x y / x < y}}. Proof. -by move=> n_gt0 x y x_ge0 y_ge0; rewrite !ltr_neqAle !eqr_le !ler_pexpn2r. +by move=> n_gt0 x y x_ge0 y_ge0; rewrite !lt_neqAle !eq_le !ler_pexpn2r. Qed. Definition lter_pexpn2r := (ler_pexpn2r, ltr_pexpn2r). Lemma pexpIrn n : (0 < n)%N -> {in nneg &, injective ((@GRing.exp R)^~ n)}. -Proof. by move=> n_gt0; apply: incr_inj_in (ler_pexpn2r _). Qed. +Proof. by move=> n_gt0; apply: inc_inj_in (ler_pexpn2r _). Qed. (* expr and ler/ltr *) Lemma expr_le1 n x : (0 < n)%N -> 0 <= x -> (x ^+ n <= 1) = (x <= 1). @@ -2663,7 +2647,7 @@ Qed. Definition expr_gte1 := (expr_ge1, expr_gt1). Lemma pexpr_eq1 x n : (0 < n)%N -> 0 <= x -> (x ^+ n == 1) = (x == 1). -Proof. by move=> ngt0 xge0; rewrite !eqr_le expr_le1 // expr_ge1. Qed. +Proof. by move=> ngt0 xge0; rewrite !eq_le expr_le1 // expr_ge1. Qed. Lemma pexprn_eq1 x n : 0 <= x -> (x ^+ n == 1) = (n == 0%N) || (x == 1). Proof. by case: n => [|n] xge0; rewrite ?eqxx // pexpr_eq1 ?gtn_eqF. Qed. @@ -2702,11 +2686,11 @@ Qed. Lemma ltr_pinv : {in [pred x in GRing.unit | 0 < x] &, {mono (@GRing.inv R) : x y /~ x < y}}. -Proof. exact: lerW_nmono_in ler_pinv. Qed. +Proof. exact: leW_nmono_in ler_pinv. Qed. Lemma ltr_ninv : {in [pred x in GRing.unit | x < 0] &, {mono (@GRing.inv R) : x y /~ x < y}}. -Proof. exact: lerW_nmono_in ler_ninv. Qed. +Proof. exact: leW_nmono_in ler_ninv. Qed. Lemma invr_gt1 x : x \is a GRing.unit -> 0 < x -> (1 < x^-1) = (x < 1). Proof. @@ -2734,44 +2718,61 @@ Definition invr_cp1 := (invr_gte1, invr_lte1). Lemma real_ler_norm x : x \is real -> x <= `|x|. Proof. -by case/real_ger0P=> hx //; rewrite (ler_trans (ltrW hx)) // oppr_ge0 ltrW. +by case/real_ge0P=> hx //; rewrite (le_trans (ltW hx)) // oppr_ge0 ltW. Qed. (* norm + add *) -Lemma normr_real x : `|x| \is real. Proof. by rewrite ger0_real. Qed. +Section NormedZmoduleTheory. + +Variable V : normedZmodType R. +Implicit Types (u v w : V). + +Lemma normr_real v : `|v| \is real. Proof. by apply/ger0_real. Qed. Hint Resolve normr_real : core. -Lemma ler_norm_sum I r (G : I -> R) (P : pred I): +Lemma ler_norm_sum I r (G : I -> V) (P : pred I): `|\sum_(i <- r | P i) G i| <= \sum_(i <- r | P i) `|G i|. Proof. elim/big_rec2: _ => [|i y x _]; first by rewrite normr0. -by rewrite -(ler_add2l `|G i|); apply: ler_trans; apply: ler_norm_add. +by rewrite -(ler_add2l `|G i|); apply: le_trans; apply: ler_norm_add. Qed. -Lemma ler_norm_sub x y : `|x - y| <= `|x| + `|y|. -Proof. by rewrite (ler_trans (ler_norm_add _ _)) ?normrN. Qed. +Lemma ler_norm_sub v w : `|v - w| <= `|v| + `|w|. +Proof. by rewrite (le_trans (ler_norm_add _ _)) ?normrN. Qed. -Lemma ler_dist_add z x y : `|x - y| <= `|x - z| + `|z - y|. -Proof. by rewrite (ler_trans _ (ler_norm_add _ _)) // addrA addrNK. Qed. +Lemma ler_dist_add u v w : `|v - w| <= `|v - u| + `|u - w|. +Proof. by rewrite (le_trans _ (ler_norm_add _ _)) // addrA addrNK. Qed. -Lemma ler_sub_norm_add x y : `|x| - `|y| <= `|x + y|. +Lemma ler_sub_norm_add v w : `|v| - `|w| <= `|v + w|. Proof. -rewrite -{1}[x](addrK y) lter_sub_addl. -by rewrite (ler_trans (ler_norm_add _ _)) // addrC normrN. +rewrite -{1}[v](addrK w) lter_sub_addl. +by rewrite (le_trans (ler_norm_add _ _)) // addrC normrN. Qed. -Lemma ler_sub_dist x y : `|x| - `|y| <= `|x - y|. -Proof. by rewrite -[`|y|]normrN ler_sub_norm_add. Qed. +Lemma ler_sub_dist v w : `|v| - `|w| <= `|v - w|. +Proof. by rewrite -[`|w|]normrN ler_sub_norm_add. Qed. -Lemma ler_dist_dist x y : `|`|x| - `|y| | <= `|x - y|. +Lemma ler_dist_dist v w : `| `|v| - `|w| | <= `|v - w|. Proof. -have [||_|_] // := @real_lerP `|x| `|y|; last by rewrite ler_sub_dist. +have [||_|_] // := @real_leP `|v| `|w|; last by rewrite ler_sub_dist. by rewrite distrC ler_sub_dist. Qed. -Lemma ler_dist_norm_add x y : `| `|x| - `|y| | <= `| x + y |. -Proof. by rewrite -[y]opprK normrN ler_dist_dist. Qed. +Lemma ler_dist_norm_add v w : `| `|v| - `|w| | <= `|v + w|. +Proof. by rewrite -[w]opprK normrN ler_dist_dist. Qed. + +Lemma ler_nnorml v x : x < 0 -> `|v| <= x = false. +Proof. by move=> h; rewrite lt_geF //; apply/(lt_le_trans h). Qed. + +Lemma ltr_nnorml v x : x <= 0 -> `|v| < x = false. +Proof. by move=> h; rewrite le_gtF //; apply/(le_trans h). Qed. + +Definition lter_nnormr := (ler_nnorml, ltr_nnorml). + +End NormedZmoduleTheory. + +Hint Extern 0 (is_true (norm _ \is real)) => exact: normr_real : core. Lemma real_ler_norml x y : x \is real -> (`|x| <= y) = (- y <= x <= y). Proof. @@ -2779,7 +2780,7 @@ move=> xR; wlog x_ge0 : x xR / 0 <= x => [hwlog|]. move: (xR) => /(@real_leVge 0) /orP [|/hwlog->|hx] //. by rewrite -[x]opprK normrN ler_opp2 andbC ler_oppl hwlog ?realN ?oppr_ge0. rewrite ger0_norm //; have [le_xy|] := boolP (x <= y); last by rewrite andbF. -by rewrite (ler_trans _ x_ge0) // oppr_le0 (ler_trans x_ge0). +by rewrite (le_trans _ x_ge0) // oppr_le0 (le_trans x_ge0). Qed. Lemma real_ler_normlP x y : @@ -2794,16 +2795,16 @@ Lemma real_eqr_norml x y : Proof. move=> Rx. apply/idP/idP=> [|/andP[/pred2P[]-> /ger0_norm/eqP]]; rewrite ?normrE //. -case: real_ler0P => // hx; rewrite 1?eqr_oppLR => /eqP exy. +case: real_le0P => // hx; rewrite 1?eqr_oppLR => /eqP exy. by move: hx; rewrite exy ?oppr_le0 eqxx orbT //. -by move: hx=> /ltrW; rewrite exy eqxx. +by move: hx=> /ltW; rewrite exy eqxx. Qed. Lemma real_eqr_norm2 x y : x \is real -> y \is real -> (`|x| == `|y|) = (x == y) || (x == -y). Proof. move=> Rx Ry; rewrite real_eqr_norml // normrE andbT. -by case: real_ler0P; rewrite // opprK orbC. +by case: real_le0P; rewrite // opprK orbC. Qed. Lemma real_ltr_norml x y : x \is real -> (`|x| < y) = (- y < x < y). @@ -2812,7 +2813,7 @@ move=> Rx; wlog x_ge0 : x Rx / 0 <= x => [hwlog|]. move: (Rx) => /(@real_leVge 0) /orP [|/hwlog->|hx] //. by rewrite -[x]opprK normrN ltr_opp2 andbC ltr_oppl hwlog ?realN ?oppr_ge0. rewrite ger0_norm //; have [le_xy|] := boolP (x < y); last by rewrite andbF. -by rewrite (ltr_le_trans _ x_ge0) // oppr_lt0 (ler_lt_trans x_ge0). +by rewrite (lt_le_trans _ x_ge0) // oppr_lt0 (le_lt_trans x_ge0). Qed. Definition real_lter_norml := (real_ler_norml, real_ltr_norml). @@ -2829,7 +2830,7 @@ Lemma real_ler_normr x y : y \is real -> (x <= `|y|) = (x <= y) || (x <= - y). Proof. move=> Ry. have [xR|xNR] := boolP (x \is real); last by rewrite ?Nreal_leF ?realN. -rewrite real_lerNgt ?real_ltr_norml // negb_and -?real_lerNgt ?realN //. +rewrite real_leNgt ?real_ltr_norml // negb_and -?real_leNgt ?realN //. by rewrite orbC ler_oppr. Qed. @@ -2837,20 +2838,12 @@ Lemma real_ltr_normr x y : y \is real -> (x < `|y|) = (x < y) || (x < - y). Proof. move=> Ry. have [xR|xNR] := boolP (x \is real); last by rewrite ?Nreal_ltF ?realN. -rewrite real_ltrNge ?real_ler_norml // negb_and -?real_ltrNge ?realN //. +rewrite real_ltNge ?real_ler_norml // negb_and -?real_ltNge ?realN //. by rewrite orbC ltr_oppr. Qed. Definition real_lter_normr := (real_ler_normr, real_ltr_normr). -Lemma ler_nnorml x y : y < 0 -> `|x| <= y = false. -Proof. by move=> y_lt0; rewrite ltr_geF // (ltr_le_trans y_lt0). Qed. - -Lemma ltr_nnorml x y : y <= 0 -> `|x| < y = false. -Proof. by move=> y_le0; rewrite ler_gtF // (ler_trans y_le0). Qed. - -Definition lter_nnormr := (ler_nnorml, ltr_nnorml). - Lemma real_ler_distl x y e : x - y \is real -> (`|x - y| <= e) = (y - e <= x <= y + e). Proof. by move=> Rxy; rewrite real_lter_norml // !lter_sub_addl. Qed. @@ -2868,9 +2861,9 @@ Definition eqr_norm_idVN := =^~ (ger0_def, ler0_def). Lemma real_exprn_even_ge0 n x : x \is real -> ~~ odd n -> 0 <= x ^+ n. Proof. -move=> xR even_n; have [/exprn_ge0 -> //|x_lt0] := real_ger0P xR. +move=> xR even_n; have [/exprn_ge0 -> //|x_lt0] := real_ge0P xR. rewrite -[x]opprK -mulN1r exprMn -signr_odd (negPf even_n) expr0 mul1r. -by rewrite exprn_ge0 ?oppr_ge0 ?ltrW. +by rewrite exprn_ge0 ?oppr_ge0 ?ltW. Qed. Lemma real_exprn_even_gt0 n x : @@ -2883,21 +2876,21 @@ Qed. Lemma real_exprn_even_le0 n x : x \is real -> ~~ odd n -> (x ^+ n <= 0) = (n != 0%N) && (x == 0). Proof. -move=> xR n_even; rewrite !real_lerNgt ?rpred0 ?rpredX //. +move=> xR n_even; rewrite !real_leNgt ?rpred0 ?rpredX //. by rewrite real_exprn_even_gt0 // negb_or negbK. Qed. Lemma real_exprn_even_lt0 n x : x \is real -> ~~ odd n -> (x ^+ n < 0) = false. -Proof. by move=> xR n_even; rewrite ler_gtF // real_exprn_even_ge0. Qed. +Proof. by move=> xR n_even; rewrite le_gtF // real_exprn_even_ge0. Qed. Lemma real_exprn_odd_ge0 n x : x \is real -> odd n -> (0 <= x ^+ n) = (0 <= x). Proof. -case/real_ger0P => [x_ge0|x_lt0] n_odd; first by rewrite exprn_ge0. -apply: negbTE; rewrite ltr_geF //. +case/real_ge0P => [x_ge0|x_lt0] n_odd; first by rewrite exprn_ge0. +apply: negbTE; rewrite lt_geF //. case: n n_odd => // n /= n_even; rewrite exprS pmulr_llt0 //. -by rewrite real_exprn_even_gt0 ?ler0_real ?ltrW // ltr_eqF ?orbT. +by rewrite real_exprn_even_gt0 ?ler0_real ?ltW // (lt_eqF x_lt0) ?orbT. Qed. Lemma real_exprn_odd_gt0 n x : x \is real -> odd n -> (0 < x ^+ n) = (0 < x). @@ -2907,12 +2900,12 @@ Qed. Lemma real_exprn_odd_le0 n x : x \is real -> odd n -> (x ^+ n <= 0) = (x <= 0). Proof. -by move=> xR n_odd; rewrite !real_lerNgt ?rpred0 ?rpredX // real_exprn_odd_gt0. +by move=> xR n_odd; rewrite !real_leNgt ?rpred0 ?rpredX // real_exprn_odd_gt0. Qed. Lemma real_exprn_odd_lt0 n x : x \is real -> odd n -> (x ^+ n < 0) = (x < 0). Proof. -by move=> xR n_odd; rewrite !real_ltrNge ?rpred0 ?rpredX // real_exprn_odd_ge0. +by move=> xR n_odd; rewrite !real_ltNge ?rpred0 ?rpredX // real_exprn_odd_ge0. Qed. (* GG: Could this be a better definition of "real" ? *) @@ -2924,7 +2917,7 @@ Proof. by move=> Rx; rewrite -normrX ger0_norm -?realEsqr. Qed. (* Binary sign ((-1) ^+ s). *) -Lemma normr_sign s : `|(-1) ^+ s| = 1 :> R. +Lemma normr_sign s : `|(-1) ^+ s : R| = 1. Proof. by rewrite normrX normrN1 expr1n. Qed. Lemma normrMsign s x : `|(-1) ^+ s * x| = `|x|. @@ -2940,7 +2933,7 @@ Lemma signr_ge0 (b : bool) : (0 <= (-1) ^+ b :> R) = ~~ b. Proof. by rewrite le0r signr_eq0 signr_gt0. Qed. Lemma signr_le0 (b : bool) : ((-1) ^+ b <= 0 :> R) = b. -Proof. by rewrite ler_eqVlt signr_eq0 signr_lt0. Qed. +Proof. by rewrite le_eqVlt signr_eq0 signr_lt0. Qed. (* This actually holds for char R != 2. *) Lemma signr_inj : injective (fun b : bool => (-1) ^+ b : R). @@ -2955,10 +2948,10 @@ Lemma neqr0_sign x : x != 0 -> (-1) ^+ (x < 0)%R = sgr x. Proof. by rewrite sgr_def => ->. Qed. Lemma gtr0_sg x : 0 < x -> sg x = 1. -Proof. by move=> x_gt0; rewrite /sg gtr_eqF // ltr_gtF. Qed. +Proof. by move=> x_gt0; rewrite /sg gt_eqF // lt_gtF. Qed. Lemma ltr0_sg x : x < 0 -> sg x = -1. -Proof. by move=> x_lt0; rewrite /sg x_lt0 ltr_eqF. Qed. +Proof. by move=> x_lt0; rewrite /sg x_lt0 lt_eqF. Qed. Lemma sgr0 : sg 0 = 0 :> R. Proof. by rewrite /sgr eqxx. Qed. Lemma sgr1 : sg 1 = 1 :> R. Proof. by rewrite gtr0_sg // ltr01. Qed. @@ -3002,16 +2995,16 @@ Proof. by rewrite !(fun_if sg) !sgrE. Qed. Lemma sgr_lt0 x : (sg x < 0) = (x < 0). Proof. rewrite /sg; case: eqP => [-> // | _]. -by case: ifP => _; rewrite ?ltrN10 // ltr_gtF. +by case: ifP => _; rewrite ?ltrN10 // lt_gtF. Qed. Lemma sgr_le0 x : (sgr x <= 0) = (x <= 0). -Proof. by rewrite !ler_eqVlt sgr_eq0 sgr_lt0. Qed. +Proof. by rewrite !le_eqVlt sgr_eq0 sgr_lt0. Qed. (* sign and norm *) Lemma realEsign x : x \is real -> x = (-1) ^+ (x < 0)%R * `|x|. -Proof. by case/real_ger0P; rewrite (mul1r, mulN1r) ?opprK. Qed. +Proof. by case/real_ge0P; rewrite (mul1r, mulN1r) ?opprK. Qed. Lemma realNEsign x : x \is real -> - x = (-1) ^+ (0 < x)%R * `|x|. Proof. by move=> Rx; rewrite -normrN -oppr_lt0 -realEsign ?rpredN. Qed. @@ -3036,121 +3029,72 @@ Lemma normr_sg x : `|sg x| = (x != 0)%:R. Proof. by rewrite sgr_def -mulr_natr normrMsign normr_nat. Qed. Lemma sgr_norm x : sg `|x| = (x != 0)%:R. -Proof. by rewrite /sg ler_gtF ?normr_ge0 // normr_eq0 mulrb if_neg. Qed. - -(* lerif *) - -Lemma lerif_refl x C : reflect (x <= x ?= iff C) C. -Proof. by apply: (iffP idP) => [-> | <-] //; split; rewrite ?eqxx. Qed. +Proof. by rewrite /sg le_gtF // normr_eq0 mulrb if_neg. Qed. -Lemma lerif_trans x1 x2 x3 C12 C23 : - x1 <= x2 ?= iff C12 -> x2 <= x3 ?= iff C23 -> x1 <= x3 ?= iff C12 && C23. -Proof. -move=> ltx12 ltx23; apply/lerifP; rewrite -ltx12. -case eqx12: (x1 == x2). - by rewrite (eqP eqx12) ltr_neqAle !ltx23 andbT; case C23. -by rewrite (@ltr_le_trans _ x2) ?ltx23 // ltr_neqAle eqx12 ltx12. -Qed. - -Lemma lerif_le x y : x <= y -> x <= y ?= iff (x >= y). -Proof. by move=> lexy; split=> //; rewrite eqr_le lexy. Qed. +(* leif *) -Lemma lerif_eq x y : x <= y -> x <= y ?= iff (x == y). -Proof. by []. Qed. +Lemma leif_nat_r m n C : (m%:R <= n%:R ?= iff C :> R) = (m <= n ?= iff C)%N. +Proof. by rewrite /leif !ler_nat eqr_nat. Qed. -Lemma ger_lerif x y C : x <= y ?= iff C -> (y <= x) = C. -Proof. by case=> le_xy; rewrite eqr_le le_xy. Qed. +Lemma leif_subLR x y z C : (x - y <= z ?= iff C) = (x <= z + y ?= iff C). +Proof. by rewrite /leif !eq_le ler_subr_addr ler_subl_addr. Qed. -Lemma ltr_lerif x y C : x <= y ?= iff C -> (x < y) = ~~ C. -Proof. by move=> le_xy; rewrite ltr_neqAle !le_xy andbT. Qed. - -Lemma lerif_nat m n C : (m%:R <= n%:R ?= iff C :> R) = (m <= n ?= iff C)%N. -Proof. by rewrite /lerif !ler_nat eqr_nat. Qed. - -Lemma mono_in_lerif (A : {pred R}) (f : R -> R) C : - {in A &, {mono f : x y / x <= y}} -> - {in A &, forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C)}. -Proof. -by move=> mf x y Ax Ay; rewrite /lerif mf ?(inj_in_eq (incr_inj_in mf)). -Qed. - -Lemma mono_lerif (f : R -> R) C : - {mono f : x y / x <= y} -> - forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C). -Proof. by move=> mf x y; rewrite /lerif mf (inj_eq (incr_inj _)). Qed. - -Lemma nmono_in_lerif (A : {pred R}) (f : R -> R) C : - {in A &, {mono f : x y /~ x <= y}} -> - {in A &, forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C)}. -Proof. -move=> mf x y Ax Ay; rewrite /lerif eq_sym mf //. -by rewrite ?(inj_in_eq (decr_inj_in mf)). -Qed. +Lemma leif_subRL x y z C : (x <= y - z ?= iff C) = (x + z <= y ?= iff C). +Proof. by rewrite -leif_subLR opprK. Qed. -Lemma nmono_lerif (f : R -> R) C : - {mono f : x y /~ x <= y} -> - forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C). -Proof. by move=> mf x y; rewrite /lerif eq_sym mf ?(inj_eq (decr_inj mf)). Qed. - -Lemma lerif_subLR x y z C : (x - y <= z ?= iff C) = (x <= z + y ?= iff C). -Proof. by rewrite /lerif !eqr_le ler_subr_addr ler_subl_addr. Qed. - -Lemma lerif_subRL x y z C : (x <= y - z ?= iff C) = (x + z <= y ?= iff C). -Proof. by rewrite -lerif_subLR opprK. Qed. - -Lemma lerif_add x1 y1 C1 x2 y2 C2 : +Lemma leif_add x1 y1 C1 x2 y2 C2 : x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> x1 + x2 <= y1 + y2 ?= iff C1 && C2. Proof. -rewrite -(mono_lerif _ (ler_add2r x2)) -(mono_lerif C2 (ler_add2l y1)). -exact: lerif_trans. +rewrite -(mono_leif (ler_add2r x2)) -(mono_leif (C := C2) (ler_add2l y1)). +exact: leif_trans. Qed. -Lemma lerif_sum (I : finType) (P C : pred I) (E1 E2 : I -> R) : +Lemma leif_sum (I : finType) (P C : pred I) (E1 E2 : I -> R) : (forall i, P i -> E1 i <= E2 i ?= iff C i) -> \sum_(i | P i) E1 i <= \sum_(i | P i) E2 i ?= iff [forall (i | P i), C i]. Proof. move=> leE12; rewrite -big_andE. -elim/big_rec3: _ => [|i Ci m2 m1 /leE12]; first by rewrite /lerif lerr eqxx. -exact: lerif_add. +elim/big_rec3: _ => [|i Ci m2 m1 /leE12]; first by rewrite /leif lexx eqxx. +exact: leif_add. Qed. -Lemma lerif_0_sum (I : finType) (P C : pred I) (E : I -> R) : +Lemma leif_0_sum (I : finType) (P C : pred I) (E : I -> R) : (forall i, P i -> 0 <= E i ?= iff C i) -> 0 <= \sum_(i | P i) E i ?= iff [forall (i | P i), C i]. -Proof. by move/lerif_sum; rewrite big1_eq. Qed. +Proof. by move/leif_sum; rewrite big1_eq. Qed. -Lemma real_lerif_norm x : x \is real -> x <= `|x| ?= iff (0 <= x). +Lemma real_leif_norm x : x \is real -> x <= `|x| ?= iff (0 <= x). Proof. -by move=> xR; rewrite ger0_def eq_sym; apply: lerif_eq; rewrite real_ler_norm. +by move=> xR; rewrite ger0_def eq_sym; apply: leif_eq; rewrite real_ler_norm. Qed. -Lemma lerif_pmul x1 x2 y1 y2 C1 C2 : +Lemma leif_pmul x1 x2 y1 y2 C1 C2 : 0 <= x1 -> 0 <= x2 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> x1 * x2 <= y1 * y2 ?= iff (y1 * y2 == 0) || C1 && C2. Proof. -move=> x1_ge0 x2_ge0 le_xy1 le_xy2; have [y_0 | ] := altP (_ =P 0). - apply/lerifP; rewrite y_0 /= mulf_eq0 !eqr_le x1_ge0 x2_ge0 !andbT. +move=> x1_ge0 x2_ge0 le_xy1 le_xy2; have [y_0 | ] := eqVneq _ 0. + apply/leifP; rewrite y_0 /= mulf_eq0 !eq_le x1_ge0 x2_ge0 !andbT. move/eqP: y_0; rewrite mulf_eq0. by case/pred2P=> <-; rewrite (le_xy1, le_xy2) ?orbT. rewrite /= mulf_eq0 => /norP[y1nz y2nz]. -have y1_gt0: 0 < y1 by rewrite ltr_def y1nz (ler_trans _ le_xy1). +have y1_gt0: 0 < y1 by rewrite lt_def y1nz (le_trans _ le_xy1). have [x2_0 | x2nz] := eqVneq x2 0. - apply/lerifP; rewrite -le_xy2 x2_0 eq_sym (negPf y2nz) andbF mulr0. - by rewrite mulr_gt0 // ltr_def y2nz -x2_0 le_xy2. -have:= le_xy2; rewrite -(mono_lerif _ (ler_pmul2l y1_gt0)). -by apply: lerif_trans; rewrite (mono_lerif _ (ler_pmul2r _)) // ltr_def x2nz. + apply/leifP; rewrite -le_xy2 x2_0 eq_sym (negPf y2nz) andbF mulr0. + by rewrite mulr_gt0 // lt_def y2nz -x2_0 le_xy2. +have:= le_xy2; rewrite -(mono_leif (ler_pmul2l y1_gt0)). +by apply: leif_trans; rewrite (mono_leif (ler_pmul2r _)) // lt_def x2nz. Qed. -Lemma lerif_nmul x1 x2 y1 y2 C1 C2 : +Lemma leif_nmul x1 x2 y1 y2 C1 C2 : y1 <= 0 -> y2 <= 0 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> y1 * y2 <= x1 * x2 ?= iff (x1 * x2 == 0) || C1 && C2. Proof. rewrite -!oppr_ge0 -mulrNN -[x1 * x2]mulrNN => y1le0 y2le0 le_xy1 le_xy2. -by apply: lerif_pmul => //; rewrite (nmono_lerif _ ler_opp2). +by apply: leif_pmul => //; rewrite (nmono_leif ler_opp2). Qed. -Lemma lerif_pprod (I : finType) (P C : pred I) (E1 E2 : I -> R) : +Lemma leif_pprod (I : finType) (P C : pred I) (E1 E2 : I -> R) : (forall i, P i -> 0 <= E1 i) -> (forall i, P i -> E1 i <= E2 i ?= iff C i) -> let pi E := \prod_(i | P i) E i in @@ -3158,35 +3102,35 @@ Lemma lerif_pprod (I : finType) (P C : pred I) (E1 E2 : I -> R) : Proof. move=> E1_ge0 leE12 /=; rewrite -big_andE; elim/(big_load (fun x => 0 <= x)): _. elim/big_rec3: _ => [|i Ci m2 m1 Pi [m1ge0 le_m12]]. - by split=> //; apply/lerifP; rewrite orbT. + by split=> //; apply/leifP; rewrite orbT. have Ei_ge0 := E1_ge0 i Pi; split; first by rewrite mulr_ge0. -congr (lerif _ _ _): (lerif_pmul Ei_ge0 m1ge0 (leE12 i Pi) le_m12). +congr (leif _ _ _): (leif_pmul Ei_ge0 m1ge0 (leE12 i Pi) le_m12). by rewrite mulf_eq0 -!orbA; congr (_ || _); rewrite !orb_andr orbA orbb. Qed. (* Mean inequalities. *) -Lemma real_lerif_mean_square_scaled x y : +Lemma real_leif_mean_square_scaled x y : x \is real -> y \is real -> x * y *+ 2 <= x ^+ 2 + y ^+ 2 ?= iff (x == y). Proof. -move=> Rx Ry; rewrite -[_ *+ 2]add0r -lerif_subRL addrAC -sqrrB -subr_eq0. -by rewrite -sqrf_eq0 eq_sym; apply: lerif_eq; rewrite -realEsqr rpredB. +move=> Rx Ry; rewrite -[_ *+ 2]add0r -leif_subRL addrAC -sqrrB -subr_eq0. +by rewrite -sqrf_eq0 eq_sym; apply: leif_eq; rewrite -realEsqr rpredB. Qed. -Lemma real_lerif_AGM2_scaled x y : +Lemma real_leif_AGM2_scaled x y : x \is real -> y \is real -> x * y *+ 4 <= (x + y) ^+ 2 ?= iff (x == y). Proof. -move=> Rx Ry; rewrite sqrrD addrAC (mulrnDr _ 2) -lerif_subLR addrK. -exact: real_lerif_mean_square_scaled. +move=> Rx Ry; rewrite sqrrD addrAC (mulrnDr _ 2) -leif_subLR addrK. +exact: real_leif_mean_square_scaled. Qed. -Lemma lerif_AGM_scaled (I : finType) (A : {pred I}) (E : I -> R) (n := #|A|) : +Lemma leif_AGM_scaled (I : finType) (A : {pred I}) (E : I -> R) (n := #|A|) : {in A, forall i, 0 <= E i *+ n} -> \prod_(i in A) (E i *+ n) <= (\sum_(i in A) E i) ^+ n ?= iff [forall i in A, forall j in A, E i == E j]. Proof. have [m leAm] := ubnP #|A|; elim: m => // m IHm in A leAm E n * => Ege0. -apply/lerifP; case: ifPn => [/forall_inP-Econstant | Enonconstant]. +apply/leifP; case: ifPn => [/forall_inP-Econstant | Enonconstant]. have [i /= Ai | A0] := pickP (mem A); last by rewrite [n]eq_card0 ?big_pred0. have /eqfun_inP-E_i := Econstant i Ai; rewrite -(eq_bigr _ E_i) sumr_const. by rewrite exprMn_n prodrMn -(eq_bigr _ E_i) prodr_const. @@ -3196,9 +3140,9 @@ have{Enonconstant} has_cmp_mu e (s := (-1) ^+ e): {i | i \in A & cmp_mu s i}. apply/sig2W/exists_inP; apply: contraR Enonconstant => /exists_inPn-mu_s_A. have n_gt0 i: i \in A -> (0 < n)%N by rewrite [n](cardD1 i) => ->. have{mu_s_A} mu_s_A i: i \in A -> s * En i <= s * mu. - move=> Ai; rewrite real_lerNgt ?mu_s_A ?rpredMsign ?ger0_real ?Ege0 //. + move=> Ai; rewrite real_leNgt ?mu_s_A ?rpredMsign ?ger0_real ?Ege0 //. by rewrite -(pmulrn_lge0 _ (n_gt0 i Ai)) -sumrMnl sumr_ge0. - have [_ /esym/eqfun_inP] := lerif_sum (fun i Ai => lerif_eq (mu_s_A i Ai)). + have [_ /esym/eqfun_inP] := leif_sum (fun i Ai => leif_eq (mu_s_A i Ai)). rewrite sumr_const -/n -mulr_sumr sumrMnl -/mu mulrnAr eqxx => A_mu. apply/forall_inP=> i Ai; apply/eqfun_inP=> j Aj. by apply: (pmulrnI (n_gt0 i Ai)); apply: (can_inj (signrMK e)); rewrite !A_mu. @@ -3206,17 +3150,17 @@ have [[i Ai Ei_lt_mu] [j Aj Ej_gt_mu]] := (has_cmp_mu 1, has_cmp_mu 0)%N. rewrite {cmp_mu has_cmp_mu}/= !mul1r !mulN1r ltr_opp2 in Ei_lt_mu Ej_gt_mu. pose A' := [predD1 A & i]; pose n' := #|A'|. have [Dn n_gt0]: n = n'.+1 /\ (n > 0)%N by rewrite [n](cardD1 i) Ai. -have i'j: j != i by apply: contraTneq Ej_gt_mu => ->; rewrite ltr_gtF. +have i'j: j != i by apply: contraTneq Ej_gt_mu => ->; rewrite lt_gtF. have{i'j} A'j: j \in A' by rewrite !inE Aj i'j. -have mu_gt0: 0 < mu := ler_lt_trans (Ege0 i Ai) Ei_lt_mu. +have mu_gt0: 0 < mu := le_lt_trans (Ege0 i Ai) Ei_lt_mu. rewrite (bigD1 i) // big_andbC (bigD1 j) //= mulrA; set pi := \prod_(k | _) _. have [-> | nz_pi] := eqVneq pi 0; first by rewrite !mulr0 exprn_gt0. have{nz_pi} pi_gt0: 0 < pi. - by rewrite ltr_def nz_pi prodr_ge0 // => k /andP[/andP[_ /Ege0]]. + by rewrite lt_def nz_pi prodr_ge0 // => k /andP[/andP[_ /Ege0]]. rewrite -/(En i) -/(En j); pose E' := [eta En with j |-> En i + En j - mu]. have E'ge0 k: k \in A' -> E' k *+ n' >= 0. case/andP=> /= _ Ak; apply: mulrn_wge0; case: ifP => _; last exact: Ege0. - by rewrite subr_ge0 ler_paddl ?Ege0 // ltrW. + by rewrite subr_ge0 ler_paddl ?Ege0 // ltW. rewrite -/n Dn in leAm; have{leAm IHm E'ge0}: _ <= _ := IHm _ leAm _ E'ge0. have ->: \sum_(k in A') E' k = mu *+ n'. apply: (addrI mu); rewrite -mulrS -Dn -sumrMnl (bigD1 i Ai) big_andbC /=. @@ -3225,7 +3169,7 @@ have ->: \sum_(k in A') E' k = mu *+ n'. rewrite prodrMn exprMn_n -/n' ler_pmuln2r ?expn_gt0; last by case: (n'). have ->: \prod_(k in A') E' k = E' j * pi. by rewrite (bigD1 j) //=; congr *%R; apply: eq_bigr => k /andP[_ /negPf->]. -rewrite -(ler_pmul2l mu_gt0) -exprS -Dn mulrA; apply: ltr_le_trans. +rewrite -(ler_pmul2l mu_gt0) -exprS -Dn mulrA; apply: lt_le_trans. rewrite ltr_pmul2r //= eqxx -addrA mulrDr mulrC -ltr_subl_addl -mulrBl. by rewrite mulrC ltr_pmul2r ?subr_gt0. Qed. @@ -3237,9 +3181,9 @@ Implicit Type p : {poly R}. Lemma poly_disk_bound p b : {ub | forall x, `|x| <= b -> `|p.[x]| <= ub}. Proof. exists (\sum_(j < size p) `|p`_j| * b ^+ j) => x le_x_b. -rewrite horner_coef (ler_trans (ler_norm_sum _ _ _)) ?ler_sum // => j _. -rewrite normrM normrX ler_wpmul2l ?ler_expn2r ?unfold_in ?normr_ge0 //. -exact: ler_trans (normr_ge0 x) le_x_b. +rewrite horner_coef (le_trans (ler_norm_sum _ _ _)) ?ler_sum // => j _. +rewrite normrM normrX ler_wpmul2l ?ler_expn2r ?unfold_in //. +exact: le_trans (normr_ge0 x) le_x_b. Qed. End NumDomainOperationTheory. @@ -3250,13 +3194,9 @@ Arguments ltr_sqr {R} [x y]. Arguments signr_inj {R} [x1 x2]. Arguments real_ler_normlP {R x y}. Arguments real_ltr_normlP {R x y}. -Arguments lerif_refl {R x C}. -Arguments mono_in_lerif [R A f C]. -Arguments nmono_in_lerif [R A f C]. -Arguments mono_lerif [R f C]. -Arguments nmono_lerif [R f C]. Section NumDomainMonotonyTheoryForReals. +Local Open Scope order_scope. Variables (R R' : numDomainType) (D : pred R) (f : R -> R') (f' : R -> nat). Implicit Types (m n p : nat) (x y z : R) (u v w : R'). @@ -3264,17 +3204,17 @@ Implicit Types (m n p : nat) (x y z : R) (u v w : R'). Lemma real_mono : {homo f : x y / x < y} -> {in real &, {mono f : x y / x <= y}}. Proof. -move=> mf x y xR yR /=; have [lt_xy | le_yx] := real_lerP xR yR. - by rewrite ltrW_homo. -by rewrite ltr_geF ?mf. +move=> mf x y xR yR /=; have [lt_xy | le_yx] := real_leP xR yR. + by rewrite ltW_homo. +by rewrite lt_geF ?mf. Qed. Lemma real_nmono : {homo f : x y /~ x < y} -> {in real &, {mono f : x y /~ x <= y}}. Proof. -move=> mf x y xR yR /=; have [lt_xy|le_yx] := real_ltrP xR yR. - by rewrite ltr_geF ?mf. -by rewrite ltrW_nhomo. +move=> mf x y xR yR /=; have [lt_xy|le_yx] := real_ltP xR yR. + by rewrite lt_geF ?mf. +by rewrite ltW_nhomo. Qed. Lemma real_mono_in : @@ -3282,8 +3222,8 @@ Lemma real_mono_in : {in [pred x in D | x \is real] &, {mono f : x y / x <= y}}. Proof. move=> Dmf x y /andP[hx xR] /andP[hy yR] /=. -have [lt_xy|le_yx] := real_lerP xR yR; first by rewrite (ltrW_homo_in Dmf). -by rewrite ltr_geF ?Dmf. +have [lt_xy|le_yx] := real_leP xR yR; first by rewrite (ltW_homo_in Dmf). +by rewrite lt_geF ?Dmf. Qed. Lemma real_nmono_in : @@ -3291,40 +3231,40 @@ Lemma real_nmono_in : {in [pred x in D | x \is real] &, {mono f : x y /~ x <= y}}. Proof. move=> Dmf x y /andP[hx xR] /andP[hy yR] /=. -have [lt_xy|le_yx] := real_ltrP xR yR; last by rewrite (ltrW_nhomo_in Dmf). -by rewrite ltr_geF ?Dmf. +have [lt_xy|le_yx] := real_ltP xR yR; last by rewrite (ltW_nhomo_in Dmf). +by rewrite lt_geF ?Dmf. Qed. -Lemma realn_mono : {homo f' : x y / x < y >-> (x < y)%N} -> - {in real &, {mono f' : x y / x <= y >-> (x <= y)%N}}. +Lemma realn_mono : {homo f' : x y / x < y >-> (x < y)} -> + {in real &, {mono f' : x y / x <= y >-> (x <= y)}}. Proof. -move=> mf x y xR yR /=; have [lt_xy | le_yx] := real_lerP xR yR. - by rewrite ltrnW_homo. -by rewrite ltn_geF ?mf. +move=> mf x y xR yR /=; have [lt_xy | le_yx] := real_leP xR yR. + by rewrite ltW_homo. +by rewrite lt_geF ?mf. Qed. -Lemma realn_nmono : {homo f' : x y / y < x >-> (x < y)%N} -> - {in real &, {mono f' : x y / y <= x >-> (x <= y)%N}}. +Lemma realn_nmono : {homo f' : x y / y < x >-> (x < y)} -> + {in real &, {mono f' : x y / y <= x >-> (x <= y)}}. Proof. -move=> mf x y xR yR /=; have [lt_xy|le_yx] := real_ltrP xR yR. - by rewrite ltn_geF ?mf. -by rewrite ltrnW_nhomo. +move=> mf x y xR yR /=; have [lt_xy|le_yx] := real_ltP xR yR. + by rewrite lt_geF ?mf. +by rewrite ltW_nhomo. Qed. -Lemma realn_mono_in : {in D &, {homo f' : x y / x < y >-> (x < y)%N}} -> - {in [pred x in D | x \is real] &, {mono f' : x y / x <= y >-> (x <= y)%N}}. +Lemma realn_mono_in : {in D &, {homo f' : x y / x < y >-> (x < y)}} -> + {in [pred x in D | x \is real] &, {mono f' : x y / x <= y >-> (x <= y)}}. Proof. move=> Dmf x y /andP[hx xR] /andP[hy yR] /=. -have [lt_xy|le_yx] := real_lerP xR yR; first by rewrite (ltrnW_homo_in Dmf). -by rewrite ltn_geF ?Dmf. +have [lt_xy|le_yx] := real_leP xR yR; first by rewrite (ltW_homo_in Dmf). +by rewrite lt_geF ?Dmf. Qed. -Lemma realn_nmono_in : {in D &, {homo f' : x y / y < x >-> (x < y)%N}} -> - {in [pred x in D | x \is real] &, {mono f' : x y / y <= x >-> (x <= y)%N}}. +Lemma realn_nmono_in : {in D &, {homo f' : x y / y < x >-> (x < y)}} -> + {in [pred x in D | x \is real] &, {mono f' : x y / y <= x >-> (x <= y)}}. Proof. move=> Dmf x y /andP[hx xR] /andP[hy yR] /=. -have [lt_xy|le_yx] := real_ltrP xR yR; last by rewrite (ltrnW_nhomo_in Dmf). -by rewrite ltn_geF ?Dmf. +have [lt_xy|le_yx] := real_ltP xR yR; last by rewrite (ltW_nhomo_in Dmf). +by rewrite lt_geF ?Dmf. Qed. End NumDomainMonotonyTheoryForReals. @@ -3340,13 +3280,13 @@ Lemma natrG_gt0 G : #|G|%:R > 0 :> R. Proof. by rewrite ltr0n cardG_gt0. Qed. Lemma natrG_neq0 G : #|G|%:R != 0 :> R. -Proof. by rewrite gtr_eqF // natrG_gt0. Qed. +Proof. by rewrite gt_eqF // natrG_gt0. Qed. Lemma natr_indexg_gt0 G B : #|G : B|%:R > 0 :> R. Proof. by rewrite ltr0n indexg_gt0. Qed. Lemma natr_indexg_neq0 G B : #|G : B|%:R != 0 :> R. -Proof. by rewrite gtr_eqF // natr_indexg_gt0. Qed. +Proof. by rewrite gt_eqF // natr_indexg_gt0. Qed. End FinGroup. @@ -3356,10 +3296,10 @@ Variable F : numFieldType. Implicit Types x y z t : F. Lemma unitf_gt0 x : 0 < x -> x \is a GRing.unit. -Proof. by move=> hx; rewrite unitfE eq_sym ltr_eqF. Qed. +Proof. by move=> hx; rewrite unitfE eq_sym lt_eqF. Qed. Lemma unitf_lt0 x : x < 0 -> x \is a GRing.unit. -Proof. by move=> hx; rewrite unitfE ltr_eqF. Qed. +Proof. by move=> hx; rewrite unitfE lt_eqF. Qed. Lemma lef_pinv : {in pos &, {mono (@GRing.inv F) : x y /~ x <= y}}. Proof. by move=> x y hx hy /=; rewrite ler_pinv ?inE ?unitf_gt0. Qed. @@ -3368,10 +3308,10 @@ Lemma lef_ninv : {in neg &, {mono (@GRing.inv F) : x y /~ x <= y}}. Proof. by move=> x y hx hy /=; rewrite ler_ninv ?inE ?unitf_lt0. Qed. Lemma ltf_pinv : {in pos &, {mono (@GRing.inv F) : x y /~ x < y}}. -Proof. exact: lerW_nmono_in lef_pinv. Qed. +Proof. exact: leW_nmono_in lef_pinv. Qed. Lemma ltf_ninv: {in neg &, {mono (@GRing.inv F) : x y /~ x < y}}. -Proof. exact: lerW_nmono_in lef_ninv. Qed. +Proof. exact: leW_nmono_in lef_ninv. Qed. Definition ltef_pinv := (lef_pinv, ltf_pinv). Definition ltef_ninv := (lef_ninv, ltf_ninv). @@ -3395,18 +3335,18 @@ Definition invf_cp1 := (invf_gte1, invf_lte1). (* These lemma are all combinations of mono(LR|RL) with ler_[pn]mul2[rl]. *) Lemma ler_pdivl_mulr z x y : 0 < z -> (x <= y / z) = (x * z <= y). -Proof. by move=> z_gt0; rewrite -(@ler_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed. +Proof. by move=> z_gt0; rewrite -(@ler_pmul2r _ z _ x) ?mulfVK ?gt_eqF. Qed. Lemma ltr_pdivl_mulr z x y : 0 < z -> (x < y / z) = (x * z < y). -Proof. by move=> z_gt0; rewrite -(@ltr_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed. +Proof. by move=> z_gt0; rewrite -(@ltr_pmul2r _ z _ x) ?mulfVK ?gt_eqF. Qed. Definition lter_pdivl_mulr := (ler_pdivl_mulr, ltr_pdivl_mulr). Lemma ler_pdivr_mulr z x y : 0 < z -> (y / z <= x) = (y <= x * z). -Proof. by move=> z_gt0; rewrite -(@ler_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed. +Proof. by move=> z_gt0; rewrite -(@ler_pmul2r _ z) ?mulfVK ?gt_eqF. Qed. Lemma ltr_pdivr_mulr z x y : 0 < z -> (y / z < x) = (y < x * z). -Proof. by move=> z_gt0; rewrite -(@ltr_pmul2r _ z) ?mulfVK ?gtr_eqF. Qed. +Proof. by move=> z_gt0; rewrite -(@ltr_pmul2r _ z) ?mulfVK ?gt_eqF. Qed. Definition lter_pdivr_mulr := (ler_pdivr_mulr, ltr_pdivr_mulr). @@ -3427,18 +3367,18 @@ Proof. by move=> z_gt0; rewrite mulrC ltr_pdivr_mulr ?[z * _]mulrC. Qed. Definition lter_pdivr_mull := (ler_pdivr_mull, ltr_pdivr_mull). Lemma ler_ndivl_mulr z x y : z < 0 -> (x <= y / z) = (y <= x * z). -Proof. by move=> z_lt0; rewrite -(@ler_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed. +Proof. by move=> z_lt0; rewrite -(@ler_nmul2r _ z) ?mulfVK ?lt_eqF. Qed. Lemma ltr_ndivl_mulr z x y : z < 0 -> (x < y / z) = (y < x * z). -Proof. by move=> z_lt0; rewrite -(@ltr_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed. +Proof. by move=> z_lt0; rewrite -(@ltr_nmul2r _ z) ?mulfVK ?lt_eqF. Qed. Definition lter_ndivl_mulr := (ler_ndivl_mulr, ltr_ndivl_mulr). Lemma ler_ndivr_mulr z x y : z < 0 -> (y / z <= x) = (x * z <= y). -Proof. by move=> z_lt0; rewrite -(@ler_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed. +Proof. by move=> z_lt0; rewrite -(@ler_nmul2r _ z) ?mulfVK ?lt_eqF. Qed. Lemma ltr_ndivr_mulr z x y : z < 0 -> (y / z < x) = (x * z < y). -Proof. by move=> z_lt0; rewrite -(@ltr_nmul2r _ z) ?mulfVK ?ltr_eqF. Qed. +Proof. by move=> z_lt0; rewrite -(@ltr_nmul2r _ z) ?mulfVK ?lt_eqF. Qed. Definition lter_ndivr_mulr := (ler_ndivr_mulr, ltr_ndivr_mulr). @@ -3461,13 +3401,13 @@ Definition lter_ndivr_mull := (ler_ndivr_mull, ltr_ndivr_mull). Lemma natf_div m d : (d %| m)%N -> (m %/ d)%:R = m%:R / d%:R :> F. Proof. by apply: char0_natf_div; apply: (@char_num F). Qed. -Lemma normfV : {morph (@norm F) : x / x ^-1}. +Lemma normfV : {morph (@norm F F) : x / x ^-1}. Proof. move=> x /=; have [/normrV //|Nux] := boolP (x \is a GRing.unit). by rewrite !invr_out // unitfE normr_eq0 -unitfE. Qed. -Lemma normf_div : {morph (@norm F) : x y / x / y}. +Lemma normf_div : {morph (@norm F F) : x y / x / y}. Proof. by move=> x y /=; rewrite normrM normfV. Qed. Lemma invr_sg x : (sg x)^-1 = sgr x. @@ -3496,32 +3436,32 @@ Definition midf_lte := (midf_le, midf_lt). (* The AGM, unscaled but without the nth root. *) -Lemma real_lerif_mean_square x y : +Lemma real_leif_mean_square x y : x \is real -> y \is real -> x * y <= mid (x ^+ 2) (y ^+ 2) ?= iff (x == y). Proof. -move=> Rx Ry; rewrite -(mono_lerif (ler_pmul2r (ltr_nat F 0 2))). -by rewrite divfK ?pnatr_eq0 // mulr_natr; apply: real_lerif_mean_square_scaled. +move=> Rx Ry; rewrite -(mono_leif (ler_pmul2r (ltr_nat F 0 2))). +by rewrite divfK ?pnatr_eq0 // mulr_natr; apply: real_leif_mean_square_scaled. Qed. -Lemma real_lerif_AGM2 x y : +Lemma real_leif_AGM2 x y : x \is real -> y \is real -> x * y <= mid x y ^+ 2 ?= iff (x == y). Proof. -move=> Rx Ry; rewrite -(mono_lerif (ler_pmul2r (ltr_nat F 0 4))). +move=> Rx Ry; rewrite -(mono_leif (ler_pmul2r (ltr_nat F 0 4))). rewrite mulr_natr (natrX F 2 2) -exprMn divfK ?pnatr_eq0 //. -exact: real_lerif_AGM2_scaled. +exact: real_leif_AGM2_scaled. Qed. -Lemma lerif_AGM (I : finType) (A : {pred I}) (E : I -> F) : +Lemma leif_AGM (I : finType) (A : {pred I}) (E : I -> F) : let n := #|A| in let mu := (\sum_(i in A) E i) / n%:R in {in A, forall i, 0 <= E i} -> \prod_(i in A) E i <= mu ^+ n ?= iff [forall i in A, forall j in A, E i == E j]. Proof. move=> n mu Ege0; have [n0 | n_gt0] := posnP n. - by rewrite n0 -big_andE !(big_pred0 _ _ _ _ (card0_eq n0)); apply/lerifP. + by rewrite n0 -big_andE !(big_pred0 _ _ _ _ (card0_eq n0)); apply/leifP. pose E' i := E i / n%:R. have defE' i: E' i *+ n = E i by rewrite -mulr_natr divfK ?pnatr_eq0 -?lt0n. -have /lerif_AGM_scaled (i): i \in A -> 0 <= E' i *+ n by rewrite defE' => /Ege0. +have /leif_AGM_scaled (i): i \in A -> 0 <= E' i *+ n by rewrite defE' => /Ege0. rewrite -/n -mulr_suml (eq_bigr _ (in1W defE')); congr (_ <= _ ?= iff _). by do 2![apply: eq_forallb_in => ? _]; rewrite -(eqr_pmuln2r n_gt0) !defE'. Qed. @@ -3537,17 +3477,17 @@ have [q Dp]: {q | forall x, x != 0 -> p.[x] = (a - q.[x^-1] / x) * x ^+ n}. rewrite -/n -lead_coefE; congr (_ + _); apply: eq_bigr=> i _. by rewrite exprB ?unitfE // -exprVn mulrA mulrAC exprSr mulrA. have [b ub_q] := poly_disk_bound q 1; exists (b / `|a| + 1) => x px0. -have b_ge0: 0 <= b by rewrite (ler_trans (normr_ge0 q.[1])) ?ub_q ?normr1. -have{b_ge0} ba_ge0: 0 <= b / `|a| by rewrite divr_ge0 ?normr_ge0. -rewrite real_lerNgt ?rpredD ?rpred1 ?ger0_real ?normr_ge0 //. +have b_ge0: 0 <= b by rewrite (le_trans (normr_ge0 q.[1])) ?ub_q ?normr1. +have{b_ge0} ba_ge0: 0 <= b / `|a| by rewrite divr_ge0. +rewrite real_leNgt ?rpredD ?rpred1 ?ger0_real //. apply: contraL px0 => lb_x; rewrite rootE. -have x_ge1: 1 <= `|x| by rewrite (ler_trans _ (ltrW lb_x)) // ler_paddl. -have nz_x: x != 0 by rewrite -normr_gt0 (ltr_le_trans ltr01). +have x_ge1: 1 <= `|x| by rewrite (le_trans _ (ltW lb_x)) // ler_paddl. +have nz_x: x != 0 by rewrite -normr_gt0 (lt_le_trans ltr01). rewrite {}Dp // mulf_neq0 ?expf_neq0 // subr_eq0 eq_sym. -have: (b / `|a|) < `|x| by rewrite (ltr_trans _ lb_x) // ltr_spaddr ?ltr01. +have: (b / `|a|) < `|x| by rewrite (lt_trans _ lb_x) // ltr_spaddr ?ltr01. apply: contraTneq => /(canRL (divfK nz_x))Dax. rewrite ltr_pdivr_mulr ?normr_gt0 ?lead_coef_eq0 // mulrC -normrM -{}Dax. -by rewrite ler_gtF // ub_q // normfV invf_le1 ?normr_gt0. +by rewrite le_gtF // ub_q // normfV invf_le1 ?normr_gt0. Qed. Import GroupScope. @@ -3560,73 +3500,32 @@ End NumFieldTheory. Section RealDomainTheory. -Hint Resolve lerr : core. - Variable R : realDomainType. Implicit Types x y z t : R. Lemma num_real x : x \is real. Proof. exact: num_real. Qed. Hint Resolve num_real : core. -Lemma ler_total : total (@le R). Proof. by move=> x y; apply: real_leVge. Qed. - -Lemma ltr_total x y : x != y -> (x < y) || (y < x). -Proof. by rewrite !ltr_def [_ == y]eq_sym => ->; apply: ler_total. Qed. - -Lemma wlog_ler P : - (forall a b, P b a -> P a b) -> (forall a b, a <= b -> P a b) -> - forall a b : R, P a b. -Proof. by move=> sP hP a b; apply: real_wlog_ler. Qed. - -Lemma wlog_ltr P : - (forall a, P a a) -> - (forall a b, (P b a -> P a b)) -> (forall a b, a < b -> P a b) -> - forall a b : R, P a b. -Proof. by move=> rP sP hP a b; apply: real_wlog_ltr. Qed. - -Lemma ltrNge x y : (x < y) = ~~ (y <= x). Proof. exact: real_ltrNge. Qed. - -Lemma lerNgt x y : (x <= y) = ~~ (y < x). Proof. exact: real_lerNgt. Qed. - Lemma lerP x y : ler_xor_gt x y `|x - y| `|y - x| (x <= y) (y < x). -Proof. exact: real_lerP. Qed. +Proof. exact: real_leP. Qed. Lemma ltrP x y : ltr_xor_ge x y `|x - y| `|y - x| (y <= x) (x < y). -Proof. exact: real_ltrP. Qed. +Proof. exact: real_ltP. Qed. Lemma ltrgtP x y : comparer x y `|x - y| `|y - x| (y == x) (x == y) - (x <= y) (y <= x) (x < y) (x > y) . -Proof. exact: real_ltrgtP. Qed. + (x >= y) (x <= y) (x > y) (x < y) . +Proof. exact: real_ltgtP. Qed. Lemma ger0P x : ger0_xor_lt0 x `|x| (x < 0) (0 <= x). -Proof. exact: real_ger0P. Qed. +Proof. exact: real_ge0P. Qed. Lemma ler0P x : ler0_xor_gt0 x `|x| (0 < x) (x <= 0). -Proof. exact: real_ler0P. Qed. +Proof. exact: real_le0P. Qed. Lemma ltrgt0P x : comparer0 x `|x| (0 == x) (x == 0) (x <= 0) (0 <= x) (x < 0) (x > 0). -Proof. exact: real_ltrgt0P. Qed. - -Lemma neqr_lt x y : (x != y) = (x < y) || (y < x). -Proof. exact: real_neqr_lt. Qed. - -Lemma eqr_leLR x y z t : - (x <= y -> z <= t) -> (y < x -> t < z) -> (x <= y) = (z <= t). -Proof. by move=> *; apply/idP/idP; rewrite // !lerNgt; apply: contra. Qed. - -Lemma eqr_leRL x y z t : - (x <= y -> z <= t) -> (y < x -> t < z) -> (z <= t) = (x <= y). -Proof. by move=> *; symmetry; apply: eqr_leLR. Qed. - -Lemma eqr_ltLR x y z t : - (x < y -> z < t) -> (y <= x -> t <= z) -> (x < y) = (z < t). -Proof. by move=> *; rewrite !ltrNge; congr negb; apply: eqr_leLR. Qed. - -Lemma eqr_ltRL x y z t : - (x < y -> z < t) -> (y <= x -> t <= z) -> (z < t) = (x < y). -Proof. by move=> *; symmetry; apply: eqr_ltLR. Qed. +Proof. exact: real_ltgt0P. Qed. (* sign *) @@ -3646,7 +3545,7 @@ Lemma mulr_sign_lt0 (b : bool) x : ((-1) ^+ b * x < 0) = (x != 0) && (b (+) (x < 0)%R). Proof. by rewrite mulr_lt0 signr_lt0 signr_eq0. Qed. -(* sign & norm*) +(* sign & norm *) Lemma mulr_sign_norm x : (-1) ^+ (x < 0)%R * `|x| = x. Proof. by rewrite real_mulr_sign_norm. Qed. @@ -3667,103 +3566,35 @@ End RealDomainTheory. Hint Resolve num_real : core. -Section RealDomainMonotony. - -Variables (R : realDomainType) (R' : numDomainType) (D : pred R). -Variables (f : R -> R') (f' : R -> nat). -Implicit Types (m n p : nat) (x y z : R) (u v w : R'). - -Hint Resolve (@num_real R) : core. - -Lemma ler_mono : {homo f : x y / x < y} -> {mono f : x y / x <= y}. -Proof. by move=> mf x y; apply: real_mono. Qed. - -Lemma ler_nmono : {homo f : x y /~ x < y} -> {mono f : x y /~ x <= y}. -Proof. by move=> mf x y; apply: real_nmono. Qed. - -Lemma ler_mono_in : - {in D &, {homo f : x y / x < y}} -> {in D &, {mono f : x y / x <= y}}. -Proof. -by move=> mf x y Dx Dy; apply: (real_mono_in mf); rewrite ?inE ?Dx ?Dy /=. -Qed. - -Lemma ler_nmono_in : - {in D &, {homo f : x y /~ x < y}} -> {in D &, {mono f : x y /~ x <= y}}. -Proof. -by move=> mf x y Dx Dy; apply: (real_nmono_in mf); rewrite ?inE ?Dx ?Dy /=. -Qed. - -Lemma lern_mono : {homo f' : m n / m < n >-> (m < n)%N} -> - {mono f' : m n / m <= n >-> (m <= n)%N}. -Proof. by move=> mf x y; apply: realn_mono. Qed. - -Lemma lern_nmono : {homo f' : m n / n < m >-> (m < n)%N} -> - {mono f' : m n / n <= m >-> (m <= n)%N}. -Proof. by move=> mf x y; apply: realn_nmono. Qed. - -Lemma lern_mono_in : {in D &, {homo f' : m n / m < n >-> (m < n)%N}} -> - {in D &, {mono f' : m n / m <= n >-> (m <= n)%N}}. -Proof. -by move=> mf x y Dx Dy; apply: (realn_mono_in mf); rewrite ?inE ?Dx ?Dy /=. -Qed. - -Lemma lern_nmono_in : {in D &, {homo f' : m n / n < m >-> (m < n)%N}} -> - {in D &, {mono f' : m n / n <= m >-> (m <= n)%N}}. -Proof. -by move=> mf x y Dx Dy; apply: (realn_nmono_in mf); rewrite ?inE ?Dx ?Dy /=. -Qed. - -End RealDomainMonotony. - -Section RealDomainArgExtremum. - -Context {R : realDomainType} {I : finType} (i0 : I). -Context (P : pred I) (F : I -> R) (Pi0 : P i0). - -Definition arg_minr := extremum <=%R i0 P F. -Definition arg_maxr := extremum >=%R i0 P F. - -Lemma arg_minrP: extremum_spec <=%R P F arg_minr. -Proof. by apply: extremumP => //; [apply: ler_trans|apply: ler_total]. Qed. - -Lemma arg_maxrP: extremum_spec >=%R P F arg_maxr. -Proof. -apply: extremumP => //; first exact: lerr. - by move=> ??? /(ler_trans _) le /le. -by move=> ??; apply: ler_total. -Qed. - -End RealDomainArgExtremum. +Section RealDomainOperations. -Notation "[ 'arg' 'minr_' ( i < i0 | P ) F ]" := - (arg_minr i0 (fun i => P%B) (fun i => F)) +Notation "[ 'arg' 'min_' ( i < i0 | P ) F ]" := + (Order.arg_min (disp := ring_display) i0 (fun i => P%B) (fun i => F)) (at level 0, i, i0 at level 10, - format "[ 'arg' 'minr_' ( i < i0 | P ) F ]") : form_scope. + format "[ 'arg' 'min_' ( i < i0 | P ) F ]") : ring_scope. -Notation "[ 'arg' 'minr_' ( i < i0 'in' A ) F ]" := - [arg minr_(i < i0 | i \in A) F] +Notation "[ 'arg' 'min_' ( i < i0 'in' A ) F ]" := + [arg min_(i < i0 | i \in A) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'minr_' ( i < i0 'in' A ) F ]") : form_scope. + format "[ 'arg' 'min_' ( i < i0 'in' A ) F ]") : ring_scope. -Notation "[ 'arg' 'minr_' ( i < i0 ) F ]" := [arg minr_(i < i0 | true) F] +Notation "[ 'arg' 'min_' ( i < i0 ) F ]" := [arg min_(i < i0 | true) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'minr_' ( i < i0 ) F ]") : form_scope. + format "[ 'arg' 'min_' ( i < i0 ) F ]") : ring_scope. -Notation "[ 'arg' 'maxr_' ( i > i0 | P ) F ]" := - (arg_maxr i0 (fun i => P%B) (fun i => F)) +Notation "[ 'arg' 'max_' ( i > i0 | P ) F ]" := + (Order.arg_max (disp := ring_display) i0 (fun i => P%B) (fun i => F)) (at level 0, i, i0 at level 10, - format "[ 'arg' 'maxr_' ( i > i0 | P ) F ]") : form_scope. + format "[ 'arg' 'max_' ( i > i0 | P ) F ]") : ring_scope. -Notation "[ 'arg' 'maxr_' ( i > i0 'in' A ) F ]" := - [arg maxr_(i > i0 | i \in A) F] +Notation "[ 'arg' 'max_' ( i > i0 'in' A ) F ]" := + [arg max_(i > i0 | i \in A) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'maxr_' ( i > i0 'in' A ) F ]") : form_scope. + format "[ 'arg' 'max_' ( i > i0 'in' A ) F ]") : ring_scope. -Notation "[ 'arg' 'maxr_' ( i > i0 ) F ]" := [arg maxr_(i > i0 | true) F] +Notation "[ 'arg' 'max_' ( i > i0 ) F ]" := [arg max_(i > i0 | true) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'maxr_' ( i > i0 ) F ]") : form_scope. - -Section RealDomainOperations. + format "[ 'arg' 'max_' ( i > i0 ) F ]") : ring_scope. (* sgr section *) @@ -3776,8 +3607,8 @@ Lemma sgr_cp0 x : ((sg x == -1) = (x < 0)) * ((sg x == 0) = (x == 0)). Proof. -rewrite -[1]/((-1) ^+ false) -signrN lt0r lerNgt sgr_def. -case: (x =P 0) => [-> | _]; first by rewrite !(eq_sym 0) !signr_eq0 ltrr eqxx. +rewrite -[1]/((-1) ^+ false) -signrN lt0r leNgt sgr_def. +case: (x =P 0) => [-> | _]; first by rewrite !(eq_sym 0) !signr_eq0 ltxx eqxx. by rewrite !(inj_eq signr_inj) eqb_id eqbF_neg signr_eq0 //. Qed. @@ -3826,7 +3657,7 @@ Lemma sgr_gt0 x : (sg x > 0) = (x > 0). Proof. by rewrite -sgr_cp0 sgr_id sgr_cp0. Qed. Lemma sgr_ge0 x : (sgr x >= 0) = (x >= 0). -Proof. by rewrite !lerNgt sgr_lt0. Qed. +Proof. by rewrite !leNgt sgr_lt0. Qed. (* norm section *) @@ -3856,10 +3687,10 @@ Proof. exact: real_ltr_normlP. Qed. Arguments ltr_normlP {x y}. Lemma ler_normr x y : (x <= `|y|) = (x <= y) || (x <= - y). -Proof. by rewrite lerNgt ltr_norml negb_and -!lerNgt orbC ler_oppr. Qed. +Proof. by rewrite leNgt ltr_norml negb_and -!leNgt orbC ler_oppr. Qed. Lemma ltr_normr x y : (x < `|y|) = (x < y) || (x < - y). -Proof. by rewrite ltrNge ler_norml negb_and -!ltrNge orbC ltr_oppr. Qed. +Proof. by rewrite ltNge ler_norml negb_and -!ltNge orbC ltr_oppr. Qed. Definition lter_normr := (ler_normr, ltr_normr). @@ -3902,45 +3733,23 @@ Lemma sqr_ge0 x : 0 <= x ^+ 2. Proof. by rewrite exprn_even_ge0. Qed. Lemma sqr_norm_eq1 x : (x ^+ 2 == 1) = (`|x| == 1). Proof. by rewrite sqrf_eq1 eqr_norml ler01 andbT. Qed. -Lemma lerif_mean_square_scaled x y : +Lemma leif_mean_square_scaled x y : x * y *+ 2 <= x ^+ 2 + y ^+ 2 ?= iff (x == y). -Proof. exact: real_lerif_mean_square_scaled. Qed. +Proof. exact: real_leif_mean_square_scaled. Qed. -Lemma lerif_AGM2_scaled x y : x * y *+ 4 <= (x + y) ^+ 2 ?= iff (x == y). -Proof. exact: real_lerif_AGM2_scaled. Qed. +Lemma leif_AGM2_scaled x y : x * y *+ 4 <= (x + y) ^+ 2 ?= iff (x == y). +Proof. exact: real_leif_AGM2_scaled. Qed. Section MinMax. (* GG: Many of the first lemmas hold unconditionally, and others hold for *) (* the real subset of a general domain. *) -Lemma minrC : @commutative R R min. -Proof. by move=> x y; rewrite /min; case: ltrgtP. Qed. - -Lemma minrr : @idempotent R min. -Proof. by move=> x; rewrite /min if_same. Qed. - -Lemma minr_l x y : x <= y -> min x y = x. -Proof. by rewrite /minr => ->. Qed. - -Lemma minr_r x y : y <= x -> min x y = y. -Proof. by move/minr_l; rewrite minrC. Qed. - -Lemma maxrC : @commutative R R max. -Proof. by move=> x y; rewrite /maxr; case: ltrgtP. Qed. - -Lemma maxrr : @idempotent R max. -Proof. by move=> x; rewrite /max if_same. Qed. - -Lemma maxr_l x y : y <= x -> max x y = x. -Proof. by move=> hxy; rewrite /max hxy. Qed. - -Lemma maxr_r x y : x <= y -> max x y = y. -Proof. by move=> hxy; rewrite maxrC maxr_l. Qed. Lemma addr_min_max x y : min x y + max x y = x + y. Proof. -case: (lerP x y)=> hxy; first by rewrite maxr_r ?minr_l. -by rewrite maxr_l ?minr_r ?ltrW // addrC. +case: (lerP x y)=> [| /ltW] hxy; + first by rewrite (meet_idPl hxy) (join_idPl hxy). +by rewrite (meet_idPr hxy) (join_idPr hxy) addrC. Qed. Lemma addr_max_min x y : max x y + min x y = x + y. @@ -3952,164 +3761,35 @@ Proof. by rewrite -[x + y]addr_min_max addrK. Qed. Lemma maxr_to_min x y : max x y = x + y - min x y. Proof. by rewrite -[x + y]addr_max_min addrK. Qed. -Lemma minrA x y z : min x (min y z) = min (min x y) z. -Proof. -rewrite /min; case: (lerP y z) => [hyz | /ltrW hyz]. - by case: lerP => hxy; rewrite ?hyz // (@ler_trans _ y). -case: lerP=> hxz; first by rewrite !(ler_trans hxz). -case: (lerP x y)=> hxy; first by rewrite lerNgt hxz. -by case: ltrgtP hyz. -Qed. - -Lemma minrCA : @left_commutative R R min. -Proof. by move=> x y z; rewrite !minrA [minr x y]minrC. Qed. - -Lemma minrAC : @right_commutative R R min. -Proof. by move=> x y z; rewrite -!minrA [minr y z]minrC. Qed. - -Variant minr_spec x y : bool -> bool -> R -> Type := -| Minr_r of x <= y : minr_spec x y true false x -| Minr_l of y < x : minr_spec x y false true y. - -Lemma minrP x y : minr_spec x y (x <= y) (y < x) (min x y). -Proof. -case: lerP=> hxy; first by rewrite minr_l //; constructor. -by rewrite minr_r 1?ltrW //; constructor. -Qed. - -Lemma oppr_max x y : - max x y = min (- x) (- y). -Proof. -case: minrP; rewrite lter_opp2 => hxy; first by rewrite maxr_l. -by rewrite maxr_r // ltrW. -Qed. - -Lemma oppr_min x y : - min x y = max (- x) (- y). -Proof. by rewrite -[maxr _ _]opprK oppr_max !opprK. Qed. - -Lemma maxrA x y z : max x (max y z) = max (max x y) z. -Proof. by apply/eqP; rewrite -eqr_opp !oppr_max minrA. Qed. - -Lemma maxrCA : @left_commutative R R max. -Proof. by move=> x y z; rewrite !maxrA [maxr x y]maxrC. Qed. - -Lemma maxrAC : @right_commutative R R max. -Proof. by move=> x y z; rewrite -!maxrA [maxr y z]maxrC. Qed. - -Variant maxr_spec x y : bool -> bool -> R -> Type := -| Maxr_r of y <= x : maxr_spec x y true false x -| Maxr_l of x < y : maxr_spec x y false true y. - -Lemma maxrP x y : maxr_spec x y (y <= x) (x < y) (maxr x y). -Proof. -case: lerP => hxy; first by rewrite maxr_l //; constructor. -by rewrite maxr_r 1?ltrW //; constructor. -Qed. - -Lemma eqr_minl x y : (min x y == x) = (x <= y). -Proof. by case: minrP=> hxy; rewrite ?eqxx // ltr_eqF. Qed. - -Lemma eqr_minr x y : (min x y == y) = (y <= x). -Proof. by rewrite minrC eqr_minl. Qed. - -Lemma eqr_maxl x y : (max x y == x) = (y <= x). -Proof. by case: maxrP=> hxy; rewrite ?eqxx // eq_sym ltr_eqF. Qed. - -Lemma eqr_maxr x y : (max x y == y) = (x <= y). -Proof. by rewrite maxrC eqr_maxl. Qed. - -Lemma ler_minr x y z : (x <= min y z) = (x <= y) && (x <= z). -Proof. -case: minrP=> hyz. - by case: lerP=> hxy //; rewrite (ler_trans _ hyz). -by case: lerP=> hxz; rewrite andbC // (ler_trans hxz) // ltrW. -Qed. - -Lemma ler_minl x y z : (min y z <= x) = (y <= x) || (z <= x). +Lemma oppr_max : {morph -%R : x y / max x y >-> min x y : R}. Proof. -case: minrP => hyz. - case: lerP => hyx //=; symmetry; apply: negbTE. - by rewrite -ltrNge (@ltr_le_trans _ y). -case: lerP => hzx; rewrite orbC //=; symmetry; apply: negbTE. -by rewrite -ltrNge (@ltr_trans _ z). +by move=> x y; case: leP; rewrite -lter_opp2 => hxy; + rewrite ?(meet_idPr hxy) ?(meet_idPl (ltW hxy)). Qed. -Lemma ler_maxr x y z : (x <= max y z) = (x <= y) || (x <= z). -Proof. by rewrite -lter_opp2 oppr_max ler_minl !ler_opp2. Qed. - -Lemma ler_maxl x y z : (max y z <= x) = (y <= x) && (z <= x). -Proof. by rewrite -lter_opp2 oppr_max ler_minr !ler_opp2. Qed. - -Lemma ltr_minr x y z : (x < min y z) = (x < y) && (x < z). -Proof. by rewrite !ltrNge ler_minl negb_or. Qed. - -Lemma ltr_minl x y z : (min y z < x) = (y < x) || (z < x). -Proof. by rewrite !ltrNge ler_minr negb_and. Qed. - -Lemma ltr_maxr x y z : (x < max y z) = (x < y) || (x < z). -Proof. by rewrite !ltrNge ler_maxl negb_and. Qed. - -Lemma ltr_maxl x y z : (max y z < x) = (y < x) && (z < x). -Proof. by rewrite !ltrNge ler_maxr negb_or. Qed. - -Definition lter_minr := (ler_minr, ltr_minr). -Definition lter_minl := (ler_minl, ltr_minl). -Definition lter_maxr := (ler_maxr, ltr_maxr). -Definition lter_maxl := (ler_maxl, ltr_maxl). +Lemma oppr_min : {morph -%R : x y / min x y >-> max x y : R}. +Proof. by move=> x y; rewrite -[max _ _]opprK oppr_max !opprK. Qed. Lemma addr_minl : @left_distributive R R +%R min. Proof. -move=> x y z; case: minrP=> hxy; first by rewrite minr_l // ler_add2r. -by rewrite minr_r // ltrW // ltr_add2r. +by move=> x y z; case: leP; case: leP => //; rewrite lter_add2; case: leP. Qed. Lemma addr_minr : @right_distributive R R +%R min. -Proof. -move=> x y z; case: minrP=> hxy; first by rewrite minr_l // ler_add2l. -by rewrite minr_r // ltrW // ltr_add2l. -Qed. +Proof. by move=> x y z; rewrite !(addrC x) addr_minl. Qed. Lemma addr_maxl : @left_distributive R R +%R max. Proof. -move=> x y z; rewrite -[_ + _]opprK opprD oppr_max. -by rewrite addr_minl -!opprD oppr_min !opprK. +by move=> x y z; case: leP; case: leP => //; rewrite lter_add2; case: leP. Qed. Lemma addr_maxr : @right_distributive R R +%R max. -Proof. -move=> x y z; rewrite -[_ + _]opprK opprD oppr_max. -by rewrite addr_minr -!opprD oppr_min !opprK. -Qed. - -Lemma minrK x y : max (min x y) x = x. -Proof. by case: minrP => hxy; rewrite ?maxrr ?maxr_r // ltrW. Qed. - -Lemma minKr x y : min y (max x y) = y. -Proof. by case: maxrP => hxy; rewrite ?minrr ?minr_l. Qed. - -Lemma maxr_minl : @left_distributive R R max min. -Proof. -move=> x y z; case: minrP => hxy. - by case: maxrP => hm; rewrite minr_l // ler_maxr (hxy, lerr) ?orbT. -by case: maxrP => hyz; rewrite minr_r // ler_maxr (ltrW hxy, lerr) ?orbT. -Qed. - -Lemma maxr_minr : @right_distributive R R max min. -Proof. by move=> x y z; rewrite maxrC maxr_minl ![_ _ x]maxrC. Qed. - -Lemma minr_maxl : @left_distributive R R min max. -Proof. -move=> x y z; rewrite -[min _ _]opprK !oppr_min [- max x y]oppr_max. -by rewrite maxr_minl !(oppr_max, oppr_min, opprK). -Qed. - -Lemma minr_maxr : @right_distributive R R min max. -Proof. by move=> x y z; rewrite minrC minr_maxl ![_ _ x]minrC. Qed. +Proof. by move=> x y z; rewrite !(addrC x) addr_maxl. Qed. Lemma minr_pmulr x y z : 0 <= x -> x * min y z = min (x * y) (x * z). Proof. -case: sgrP=> // hx _; first by rewrite hx !mul0r minrr. -case: minrP=> hyz; first by rewrite minr_l // ler_pmul2l. -by rewrite minr_r // ltrW // ltr_pmul2l. +case: sgrP=> // hx _; first by rewrite hx !mul0r meetxx. +by case: leP; case: leP => //; rewrite lter_pmul2l //; case: leP. Qed. Lemma minr_nmulr x y z : x <= 0 -> x * min y z = max (x * y) (x * z). @@ -4143,19 +3823,16 @@ Lemma maxr_nmull x y z : x <= 0 -> max y z * x = min (y * x) (z * x). Proof. by move=> *; rewrite mulrC maxr_nmulr // ![_ * x]mulrC. Qed. Lemma maxrN x : max x (- x) = `|x|. -Proof. -case: ger0P=> hx; first by rewrite maxr_l // ge0_cp //. -by rewrite maxr_r // le0_cp // ltrW. -Qed. +Proof. by case: ger0P=> [/ge0_cp [] | /lt0_cp []]; case: (leP (- x) x). Qed. Lemma maxNr x : max (- x) x = `|x|. -Proof. by rewrite maxrC maxrN. Qed. +Proof. by rewrite joinC maxrN. Qed. Lemma minrN x : min x (- x) = - `|x|. -Proof. by rewrite -[minr _ _]opprK oppr_min opprK maxNr. Qed. +Proof. by rewrite -[min _ _]opprK oppr_min opprK maxNr. Qed. Lemma minNr x : min (- x) x = - `|x|. -Proof. by rewrite -[minr _ _]opprK oppr_min opprK maxrN. Qed. +Proof. by rewrite -[min _ _]opprK oppr_min opprK maxrN. Qed. End MinMax. @@ -4166,7 +3843,7 @@ Variable p : {poly R}. Lemma poly_itv_bound a b : {ub | forall x, a <= x <= b -> `|p.[x]| <= ub}. Proof. have [ub le_p_ub] := poly_disk_bound p (Num.max `|a| `|b|). -exists ub => x /andP[le_a_x le_x_b]; rewrite le_p_ub // ler_maxr !ler_normr. +exists ub => x /andP[le_a_x le_x_b]; rewrite le_p_ub // lexU !ler_normr. by have [_|_] := ler0P x; rewrite ?ler_opp2 ?le_a_x ?le_x_b orbT. Qed. @@ -4177,16 +3854,16 @@ have [p_le1 | p_gt1] := leqP (size p) 1. exists 0 => x _; rewrite (size1_polyC p_le1) hornerC. by rewrite -[p`_0]lead_coefC -size1_polyC // mon_p ltr01. pose lb := \sum_(j < n.+1) `|p`_j|; exists (lb + 1) => x le_ub_x. -have x_ge1: 1 <= x; last have x_gt0 := ltr_le_trans ltr01 x_ge1. - by rewrite -(ler_add2l lb) ler_paddl ?sumr_ge0 // => j _; apply: normr_ge0. +have x_ge1: 1 <= x; last have x_gt0 := lt_le_trans ltr01 x_ge1. + by rewrite -(ler_add2l lb) ler_paddl ?sumr_ge0 // => j _. rewrite horner_coef -(subnK p_gt1) -/n addnS big_ord_recr /= addn1. rewrite [in p`__]subnSK // subn1 -lead_coefE mon_p mul1r -ltr_subl_addl sub0r. -apply: ler_lt_trans (_ : lb * x ^+ n < _); last first. +apply: le_lt_trans (_ : lb * x ^+ n < _); last first. rewrite exprS ltr_pmul2r ?exprn_gt0 ?(ltr_le_trans ltr01) //. by rewrite -(ltr_add2r 1) ltr_spaddr ?ltr01. -rewrite -sumrN mulr_suml ler_sum // => j _; apply: ler_trans (ler_norm _) _. -rewrite normrN normrM ler_wpmul2l ?normr_ge0 // normrX. -by rewrite ger0_norm ?(ltrW x_gt0) // ler_weexpn2l ?leq_ord. +rewrite -sumrN mulr_suml ler_sum // => j _; apply: le_trans (ler_norm _) _. +rewrite normrN normrM ler_wpmul2l // normrX. +by rewrite ger0_norm ?(ltW x_gt0) // ler_weexpn2l ?leq_ord. Qed. End PolyBounds. @@ -4197,11 +3874,11 @@ Section RealField. Variables (F : realFieldType) (x y : F). -Lemma lerif_mean_square : x * y <= (x ^+ 2 + y ^+ 2) / 2%:R ?= iff (x == y). -Proof. by apply: real_lerif_mean_square; apply: num_real. Qed. +Lemma leif_mean_square : x * y <= (x ^+ 2 + y ^+ 2) / 2%:R ?= iff (x == y). +Proof. by apply: real_leif_mean_square; apply: num_real. Qed. -Lemma lerif_AGM2 : x * y <= ((x + y) / 2%:R)^+ 2 ?= iff (x == y). -Proof. by apply: real_lerif_AGM2; apply: num_real. Qed. +Lemma leif_AGM2 : x * y <= ((x + y) / 2%:R)^+ 2 ?= iff (x == y). +Proof. by apply: real_leif_AGM2; apply: num_real. Qed. End RealField. @@ -4215,7 +3892,7 @@ Proof. by move/ger0_norm=> {1}<-; rewrite /bound; case: (sigW _). Qed. Lemma upper_nthrootP i : (bound x <= i)%N -> x < 2%:R ^+ i. Proof. rewrite /bound; case: (sigW _) => /= b le_x_b le_b_i. -apply: ler_lt_trans (ler_norm x) (ltr_trans le_x_b _ ). +apply: le_lt_trans (ler_norm x) (lt_trans le_x_b _ ). by rewrite -natrX ltr_nat (leq_ltn_trans le_b_i) // ltn_expl. Qed. @@ -4246,7 +3923,7 @@ by have [//|_ /eqP//|->] := ltrgt0P a; rewrite mulf_eq0 orbb => /eqP. Qed. Lemma ltr0_sqrtr a : a < 0 -> sqrt a = 0. -Proof. by move=> /ltrW; apply: ler0_sqrtr. Qed. +Proof. by move=> /ltW; apply: ler0_sqrtr. Qed. Variant sqrtr_spec a : R -> bool -> bool -> R -> Type := | IsNoSqrtr of a < 0 : sqrtr_spec a a false true 0 @@ -4283,13 +3960,13 @@ Proof. by move: (sqrtr_sqr 1); rewrite expr1n => ->; rewrite normr1. Qed. Lemma sqrtr_eq0 a : (sqrt a == 0) = (a <= 0). Proof. -case: sqrtrP => [/ltrW ->|b]; first by rewrite eqxx. -case: ltrgt0P => [b_gt0|//|->]; last by rewrite exprS mul0r lerr. -by rewrite ltr_geF ?pmulr_rgt0. +case: sqrtrP => [/ltW ->|b]; first by rewrite eqxx. +case: ltrgt0P => [b_gt0|//|->]; last by rewrite exprS mul0r lexx. +by rewrite lt_geF ?pmulr_rgt0. Qed. Lemma sqrtr_gt0 a : (0 < sqrt a) = (0 < a). -Proof. by rewrite lt0r sqrtr_ge0 sqrtr_eq0 -ltrNge andbT. Qed. +Proof. by rewrite lt0r sqrtr_ge0 sqrtr_eq0 -ltNge andbT. Qed. Lemma eqr_sqrt a b : 0 <= a -> 0 <= b -> (sqrt a == sqrt b) = (a == b). Proof. @@ -4300,29 +3977,29 @@ Qed. Lemma ler_wsqrtr : {homo @sqrt R : a b / a <= b}. Proof. move=> a b /= le_ab; case: (boolP (0 <= a))=> [pa|]; last first. - by rewrite -ltrNge; move/ltrW; rewrite -sqrtr_eq0; move/eqP->. + by rewrite -ltNge; move/ltW; rewrite -sqrtr_eq0; move/eqP->. rewrite -(@ler_pexpn2r R 2) ?nnegrE ?sqrtr_ge0 //. -by rewrite !sqr_sqrtr // (ler_trans pa). +by rewrite !sqr_sqrtr // (le_trans pa). Qed. Lemma ler_psqrt : {in @pos R &, {mono sqrt : a b / a <= b}}. Proof. -apply: ler_mono_in => x y x_gt0 y_gt0. -rewrite !ltr_neqAle => /andP[neq_xy le_xy]. -by rewrite ler_wsqrtr // eqr_sqrt ?ltrW // neq_xy. +apply: le_mono_in => x y x_gt0 y_gt0. +rewrite !lt_neqAle => /andP[neq_xy le_xy]. +by rewrite ler_wsqrtr // eqr_sqrt ?ltW // neq_xy. Qed. Lemma ler_sqrt a b : 0 < b -> (sqrt a <= sqrt b) = (a <= b). Proof. move=> b_gt0; have [a_le0|a_gt0] := ler0P a; last by rewrite ler_psqrt. -by rewrite ler0_sqrtr // sqrtr_ge0 (ler_trans a_le0) ?ltrW. +by rewrite ler0_sqrtr // sqrtr_ge0 (le_trans a_le0) ?ltW. Qed. Lemma ltr_sqrt a b : 0 < b -> (sqrt a < sqrt b) = (a < b). Proof. move=> b_gt0; have [a_le0|a_gt0] := ler0P a; last first. - by rewrite (lerW_mono_in ler_psqrt). -by rewrite ler0_sqrtr // sqrtr_gt0 b_gt0 (ler_lt_trans a_le0). + by rewrite (leW_mono_in ler_psqrt). +by rewrite ler0_sqrtr // sqrtr_gt0 b_gt0 (le_lt_trans a_le0). Qed. End RealClosedFieldTheory. @@ -4341,10 +4018,10 @@ Variable C : numClosedFieldType. Implicit Types a x y z : C. Definition normCK x : `|x| ^+ 2 = x * x^*. -Proof. by case: C x => ? [? ? []]. Qed. +Proof. by case: C x => ? [? ? ? []]. Qed. Lemma sqrCi : 'i ^+ 2 = -1 :> C. -Proof. by case: C => ? [? ? []]. Qed. +Proof. by case: C => ? [? ? ? []]. Qed. Lemma conjCK : involutive (@conjC C). Proof. @@ -4367,15 +4044,14 @@ Variant rootC_spec n (x : C) : Type := Fact rootC_subproof n x : rootC_spec n x. Proof. -have realRe2 u : Re2 u \is Num.real. - rewrite realEsqr expr2 {2}/Re2 -{2}[u]conjCK addrC -rmorphD -normCK. - by rewrite exprn_ge0 ?normr_ge0. +have realRe2 u : Re2 u \is Num.real by + rewrite realEsqr expr2 {2}/Re2 -{2}[u]conjCK addrC -rmorphD -normCK exprn_ge0. have argCle_total : total argCle. move=> u v; rewrite /total /argCle. by do 2!case: (nnegIm _) => //; rewrite ?orbT //= real_leVge. have argCle_trans : transitive argCle. move=> u v w /implyP geZuv /implyP geZvw; apply/implyP. - by case/geZvw/andP=> /geZuv/andP[-> geRuv] /ler_trans->. + by case/geZvw/andP=> /geZuv/andP[-> geRuv] /le_trans->. pose p := 'X^n - (x *+ (n > 0))%:P; have [r0 Dp] := closed_field_poly_normal p. have sz_p: size p = n.+1. rewrite size_addl ?size_polyXn // ltnS size_opp size_polyC mulrn_eq0. @@ -4392,7 +4068,7 @@ exists r`_0 => [|z n_gt0 /(mem_rP z n_gt0) r_z]. case: posnP => [n0 | n_gt0]; first by rewrite nth_default // sz_r n0. by apply/mem_rP=> //; rewrite mem_nth ?sz_r. case: {Dp mem_rP}r r_z r_arg => // y r1; rewrite inE => /predU1P[-> _|r1z]. - by apply/implyP=> ->; rewrite lerr. + by apply/implyP=> ->; rewrite lexx. by move/(order_path_min argCle_trans)/allP->. Qed. @@ -4411,11 +4087,11 @@ Let nz2 : 2%:R != 0 :> C. Proof. by rewrite pnatr_eq0. Qed. Lemma normCKC x : `|x| ^+ 2 = x^* * x. Proof. by rewrite normCK mulrC. Qed. Lemma mul_conjC_ge0 x : 0 <= x * x^*. -Proof. by rewrite -normCK exprn_ge0 ?normr_ge0. Qed. +Proof. by rewrite -normCK exprn_ge0. Qed. Lemma mul_conjC_gt0 x : (0 < x * x^*) = (x != 0). Proof. -have [->|x_neq0] := altP eqP; first by rewrite rmorph0 mulr0. +have [->|x_neq0] := eqVneq; first by rewrite rmorph0 mulr0. by rewrite -normCK exprn_gt0 ?normr_gt0. Qed. @@ -4493,16 +4169,13 @@ Proof. by move=> n_gt0; rewrite -{1}(rootC0 n) eqr_rootC. Qed. (* Rectangular coordinates. *) Lemma nonRealCi : ('i : C) \isn't real. -Proof. by rewrite realEsqr sqrCi oppr_ge0 ltr_geF ?ltr01. Qed. +Proof. by rewrite realEsqr sqrCi oppr_ge0 lt_geF ?ltr01. Qed. Lemma neq0Ci : 'i != 0 :> C. Proof. by apply: contraNneq nonRealCi => ->; apply: real0. Qed. Lemma normCi : `|'i| = 1 :> C. -Proof. -apply/eqP; rewrite -(@pexpr_eq1 _ _ 2) ?normr_ge0 //. -by rewrite -normrX sqrCi normrN1. -Qed. +Proof. by apply/eqP; rewrite -(@pexpr_eq1 _ _ 2) // -normrX sqrCi normrN1. Qed. Lemma invCi : 'i^-1 = - 'i :> C. Proof. by rewrite -div1r -[1]opprK -sqrCi mulNr mulfK ?neq0Ci. Qed. @@ -4630,19 +4303,19 @@ Proof. by move=> x y Rx Ry; rewrite /= invC_norm conjC_rect // mulrC normC2_rect. Qed. -Lemma lerif_normC_Re_Creal z : `|'Re z| <= `|z| ?= iff (z \is real). +Lemma leif_normC_Re_Creal z : `|'Re z| <= `|z| ?= iff (z \is real). Proof. -rewrite -(mono_in_lerif ler_sqr); try by rewrite qualifE normr_ge0. -rewrite normCK conj_Creal // normC2_Re_Im -expr2. -rewrite addrC -lerif_subLR subrr (sameP (Creal_ImP _) eqP) -sqrf_eq0 eq_sym. -by apply: lerif_eq; rewrite -realEsqr. +rewrite -(mono_in_leif ler_sqr); try by rewrite qualifE. +rewrite [`|'Re _| ^+ 2]normCK conj_Creal // normC2_Re_Im -expr2. +rewrite addrC -leif_subLR subrr (sameP (Creal_ImP _) eqP) -sqrf_eq0 eq_sym. +by apply: leif_eq; rewrite -realEsqr. Qed. -Lemma lerif_Re_Creal z : 'Re z <= `|z| ?= iff (0 <= z). +Lemma leif_Re_Creal z : 'Re z <= `|z| ?= iff (0 <= z). Proof. have ubRe: 'Re z <= `|'Re z| ?= iff (0 <= 'Re z). - by rewrite ger0_def eq_sym; apply/lerif_eq/real_ler_norm. -congr (_ <= _ ?= iff _): (lerif_trans ubRe (lerif_normC_Re_Creal z)). + by rewrite ger0_def eq_sym; apply/leif_eq/real_ler_norm. +congr (_ <= _ ?= iff _): (leif_trans ubRe (leif_normC_Re_Creal z)). apply/andP/idP=> [[zRge0 /Creal_ReP <- //] | z_ge0]. by have Rz := ger0_real z_ge0; rewrite (Creal_ReP _ _). Qed. @@ -4693,50 +4366,50 @@ suffices /existsP[i ltRwi0]: [exists i : 'I_n, 'Re (w ^+ i) < 0]. by exists (w ^+ i) => //; rewrite exprAC wn1 expr1n. apply: contra_eqT (congr1 Re pw_0) => /existsPn geRw0. rewrite raddf_sum raddf0 /= (bigD1 (Ordinal (ltnW n_gt1))) //=. -rewrite (Creal_ReP _ _) ?rpred1 // gtr_eqF ?ltr_paddr ?ltr01 //=. -by apply: sumr_ge0 => i _; rewrite real_lerNgt ?rpred0. +rewrite (Creal_ReP _ _) ?rpred1 // gt_eqF ?ltr_paddr ?ltr01 //=. +by apply: sumr_ge0 => i _; rewrite real_leNgt ?rpred0. Qed. Lemma Im_rootC_ge0 n x : (n > 1)%N -> 0 <= 'Im (n.-root x). Proof. set y := n.-root x => n_gt1; have n_gt0 := ltnW n_gt1. -apply: wlog_neg; rewrite -real_ltrNge ?rpred0 // => ltIy0. +apply: wlog_neg; rewrite -real_ltNge ?rpred0 // => ltIy0. suffices [z zn_x leI0z]: exists2 z, z ^+ n = x & 'Im z >= 0. by rewrite /y; case_rootC => /= y1 _ /(_ z n_gt0 zn_x)/argCleP[]. have [w wn1 ltRw0] := neg_unity_root n_gt1. wlog leRI0yw: w wn1 ltRw0 / 0 <= 'Re y * 'Im w. move=> IHw; have: 'Re y * 'Im w \is real by rewrite rpredM. - case/real_ger0P=> [|/ltrW leRIyw0]; first exact: IHw. + case/real_ge0P=> [|/ltW leRIyw0]; first exact: IHw. apply: (IHw w^*); rewrite ?Re_conj ?Im_conj ?mulrN ?oppr_ge0 //. by rewrite -rmorphX wn1 rmorph1. exists (w * y); first by rewrite exprMn wn1 mul1r rootCK. rewrite [w]Crect [y]Crect mulC_rect. -by rewrite Im_rect ?rpredD ?rpredN 1?rpredM // addr_ge0 // ltrW ?nmulr_rgt0. +by rewrite Im_rect ?rpredD ?rpredN 1?rpredM // addr_ge0 // ltW ?nmulr_rgt0. Qed. Lemma rootC_lt0 n x : (1 < n)%N -> (n.-root x < 0) = false. Proof. set y := n.-root x => n_gt1; have n_gt0 := ltnW n_gt1. -apply: negbTE; apply: wlog_neg => /negbNE lt0y; rewrite ler_gtF //. +apply: negbTE; apply: wlog_neg => /negbNE lt0y; rewrite le_gtF //. have Rx: x \is real by rewrite -[x](rootCK n_gt0) rpredX // ltr0_real. have Re_y: 'Re y = y by apply/Creal_ReP; rewrite ltr0_real. have [z zn_x leR0z]: exists2 z, z ^+ n = x & 'Re z >= 0. have [w wn1 ltRw0] := neg_unity_root n_gt1. exists (w * y); first by rewrite exprMn wn1 mul1r rootCK. - by rewrite ReMr ?ltr0_real // ltrW // nmulr_lgt0. + by rewrite ReMr ?ltr0_real // ltW // nmulr_lgt0. without loss leI0z: z zn_x leR0z / 'Im z >= 0. move=> IHz; have: 'Im z \is real by []. - case/real_ger0P=> [|/ltrW leIz0]; first exact: IHz. + case/real_ge0P=> [|/ltW leIz0]; first exact: IHz. apply: (IHz z^*); rewrite ?Re_conj ?Im_conj ?oppr_ge0 //. by rewrite -rmorphX zn_x conj_Creal. -by apply: ler_trans leR0z _; rewrite -Re_y ?rootC_Re_max ?ltr0_real. +by apply: le_trans leR0z _; rewrite -Re_y ?rootC_Re_max ?ltr0_real. Qed. Lemma rootC_ge0 n x : (n > 0)%N -> (0 <= n.-root x) = (0 <= x). Proof. set y := n.-root x => n_gt0. apply/idP/idP=> [/(exprn_ge0 n) | x_ge0]; first by rewrite rootCK. -rewrite -(ger_lerif (lerif_Re_Creal y)). +rewrite -(ge_leif (leif_Re_Creal y)). have Ray: `|y| \is real by apply: normr_real. rewrite -(Creal_ReP _ Ray) rootC_Re_max ?(Creal_ImP _ Ray) //. by rewrite -normrX rootCK // ger0_norm. @@ -4747,25 +4420,25 @@ Proof. by move=> n_gt0; rewrite !lt0r rootC_ge0 ?rootC_eq0. Qed. Lemma rootC_le0 n x : (1 < n)%N -> (n.-root x <= 0) = (x == 0). Proof. -by move=> n_gt1; rewrite ler_eqVlt rootC_lt0 // orbF rootC_eq0 1?ltnW. +by move=> n_gt1; rewrite le_eqVlt rootC_lt0 // orbF rootC_eq0 1?ltnW. Qed. Lemma ler_rootCl n : (n > 0)%N -> {in Num.nneg, {mono n.-root : x y / x <= y}}. Proof. move=> n_gt0 x x_ge0 y; have [y_ge0 | not_y_ge0] := boolP (0 <= y). by rewrite -(ler_pexpn2r n_gt0) ?qualifE ?rootC_ge0 ?rootCK. -rewrite (contraNF (@ler_trans _ _ 0 _ _)) ?rootC_ge0 //. -by rewrite (contraNF (ler_trans x_ge0)). +rewrite (contraNF (@le_trans _ _ _ 0 _ _)) ?rootC_ge0 //. +by rewrite (contraNF (le_trans x_ge0)). Qed. Lemma ler_rootC n : (n > 0)%N -> {in Num.nneg &, {mono n.-root : x y / x <= y}}. Proof. by move=> n_gt0 x y x_ge0 _; apply: ler_rootCl. Qed. Lemma ltr_rootCl n : (n > 0)%N -> {in Num.nneg, {mono n.-root : x y / x < y}}. -Proof. by move=> n_gt0 x x_ge0 y; rewrite !ltr_def ler_rootCl ?eqr_rootC. Qed. +Proof. by move=> n_gt0 x x_ge0 y; rewrite !lt_def ler_rootCl ?eqr_rootC. Qed. Lemma ltr_rootC n : (n > 0)%N -> {in Num.nneg &, {mono n.-root : x y / x < y}}. -Proof. by move/ler_rootC/lerW_mono_in. Qed. +Proof. by move/ler_rootC/leW_mono_in. Qed. Lemma exprCK n x : (0 < n)%N -> 0 <= x -> n.-root (x ^+ n) = x. Proof. @@ -4776,8 +4449,7 @@ Qed. Lemma norm_rootC n x : `|n.-root x| = n.-root `|x|. Proof. have [-> | n_gt0] := posnP n; first by rewrite !root0C normr0. -apply/eqP; rewrite -(eqr_expn2 n_gt0) ?rootC_ge0 ?normr_ge0 //. -by rewrite -normrX !rootCK. +by apply/eqP; rewrite -(eqr_expn2 n_gt0) ?rootC_ge0 // -normrX !rootCK. Qed. Lemma rootCX n x k : (n > 0)%N -> 0 <= x -> n.-root (x ^+ k) = n.-root x ^+ k. @@ -4809,31 +4481,31 @@ by move=> n_gt0; rewrite -{1}(rootC1 n_gt0) ler_rootCl // qualifE ler01. Qed. Lemma rootC_gt1 n x : (n > 0)%N -> (n.-root x > 1) = (x > 1). -Proof. by move=> n_gt0; rewrite !ltr_def rootC_eq1 ?rootC_ge1. Qed. +Proof. by move=> n_gt0; rewrite !lt_def rootC_eq1 ?rootC_ge1. Qed. Lemma rootC_le1 n x : (n > 0)%N -> 0 <= x -> (n.-root x <= 1) = (x <= 1). Proof. by move=> n_gt0 x_ge0; rewrite -{1}(rootC1 n_gt0) ler_rootCl. Qed. Lemma rootC_lt1 n x : (n > 0)%N -> 0 <= x -> (n.-root x < 1) = (x < 1). -Proof. by move=> n_gt0 x_ge0; rewrite !ltr_neqAle rootC_eq1 ?rootC_le1. Qed. +Proof. by move=> n_gt0 x_ge0; rewrite !lt_neqAle rootC_eq1 ?rootC_le1. Qed. Lemma rootCMl n x z : 0 <= x -> n.-root (x * z) = n.-root x * n.-root z. Proof. rewrite le0r => /predU1P[-> | x_gt0]; first by rewrite !(mul0r, rootC0). have [| n_gt1 | ->] := ltngtP n 1; last by rewrite !root1C. by case: n => //; rewrite !root0C mul0r. -have [x_ge0 n_gt0] := (ltrW x_gt0, ltnW n_gt1). +have [x_ge0 n_gt0] := (ltW x_gt0, ltnW n_gt1). have nx_gt0: 0 < n.-root x by rewrite rootC_gt0. -have Rnx: n.-root x \is real by rewrite ger0_real ?ltrW. +have Rnx: n.-root x \is real by rewrite ger0_real ?ltW. apply: eqC_semipolar; last 1 first; try apply/eqP. - by rewrite ImMl // !(Im_rootC_ge0, mulr_ge0, rootC_ge0). -- by rewrite -(eqr_expn2 n_gt0) ?normr_ge0 // -!normrX exprMn !rootCK. -rewrite eqr_le; apply/andP; split; last first. +- by rewrite -(eqr_expn2 n_gt0) // -!normrX exprMn !rootCK. +rewrite eq_le; apply/andP; split; last first. rewrite rootC_Re_max ?exprMn ?rootCK ?ImMl //. - by rewrite mulr_ge0 ?Im_rootC_ge0 ?ltrW. -rewrite -[n.-root _](mulVKf (negbT (gtr_eqF nx_gt0))) !(ReMl Rnx) //. -rewrite ler_pmul2l // rootC_Re_max ?exprMn ?exprVn ?rootCK ?mulKf ?gtr_eqF //. -by rewrite ImMl ?rpredV // mulr_ge0 ?invr_ge0 ?Im_rootC_ge0 ?ltrW. + by rewrite mulr_ge0 ?Im_rootC_ge0 ?ltW. +rewrite -[n.-root _](mulVKf (negbT (gt_eqF nx_gt0))) !(ReMl Rnx) //. +rewrite ler_pmul2l // rootC_Re_max ?exprMn ?exprVn ?rootCK ?mulKf ?gt_eqF //. +by rewrite ImMl ?rpredV // mulr_ge0 ?invr_ge0 ?Im_rootC_ge0 ?ltW. Qed. Lemma rootCMr n x z : 0 <= x -> n.-root (z * x) = n.-root z * n.-root x. @@ -4850,18 +4522,18 @@ Qed. (* The proper form of the Arithmetic - Geometric Mean inequality. *) -Lemma lerif_rootC_AGM (I : finType) (A : {pred I}) (n := #|A|) E : +Lemma leif_rootC_AGM (I : finType) (A : {pred I}) (n := #|A|) E : {in A, forall i, 0 <= E i} -> n.-root (\prod_(i in A) E i) <= (\sum_(i in A) E i) / n%:R ?= iff [forall i in A, forall j in A, E i == E j]. Proof. move=> Ege0; have [n0 | n_gt0] := posnP n. - rewrite n0 root0C invr0 mulr0; apply/lerif_refl/forall_inP=> i. + rewrite n0 root0C invr0 mulr0; apply/leif_refl/forall_inP=> i. by rewrite (card0_eq n0). -rewrite -(mono_in_lerif (ler_pexpn2r n_gt0)) ?rootCK //=; first 1 last. +rewrite -(mono_in_leif (ler_pexpn2r n_gt0)) ?rootCK //=; first 1 last. - by rewrite qualifE rootC_ge0 // prodr_ge0. - by rewrite rpred_div ?rpred_nat ?rpred_sum. -exact: lerif_AGM. +exact: leif_AGM. Qed. (* Square root. *) @@ -4893,13 +4565,13 @@ Proof. apply: (iffP andP) => [[leI0x not_gt0x] | <-]; last first. by rewrite sqrtC_lt0 Im_rootC_ge0. have /eqP := sqrtCK (x ^+ 2); rewrite eqf_sqr => /pred2P[] // defNx. -apply: sqrCK; rewrite -real_lerNgt ?rpred0 // in not_gt0x; -apply/Creal_ImP/ler_anti; +apply: sqrCK; rewrite -real_leNgt ?rpred0 // in not_gt0x; +apply/Creal_ImP/le_anti; by rewrite leI0x -oppr_ge0 -raddfN -defNx Im_rootC_ge0. Qed. Lemma normC_def x : `|x| = sqrtC (x * x^*). -Proof. by rewrite -normCK sqrCK ?normr_ge0. Qed. +Proof. by rewrite -normCK sqrCK. Qed. Lemma norm_conjC x : `|x^*| = `|x|. Proof. by rewrite !normC_def conjCK mulrC. Qed. @@ -4919,7 +4591,7 @@ Lemma normC_add_eq x y : Proof. move=> lin_xy; apply: sig2_eqW; pose u z := if z == 0 then 1 else z / `|z|. have uE z: (`|u z| = 1) * (`|z| * u z = z). - rewrite /u; have [->|nz_z] := altP eqP; first by rewrite normr0 normr1 mul0r. + rewrite /u; have [->|nz_z] := eqVneq; first by rewrite normr0 normr1 mul0r. by rewrite normf_div normr_id mulrCA divff ?mulr1 ?normr_eq0. have [->|nz_x] := eqVneq x 0; first by exists (u y); rewrite uE ?normr0 ?mul0r. exists (u x); rewrite uE // /u (negPf nz_x); congr (_ , _). @@ -4950,8 +4622,8 @@ have [Qj | /nandP[/negP[]// | /negbNE/eqP->]] := boolP (Q j); last first. by rewrite mulrC divfK. have: `|F i + F j| = `|F i| + `|F j|. do [rewrite !(bigD1 j Qj) /=; set z := \sum_(k | _) `|_|] in norm_sumF. - apply/eqP; rewrite eqr_le ler_norm_add -(ler_add2r z) -addrA -norm_sumF addrA. - by rewrite (ler_trans (ler_norm_add _ _)) // ler_add2l ler_norm_sum. + apply/eqP; rewrite eq_le ler_norm_add -(ler_add2r z) -addrA -norm_sumF addrA. + by rewrite (le_trans (ler_norm_add _ _)) // ler_add2l ler_norm_sum. by case/normC_add_eq=> k _ [/(canLR (mulKf nzFi)) <-]; rewrite -(mulrC (F i)). Qed. @@ -4970,14 +4642,14 @@ Lemma normC_sum_upper (I : finType) (P : pred I) (F G : I -> C) : forall i, P i -> F i = G i. Proof. set sumF := \sum_(i | _) _; set sumG := \sum_(i | _) _ => leFG eq_sumFG. -have posG i: P i -> 0 <= G i by move/leFG; apply: ler_trans; apply: normr_ge0. +have posG i: P i -> 0 <= G i by move/leFG; apply: le_trans. have norm_sumG: `|sumG| = sumG by rewrite ger0_norm ?sumr_ge0. have norm_sumF: `|sumF| = \sum_(i | P i) `|F i|. - apply/eqP; rewrite eqr_le ler_norm_sum eq_sumFG norm_sumG -subr_ge0 -sumrB. + apply/eqP; rewrite eq_le ler_norm_sum eq_sumFG norm_sumG -subr_ge0 -sumrB. by rewrite sumr_ge0 // => i Pi; rewrite subr_ge0 ?leFG. have [t _ defF] := normC_sum_eq norm_sumF. have [/(psumr_eq0P posG) G0 i Pi | nz_sumG] := eqVneq sumG 0. - by apply/eqP; rewrite G0 // -normr_eq0 eqr_le normr_ge0 -(G0 i Pi) leFG. + by apply/eqP; rewrite G0 // -normr_eq0 eq_le normr_ge0 -(G0 i Pi) leFG. have t1: t = 1. apply: (mulfI nz_sumG); rewrite mulr1 -{1}norm_sumG -eq_sumFG norm_sumF. by rewrite mulr_suml -(eq_bigr _ defF). @@ -5010,41 +4682,183 @@ Arguments sqrCK_P {C x}. End Theory. +(*************) +(* FACTORIES *) +(*************) + +Module NumMixin. +Section NumMixin. +Variable (R : idomainType). + +Record of_ := Mixin { + le : rel R; + lt : rel R; + norm : R -> R; + normD : forall x y, le (norm (x + y)) (norm x + norm y); + addr_gt0 : forall x y, lt 0 x -> lt 0 y -> lt 0 (x + y); + norm_eq0 : forall x, norm x = 0 -> x = 0; + ger_total : forall x y, le 0 x -> le 0 y -> le x y || le y x; + normM : {morph norm : x y / x * y}; + le_def : forall x y, (le x y) = (norm (y - x) == y - x); + lt_def : forall x y, (lt x y) = (y != x) && (le x y) +}. + +Variable (m : of_). + +Local Notation "x <= y" := (le m x y) : ring_scope. +Local Notation "x < y" := (lt m x y) : ring_scope. +Local Notation "`| x |" := (norm m x) : ring_scope. + +Lemma ltrr x : x < x = false. Proof. by rewrite lt_def eqxx. Qed. + +Lemma ge0_def x : (0 <= x) = (`|x| == x). +Proof. by rewrite le_def subr0. Qed. + +Lemma subr_ge0 x y : (0 <= x - y) = (y <= x). +Proof. by rewrite ge0_def -le_def. Qed. + +Lemma subr_gt0 x y : (0 < y - x) = (x < y). +Proof. by rewrite !lt_def subr_eq0 subr_ge0. Qed. + +Lemma lt_trans : transitive (lt m). +Proof. +move=> y x z le_xy le_yz. +by rewrite -subr_gt0 -(subrK y z) -addrA addr_gt0 // subr_gt0. +Qed. + +Lemma le01 : 0 <= 1. +Proof. +have n1_nz: `|1| != 0 :> R by apply: contraNneq (@oner_neq0 R) => /norm_eq0->. +by rewrite ge0_def -(inj_eq (mulfI n1_nz)) -normM !mulr1. +Qed. + +Lemma lt01 : 0 < 1. +Proof. by rewrite lt_def oner_neq0 le01. Qed. + +Lemma ltW x y : x < y -> x <= y. Proof. by rewrite lt_def => /andP[]. Qed. + +Lemma lerr x : x <= x. +Proof. +have n2: `|2%:R| == 2%:R :> R by rewrite -ge0_def ltW ?addr_gt0 ?lt01. +rewrite le_def subrr -(inj_eq (addrI `|0|)) addr0 -mulr2n -mulr_natr. +by rewrite -(eqP n2) -normM mul0r. +Qed. + +Lemma le_def' x y : (x <= y) = (x == y) || (x < y). +Proof. by rewrite lt_def; case: eqVneq => //= ->; rewrite lerr. Qed. + +Lemma le_trans : transitive (le m). +by move=> y x z; rewrite !le_def' => /predU1P [->|hxy] // /predU1P [<-|hyz]; + rewrite ?hxy ?(lt_trans hxy hyz) orbT. +Qed. + +Lemma normrMn x n : `|x *+ n| = `|x| *+ n. +Proof. +rewrite -mulr_natr -[RHS]mulr_natr normM. +congr (_ * _); apply/eqP; rewrite -ge0_def. +elim: n => [|n ih]; [exact: lerr | apply: (le_trans ih)]. +by rewrite le_def -natrB // subSnn -[_%:R]subr0 -le_def mulr1n le01. +Qed. + +Lemma normrN1 : `|-1| = 1 :> R. +Proof. +have: `|-1| ^+ 2 == 1 :> R + by rewrite expr2 /= -normM mulrNN mul1r -[1]subr0 -le_def le01. +rewrite sqrf_eq1 => /predU1P [] //; rewrite -[-1]subr0 -le_def. +have ->: 0 <= -1 = (-1 == 0 :> R) || (0 < -1) + by rewrite lt_def; case: eqP => // ->; rewrite lerr. +by rewrite oppr_eq0 oner_eq0 => /(addr_gt0 lt01); rewrite subrr ltrr. +Qed. + +Lemma normrN x : `|- x| = `|x|. +Proof. by rewrite -mulN1r normM -[RHS]mul1r normrN1. Qed. + +Definition lePOrderMixin : ltPOrderMixin R := + LtPOrderMixin le_def' ltrr lt_trans. + +Definition normedZmodMixin : + @normed_mixin_of R R lePOrderMixin := + @Num.NormedMixin _ _ lePOrderMixin (norm m) + (normD m) (@norm_eq0 m) normrMn normrN. + +Definition numDomainMixin : + @mixin_of R lePOrderMixin normedZmodMixin := + @Num.Mixin _ lePOrderMixin normedZmodMixin (@addr_gt0 m) + (@ger_total m) (@normM m) (@le_def m). + +End NumMixin. + +Module Exports. +Notation numMixin := of_. +Notation NumMixin := Mixin. +Coercion lePOrderMixin : numMixin >-> ltPOrderMixin. +Coercion normedZmodMixin : numMixin >-> normed_mixin_of. +Coercion numDomainMixin : numMixin >-> mixin_of. +End Exports. + +End NumMixin. +Import NumMixin.Exports. + Module RealMixin. +Section RealMixin. +Variables (R : numDomainType). -Section RealMixins. +Variable (real : real_axiom R). + +Lemma le_total : totalPOrderMixin R. +Proof. +move=> x y; move: (real (x - y)). +by rewrite unfold_in !ler_def subr0 add0r opprB orbC. +Qed. -Variables (R : idomainType) (le : rel R) (lt : rel R) (norm : R -> R). -Local Infix "<=" := le. -Local Infix "<" := lt. -Local Notation "`| x |" := (norm x) : ring_scope. +Definition totalMixin : + Order.Total.mixin_of (DistrLatticeType R le_total) := le_total. -Section LeMixin. +End RealMixin. -Hypothesis le0_add : forall x y, 0 <= x -> 0 <= y -> 0 <= x + y. -Hypothesis le0_mul : forall x y, 0 <= x -> 0 <= y -> 0 <= x * y. -Hypothesis le0_anti : forall x, 0 <= x -> x <= 0 -> x = 0. -Hypothesis sub_ge0 : forall x y, (0 <= y - x) = (x <= y). -Hypothesis le0_total : forall x, (0 <= x) || (x <= 0). -Hypothesis normN: forall x, `|- x| = `|x|. -Hypothesis ge0_norm : forall x, 0 <= x -> `|x| = x. -Hypothesis lt_def : forall x y, (x < y) = (y != x) && (x <= y). +Module Exports. +Coercion le_total : real_axiom >-> totalPOrderMixin. +Coercion totalMixin : real_axiom >-> totalOrderMixin. +End Exports. + +End RealMixin. +Import RealMixin.Exports. + +Module RealLeMixin. +Section RealLeMixin. +Variables (R : idomainType). + +Record of_ := Mixin { + le : rel R; + lt : rel R; + norm : R -> R; + le0_add : forall x y, le 0 x -> le 0 y -> le 0 (x + y); + le0_mul : forall x y, le 0 x -> le 0 y -> le 0 (x * y); + le0_anti : forall x, le 0 x -> le x 0 -> x = 0; + sub_ge0 : forall x y, le 0 (y - x) = le x y; + le0_total : forall x, le 0 x || le x 0; + normN : forall x, norm (- x) = norm x; + ge0_norm : forall x, le 0 x -> norm x = x; + lt_def : forall x y, lt x y = (y != x) && le x y; +}. + +Variable (m : of_). + +Local Notation "x <= y" := (le m x y) : ring_scope. +Local Notation "x < y" := (lt m x y) : ring_scope. +Local Notation "`| x |" := (norm m x) : ring_scope. Let le0N x : (0 <= - x) = (x <= 0). Proof. by rewrite -sub0r sub_ge0. Qed. Let leN_total x : 0 <= x \/ 0 <= - x. Proof. by apply/orP; rewrite le0N le0_total. Qed. -Let le00 : (0 <= 0). Proof. by have:= le0_total 0; rewrite orbb. Qed. -Let le01 : (0 <= 1). -Proof. -by case/orP: (le0_total 1)=> // ?; rewrite -[1]mul1r -mulrNN le0_mul ?le0N. -Qed. +Let le00 : 0 <= 0. Proof. by have:= le0_total m 0; rewrite orbb. Qed. Fact lt0_add x y : 0 < x -> 0 < y -> 0 < x + y. Proof. -rewrite !lt_def => /andP[x_neq0 l0x] /andP[y_neq0 l0y]; rewrite le0_add //. +rewrite !lt_def => /andP [x_neq0 l0x] /andP [y_neq0 l0y]; rewrite le0_add //. rewrite andbT addr_eq0; apply: contraNneq x_neq0 => hxy. -by rewrite [x]le0_anti // hxy -le0N opprK. +by rewrite [x](@le0_anti m) // hxy -le0N opprK. Qed. Fact eq0_norm x : `|x| = 0 -> x = 0. @@ -5060,7 +4874,7 @@ rewrite {x}subr0; apply/idP/eqP=> [/ge0_norm// | Dy]. by have [//| ny_ge0] := leN_total y; rewrite -Dy -normN ge0_norm. Qed. -Fact normM : {morph norm : x y / x * y}. +Fact normM : {morph norm m : x y / x * y}. Proof. move=> x y /=; wlog x_ge0 : x / 0 <= x. by move=> IHx; case: (leN_total x) => /IHx//; rewrite mulNr !normN. @@ -5079,73 +4893,728 @@ rewrite -normN ge0_norm //; have [hxy|hxy] := leN_total (x + y). by rewrite -normN ge0_norm // opprK addrCA addrNK le0_add. Qed. -Lemma le_total x y : (x <= y) || (y <= x). -Proof. by rewrite -sub_ge0 -opprB le0N orbC -sub_ge0 le0_total. Qed. +Fact le_total : total (le m). +Proof. by move=> x y; rewrite -sub_ge0 -opprB le0N orbC -sub_ge0 le0_total. Qed. -Definition Le := - Mixin le_normD lt0_add eq0_norm (in2W le_total) normM le_def lt_def. +Definition numMixin : numMixin R := + NumMixin le_normD lt0_add eq0_norm (in2W le_total) normM le_def (lt_def m). -Lemma Real (R' : numDomainType) & phant R' : - R' = NumDomainType R Le -> real_axiom R'. -Proof. by move->. Qed. +Definition orderMixin : + totalPOrderMixin (POrderType ring_display R numMixin) := + le_total. -End LeMixin. +End RealLeMixin. + +Module Exports. +Notation realLeMixin := of_. +Notation RealLeMixin := Mixin. +Coercion numMixin : realLeMixin >-> NumMixin.of_. +Coercion orderMixin : realLeMixin >-> totalPOrderMixin. +End Exports. + +End RealLeMixin. + +Module RealLtMixin. +Section RealLtMixin. +Variables (R : idomainType). + +Record of_ := Mixin { + lt : rel R; + le : rel R; + norm : R -> R; + lt0_add : forall x y, lt 0 x -> lt 0 y -> lt 0 (x + y); + lt0_mul : forall x y, lt 0 x -> lt 0 y -> lt 0 (x * y); + lt0_ngt0 : forall x, lt 0 x -> ~~ (lt x 0); + sub_gt0 : forall x y, lt 0 (y - x) = lt x y; + lt0_total : forall x, x != 0 -> lt 0 x || lt x 0; + normN : forall x, norm (- x) = norm x; + ge0_norm : forall x, le 0 x -> norm x = x; + le_def : forall x y, le x y = (x == y) || lt x y; +}. -Section LtMixin. +Variable (m : of_). -Hypothesis lt0_add : forall x y, 0 < x -> 0 < y -> 0 < x + y. -Hypothesis lt0_mul : forall x y, 0 < x -> 0 < y -> 0 < x * y. -Hypothesis lt0_ngt0 : forall x, 0 < x -> ~~ (x < 0). -Hypothesis sub_gt0 : forall x y, (0 < y - x) = (x < y). -Hypothesis lt0_total : forall x, x != 0 -> (0 < x) || (x < 0). -Hypothesis normN : forall x, `|- x| = `|x|. -Hypothesis ge0_norm : forall x, 0 <= x -> `|x| = x. -Hypothesis le_def : forall x y, (x <= y) = (y == x) || (x < y). +Local Notation "x < y" := (lt m x y) : ring_scope. +Local Notation "x <= y" := (le m x y) : ring_scope. +Local Notation "`| x |" := (norm m x) : ring_scope. + +Fact lt0N x : (- x < 0) = (0 < x). +Proof. by rewrite -sub_gt0 add0r opprK. Qed. +Let leN_total x : 0 <= x \/ 0 <= - x. +Proof. +rewrite !le_def [_ == - x]eq_sym oppr_eq0 -[0 < - x]lt0N opprK. +apply/orP; case: (eqVneq x) => //=; exact: lt0_total. +Qed. + +Let le00 : (0 <= 0). Proof. by rewrite le_def eqxx. Qed. + +Fact sub_ge0 x y : (0 <= y - x) = (x <= y). +Proof. by rewrite !le_def eq_sym subr_eq0 eq_sym sub_gt0. Qed. Fact le0_add x y : 0 <= x -> 0 <= y -> 0 <= x + y. Proof. -rewrite !le_def => /predU1P[->|x_gt0]; first by rewrite add0r. -by case/predU1P=> [->|y_gt0]; rewrite ?addr0 ?x_gt0 ?lt0_add // orbT. +rewrite !le_def => /predU1P [<-|x_gt0]; first by rewrite add0r. +by case/predU1P=> [<-|y_gt0]; rewrite ?addr0 ?x_gt0 ?lt0_add // orbT. Qed. Fact le0_mul x y : 0 <= x -> 0 <= y -> 0 <= x * y. Proof. -rewrite !le_def => /predU1P[->|x_gt0]; first by rewrite mul0r eqxx. -by case/predU1P=> [->|y_gt0]; rewrite ?mulr0 ?eqxx // orbC lt0_mul. +rewrite !le_def => /predU1P [<-|x_gt0]; first by rewrite mul0r eqxx. +by case/predU1P=> [<-|y_gt0]; rewrite ?mulr0 ?eqxx ?lt0_mul // orbT. +Qed. + +Fact normM : {morph norm m : x y / x * y}. +Proof. +move=> x y /=; wlog x_ge0 : x / 0 <= x. + by move=> IHx; case: (leN_total x) => /IHx//; rewrite mulNr !normN. +wlog y_ge0 : y / 0 <= y; last by rewrite ?ge0_norm ?le0_mul. +by move=> IHy; case: (leN_total y) => /IHy//; rewrite mulrN !normN. +Qed. + +Fact le_normD x y : `|x + y| <= `|x| + `|y|. +Proof. +wlog x_ge0 : x y / 0 <= x. + by move=> IH; case: (leN_total x) => /IH// /(_ (- y)); rewrite -opprD !normN. +rewrite -sub_ge0 ge0_norm //; have [y_ge0 | ny_ge0] := leN_total y. + by rewrite !ge0_norm ?subrr ?le0_add. +rewrite -normN ge0_norm //; have [hxy|hxy] := leN_total (x + y). + by rewrite ge0_norm // opprD addrCA -addrA addKr le0_add. +by rewrite -normN ge0_norm // opprK addrCA addrNK le0_add. Qed. -Fact le0_anti x : 0 <= x -> x <= 0 -> x = 0. -Proof. by rewrite !le_def => /predU1P[] // /lt0_ngt0/negPf-> /predU1P[]. Qed. +Fact eq0_norm x : `|x| = 0 -> x = 0. +Proof. +case: (leN_total x) => /ge0_norm => [-> // | Dnx nx0]. +by rewrite -[x]opprK -Dnx normN nx0 oppr0. +Qed. -Fact sub_ge0 x y : (0 <= y - x) = (x <= y). -Proof. by rewrite !le_def subr_eq0 sub_gt0. Qed. +Fact le_def' x y : (x <= y) = (`|y - x| == y - x). +Proof. +wlog ->: x y / x = 0 by move/(_ 0 (y - x)); rewrite subr0 sub_ge0 => ->. +rewrite {x}subr0; apply/idP/eqP=> [/ge0_norm// | Dy]. +by have [//| ny_ge0] := leN_total y; rewrite -Dy -normN ge0_norm. +Qed. Fact lt_def x y : (x < y) = (y != x) && (x <= y). Proof. -rewrite le_def; case: eqP => //= ->; rewrite -sub_gt0 subrr. +rewrite le_def; case: eqVneq => //= ->; rewrite -sub_gt0 subrr. by apply/idP=> lt00; case/negP: (lt0_ngt0 lt00). Qed. -Fact le0_total x : (0 <= x) || (x <= 0). -Proof. by rewrite !le_def [0 == _]eq_sym; have [|/lt0_total] := altP eqP. Qed. +Fact le_total : total (le m). +Proof. +move=> x y; rewrite !le_def; case: eqVneq => [->|] //=; rewrite -subr_eq0. +by move/(lt0_total m); rewrite -(sub_gt0 _ (x - y)) sub0r opprB !sub_gt0 orbC. +Qed. -Definition Lt := - Le le0_add le0_mul le0_anti sub_ge0 le0_total normN ge0_norm lt_def. +Definition numMixin : numMixin R := + NumMixin le_normD (@lt0_add m) eq0_norm (in2W le_total) normM le_def' lt_def. -End LtMixin. +Definition orderMixin : + totalPOrderMixin (POrderType ring_display R numMixin) := + le_total. -End RealMixins. +End RealLtMixin. -End RealMixin. +Module Exports. +Notation realLtMixin := of_. +Notation RealLtMixin := Mixin. +Coercion numMixin : realLtMixin >-> NumMixin.of_. +Coercion orderMixin : realLtMixin >-> totalPOrderMixin. +End Exports. + +End RealLtMixin. End Num. -Export Num.NumDomain.Exports Num.NumField.Exports Num.ClosedField.Exports. +Export Num.NumDomain.Exports Num.NormedZmodule.Exports. +Export Num.NumDomain_joins.Exports. +Export Num.NumField.Exports Num.ClosedField.Exports. Export Num.RealDomain.Exports Num.RealField.Exports. Export Num.ArchimedeanField.Exports Num.RealClosedField.Exports. Export Num.Syntax Num.PredInstances. +Export Num.NumMixin.Exports Num.RealMixin.Exports. +Export Num.RealLeMixin.Exports Num.RealLtMixin.Exports. -Notation RealLeMixin := Num.RealMixin.Le. -Notation RealLtMixin := Num.RealMixin.Lt. -Notation RealLeAxiom R := (Num.RealMixin.Real (Phant R) (erefl _)). Notation ImaginaryMixin := Num.ClosedField.ImaginaryMixin. + +(* compatibility module *) +Module mc_1_10. +Module Num. +(* If you failed to process the next line in the PG or the CoqIDE, replace *) +(* all the "ssrnum.Num" with "Top.Num" in this module to process them, and *) +(* revert them in order to compile and commit. This problem will be solved *) +(* in Coq 8.10. See also: https://github.com/math-comp/math-comp/pull/270. *) +Export ssrnum.Num. +Import ssrnum.Num.Def. + +Module Import Syntax. +Notation "`| x |" := + (@norm _ (@Num.NormedZmodule.numDomain_normedZmodType _) x) : ring_scope. +End Syntax. + +Module Import Theory. +Export ssrnum.Num.Theory. + +Section NumIntegralDomainTheory. +Variable R : numDomainType. +Implicit Types x y z : R. +Definition ltr_def x y : (x < y) = (y != x) && (x <= y) := lt_def x y. +Definition gerE x y : ge x y = (y <= x) := geE x y. +Definition gtrE x y : gt x y = (y < x) := gtE x y. +Definition lerr x : x <= x := lexx x. +Definition ltrr x : x < x = false := ltxx x. +Definition ltrW x y : x < y -> x <= y := @ltW _ _ x y. +Definition ltr_neqAle x y : (x < y) = (x != y) && (x <= y) := lt_neqAle x y. +Definition ler_eqVlt x y : (x <= y) = (x == y) || (x < y) := le_eqVlt x y. +Definition gtr_eqF x y : y < x -> x == y = false := @gt_eqF _ _ x y. +Definition ltr_eqF x y : x < y -> x == y = false := @lt_eqF _ _ x y. +Definition ler_asym : antisymmetric (@ler R) := le_anti. +Definition eqr_le x y : (x == y) = (x <= y <= x) := eq_le x y. +Definition ltr_trans : transitive (@ltr R) := lt_trans. +Definition ler_lt_trans y x z : x <= y -> y < z -> x < z := + @le_lt_trans _ _ y x z. +Definition ltr_le_trans y x z : x < y -> y <= z -> x < z := + @lt_le_trans _ _ y x z. +Definition ler_trans : transitive (@ler R) := le_trans. +Definition lterr := (lerr, ltrr). +Definition lerifP x y C : + reflect (x <= y ?= iff C) (if C then x == y else x < y) := leifP. +Definition ltr_asym x y : x < y < x = false := lt_asym x y. +Definition ler_anti : antisymmetric (@ler R) := le_anti. +Definition ltr_le_asym x y : x < y <= x = false := lt_le_asym x y. +Definition ler_lt_asym x y : x <= y < x = false := le_lt_asym x y. +Definition lter_anti := (=^~ eqr_le, ltr_asym, ltr_le_asym, ler_lt_asym). +Definition ltr_geF x y : x < y -> y <= x = false := @lt_geF _ _ x y. +Definition ler_gtF x y : x <= y -> y < x = false := @le_gtF _ _ x y. +Definition ltr_gtF x y : x < y -> y < x = false := @lt_gtF _ _ x y. +Definition normr0 : `|0| = 0 :> R := normr0 _. +Definition normrMn x n : `|x *+ n| = `|x| *+ n := normrMn x n. +Definition normr0P {x} : reflect (`|x| = 0) (x == 0) := normr0P. +Definition normr_eq0 x : (`|x| == 0) = (x == 0) := normr_eq0 x. +Definition normrN x : `|- x| = `|x| := normrN x. +Definition distrC x y : `|x - y| = `|y - x| := distrC x y. +Definition normr_id x : `| `|x| | = `|x| := normr_id x. +Definition normr_ge0 x : 0 <= `|x| := normr_ge0 x. +Definition normr_le0 x : (`|x| <= 0) = (x == 0) := normr_le0 x. +Definition normr_lt0 x : `|x| < 0 = false := normr_lt0 x. +Definition normr_gt0 x : (`|x| > 0) = (x != 0) := normr_gt0 x. +Definition normrE := (normr_id, normr0, @normr1 R, @normrN1 R, normr_ge0, + normr_eq0, normr_lt0, normr_le0, normr_gt0, normrN). +End NumIntegralDomainTheory. + +Section NumIntegralDomainMonotonyTheory. +Variables R R' : numDomainType. +Section AcrossTypes. +Variables (D D' : pred R) (f : R -> R'). +Definition ltrW_homo : {homo f : x y / x < y} -> {homo f : x y / x <= y} := + ltW_homo (f := f). +Definition ltrW_nhomo : {homo f : x y /~ x < y} -> {homo f : x y /~ x <= y} := + ltW_nhomo (f := f). +Definition inj_homo_ltr : + injective f -> {homo f : x y / x <= y} -> {homo f : x y / x < y} := + inj_homo_lt (f := f). +Definition inj_nhomo_ltr : + injective f -> {homo f : x y /~ x <= y} -> {homo f : x y /~ x < y} := + inj_nhomo_lt (f := f). +Definition incr_inj : {mono f : x y / x <= y} -> injective f := + inc_inj (f := f). +Definition decr_inj : {mono f : x y /~ x <= y} -> injective f := + dec_inj (f := f). +Definition lerW_mono : {mono f : x y / x <= y} -> {mono f : x y / x < y} := + leW_mono (f := f). +Definition lerW_nmono : {mono f : x y /~ x <= y} -> {mono f : x y /~ x < y} := + leW_nmono (f := f). +Definition ltrW_homo_in : + {in D & D', {homo f : x y / x < y}} -> {in D & D', {homo f : x y / x <= y}} := + ltW_homo_in (f := f). +Definition ltrW_nhomo_in : + {in D & D', {homo f : x y /~ x < y}} -> + {in D & D', {homo f : x y /~ x <= y}} := + ltW_nhomo_in (f := f). +Definition inj_homo_ltr_in : + {in D & D', injective f} -> {in D & D', {homo f : x y / x <= y}} -> + {in D & D', {homo f : x y / x < y}} := + inj_homo_lt_in (f := f). +Definition inj_nhomo_ltr_in : + {in D & D', injective f} -> {in D & D', {homo f : x y /~ x <= y}} -> + {in D & D', {homo f : x y /~ x < y}} := + inj_nhomo_lt_in (f := f). +Definition incr_inj_in : + {in D &, {mono f : x y / x <= y}} -> {in D &, injective f} := + inc_inj_in (f := f). +Definition decr_inj_in : + {in D &, {mono f : x y /~ x <= y}} -> {in D &, injective f} := + dec_inj_in (f := f). +Definition lerW_mono_in : + {in D &, {mono f : x y / x <= y}} -> {in D &, {mono f : x y / x < y}} := + leW_mono_in (f := f). +Definition lerW_nmono_in : + {in D &, {mono f : x y /~ x <= y}} -> {in D &, {mono f : x y /~ x < y}} := + leW_nmono_in (f := f). +End AcrossTypes. +Section NatToR. +Variables (D D' : pred nat) (f : nat -> R). +Definition ltnrW_homo : + {homo f : m n / (m < n)%N >-> m < n} -> + {homo f : m n / (m <= n)%N >-> m <= n} := + ltW_homo (f := f). +Definition ltnrW_nhomo : + {homo f : m n / (n < m)%N >-> m < n} -> + {homo f : m n / (n <= m)%N >-> m <= n} := + ltW_nhomo (f := f). +Definition inj_homo_ltnr : injective f -> + {homo f : m n / (m <= n)%N >-> m <= n} -> + {homo f : m n / (m < n)%N >-> m < n} := + inj_homo_lt (f := f). +Definition inj_nhomo_ltnr : injective f -> + {homo f : m n / (n <= m)%N >-> m <= n} -> + {homo f : m n / (n < m)%N >-> m < n} := + inj_nhomo_lt (f := f). +Definition incnr_inj : + {mono f : m n / (m <= n)%N >-> m <= n} -> injective f := + inc_inj (f := f). +Definition decnr_inj : + {mono f : m n / (n <= m)%N >-> m <= n} -> injective f := + dec_inj (f := f). +Definition decnr_inj_inj := decnr_inj. +Definition lenrW_mono : {mono f : m n / (m <= n)%N >-> m <= n} -> + {mono f : m n / (m < n)%N >-> m < n} := + leW_mono (f := f). +Definition lenrW_nmono : {mono f : m n / (n <= m)%N >-> m <= n} -> + {mono f : m n / (n < m)%N >-> m < n} := + leW_nmono (f := f). +Definition lenr_mono : {homo f : m n / (m < n)%N >-> m < n} -> + {mono f : m n / (m <= n)%N >-> m <= n} := + le_mono (f := f). +Definition lenr_nmono : + {homo f : m n / (n < m)%N >-> m < n} -> + {mono f : m n / (n <= m)%N >-> m <= n} := + le_nmono (f := f). +Definition ltnrW_homo_in : + {in D & D', {homo f : m n / (m < n)%N >-> m < n}} -> + {in D & D', {homo f : m n / (m <= n)%N >-> m <= n}} := + ltW_homo_in (f := f). +Definition ltnrW_nhomo_in : + {in D & D', {homo f : m n / (n < m)%N >-> m < n}} -> + {in D & D', {homo f : m n / (n <= m)%N >-> m <= n}} := + ltW_nhomo_in (f := f). +Definition inj_homo_ltnr_in : {in D & D', injective f} -> + {in D & D', {homo f : m n / (m <= n)%N >-> m <= n}} -> + {in D & D', {homo f : m n / (m < n)%N >-> m < n}} := + inj_homo_lt_in (f := f). +Definition inj_nhomo_ltnr_in : {in D & D', injective f} -> + {in D & D', {homo f : m n / (n <= m)%N >-> m <= n}} -> + {in D & D', {homo f : m n / (n < m)%N >-> m < n}} := + inj_nhomo_lt_in (f := f). +Definition incnr_inj_in : + {in D &, {mono f : m n / (m <= n)%N >-> m <= n}} -> {in D &, injective f} := + inc_inj_in (f := f). +Definition decnr_inj_in : + {in D &, {mono f : m n / (n <= m)%N >-> m <= n}} -> {in D &, injective f} := + dec_inj_in (f := f). +Definition decnr_inj_inj_in := decnr_inj_in. +Definition lenrW_mono_in : + {in D &, {mono f : m n / (m <= n)%N >-> m <= n}} -> + {in D &, {mono f : m n / (m < n)%N >-> m < n}} := + leW_mono_in (f := f). +Definition lenrW_nmono_in : + {in D &, {mono f : m n / (n <= m)%N >-> m <= n}} -> + {in D &, {mono f : m n / (n < m)%N >-> m < n}} := + leW_nmono_in (f := f). +Definition lenr_mono_in : + {in D &, {homo f : m n / (m < n)%N >-> m < n}} -> + {in D &, {mono f : m n / (m <= n)%N >-> m <= n}} := + le_mono_in (f := f). +Definition lenr_nmono_in : + {in D &, {homo f : m n / (n < m)%N >-> m < n}} -> + {in D &, {mono f : m n / (n <= m)%N >-> m <= n}} := + le_nmono_in (f := f). +End NatToR. +Section RToNat. +Variables (D D' : pred R) (f : R -> nat). +Definition ltrnW_homo : + {homo f : m n / m < n >-> (m < n)%N} -> + {homo f : m n / m <= n >-> (m <= n)%N} := + ltW_homo (f := f). +Definition ltrnW_nhomo : + {homo f : m n / n < m >-> (m < n)%N} -> + {homo f : m n / n <= m >-> (m <= n)%N} := + ltW_nhomo (f := f). +Definition inj_homo_ltrn : injective f -> + {homo f : m n / m <= n >-> (m <= n)%N} -> + {homo f : m n / m < n >-> (m < n)%N} := + inj_homo_lt (f := f). +Definition inj_nhomo_ltrn : injective f -> + {homo f : m n / n <= m >-> (m <= n)%N} -> + {homo f : m n / n < m >-> (m < n)%N} := + inj_nhomo_lt (f := f). +Definition incrn_inj : {mono f : m n / m <= n >-> (m <= n)%N} -> injective f := + inc_inj (f := f). +Definition decrn_inj : {mono f : m n / n <= m >-> (m <= n)%N} -> injective f := + dec_inj (f := f). +Definition lernW_mono : + {mono f : m n / m <= n >-> (m <= n)%N} -> + {mono f : m n / m < n >-> (m < n)%N} := + leW_mono (f := f). +Definition lernW_nmono : + {mono f : m n / n <= m >-> (m <= n)%N} -> + {mono f : m n / n < m >-> (m < n)%N} := + leW_nmono (f := f). +Definition ltrnW_homo_in : + {in D & D', {homo f : m n / m < n >-> (m < n)%N}} -> + {in D & D', {homo f : m n / m <= n >-> (m <= n)%N}} := + ltW_homo_in (f := f). +Definition ltrnW_nhomo_in : + {in D & D', {homo f : m n / n < m >-> (m < n)%N}} -> + {in D & D', {homo f : m n / n <= m >-> (m <= n)%N}} := + ltW_nhomo_in (f := f). +Definition inj_homo_ltrn_in : {in D & D', injective f} -> + {in D & D', {homo f : m n / m <= n >-> (m <= n)%N}} -> + {in D & D', {homo f : m n / m < n >-> (m < n)%N}} := + inj_homo_lt_in (f := f). +Definition inj_nhomo_ltrn_in : {in D & D', injective f} -> + {in D & D', {homo f : m n / n <= m >-> (m <= n)%N}} -> + {in D & D', {homo f : m n / n < m >-> (m < n)%N}} := + inj_nhomo_lt_in (f := f). +Definition incrn_inj_in : + {in D &, {mono f : m n / m <= n >-> (m <= n)%N}} -> {in D &, injective f} := + inc_inj_in (f := f). +Definition decrn_inj_in : + {in D &, {mono f : m n / n <= m >-> (m <= n)%N}} -> {in D &, injective f} := + dec_inj_in (f := f). +Definition lernW_mono_in : + {in D &, {mono f : m n / m <= n >-> (m <= n)%N}} -> + {in D &, {mono f : m n / m < n >-> (m < n)%N}} := + leW_mono_in (f := f). +Definition lernW_nmono_in : + {in D &, {mono f : m n / n <= m >-> (m <= n)%N}} -> + {in D &, {mono f : m n / n < m >-> (m < n)%N}} := + leW_nmono_in (f := f). +End RToNat. +End NumIntegralDomainMonotonyTheory. + +Section NumDomainOperationTheory. +Variable R : numDomainType. +Implicit Types x y z t : R. +Definition lerif_refl x C : reflect (x <= x ?= iff C) C := leif_refl. +Definition lerif_trans x1 x2 x3 C12 C23 : + x1 <= x2 ?= iff C12 -> x2 <= x3 ?= iff C23 -> x1 <= x3 ?= iff C12 && C23 := + @leif_trans _ _ x1 x2 x3 C12 C23. +Definition lerif_le x y : x <= y -> x <= y ?= iff (x >= y) := @leif_le _ _ x y. +Definition lerif_eq x y : x <= y -> x <= y ?= iff (x == y) := @leif_eq _ _ x y. +Definition ger_lerif x y C : x <= y ?= iff C -> (y <= x) = C := + @ge_leif _ _ x y C. +Definition ltr_lerif x y C : x <= y ?= iff C -> (x < y) = ~~ C := + @lt_leif _ _ x y C. +Definition normr_real x : `|x| \is real := normr_real x. +Definition ler_norm_sum I r (G : I -> R) (P : pred I): + `|\sum_(i <- r | P i) G i| <= \sum_(i <- r | P i) `|G i| := + ler_norm_sum r G P. +Definition ler_norm_sub x y : `|x - y| <= `|x| + `|y| := ler_norm_sub x y. +Definition ler_dist_add z x y : `|x - y| <= `|x - z| + `|z - y| := + ler_dist_add z x y. +Definition ler_sub_norm_add x y : `|x| - `|y| <= `|x + y| := + ler_sub_norm_add x y. +Definition ler_sub_dist x y : `|x| - `|y| <= `|x - y| := ler_sub_dist x y. +Definition ler_dist_dist x y : `| `|x| - `|y| | <= `|x - y| := + ler_dist_dist x y. +Definition ler_dist_norm_add x y : `| `|x| - `|y| | <= `|x + y| := + ler_dist_norm_add x y. +Definition ler_nnorml x y : y < 0 -> `|x| <= y = false := @ler_nnorml _ _ x y. +Definition ltr_nnorml x y : y <= 0 -> `|x| < y = false := @ltr_nnorml _ _ x y. +Definition lter_nnormr := (ler_nnorml, ltr_nnorml). +Definition mono_in_lerif (A : pred R) (f : R -> R) C : + {in A &, {mono f : x y / x <= y}} -> + {in A &, forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C)} := + @mono_in_leif _ _ A f C. +Definition mono_lerif (f : R -> R) C : + {mono f : x y / x <= y} -> + forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C) := + @mono_leif _ _ f C. +Definition nmono_in_lerif (A : pred R) (f : R -> R) C : + {in A &, {mono f : x y /~ x <= y}} -> + {in A &, forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C)} := + @nmono_in_leif _ _ A f C. +Definition nmono_lerif (f : R -> R) C : + {mono f : x y /~ x <= y} -> + forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C) := + @nmono_leif _ _ f C. +End NumDomainOperationTheory. + +Section RealDomainTheory. +Variable R : realDomainType. +Implicit Types x y z t : R. +Definition ler_total : total (@ler R) := le_total. +Definition ltr_total x y : x != y -> (x < y) || (y < x) := + @lt_total _ _ x y. +Definition wlog_ler P : + (forall a b, P b a -> P a b) -> (forall a b, a <= b -> P a b) -> + forall a b : R, P a b := + @wlog_le _ _ P. +Definition wlog_ltr P : + (forall a, P a a) -> + (forall a b, (P b a -> P a b)) -> (forall a b, a < b -> P a b) -> + forall a b : R, P a b := + @wlog_lt _ _ P. +Definition ltrNge x y : (x < y) = ~~ (y <= x) := @ltNge _ _ x y. +Definition lerNgt x y : (x <= y) = ~~ (y < x) := @leNgt _ _ x y. +Definition neqr_lt x y : (x != y) = (x < y) || (y < x) := @neq_lt _ _ x y. +Definition eqr_leLR x y z t : + (x <= y -> z <= t) -> (y < x -> t < z) -> (x <= y) = (z <= t) := + @eq_leLR _ _ x y z t. +Definition eqr_leRL x y z t : + (x <= y -> z <= t) -> (y < x -> t < z) -> (z <= t) = (x <= y) := + @eq_leRL _ _ x y z t. +Definition eqr_ltLR x y z t : + (x < y -> z < t) -> (y <= x -> t <= z) -> (x < y) = (z < t) := + @eq_ltLR _ _ x y z t. +Definition eqr_ltRL x y z t : + (x < y -> z < t) -> (y <= x -> t <= z) -> (z < t) = (x < y) := + @eq_ltRL _ _ x y z t. +End RealDomainTheory. + +Section RealDomainMonotony. +Variables (R : realDomainType) (R' : numDomainType) (D : pred R). +Variables (f : R -> R') (f' : R -> nat). +Definition ler_mono : {homo f : x y / x < y} -> {mono f : x y / x <= y} := + le_mono (f := f). +Definition homo_mono := ler_mono. +Definition ler_nmono : {homo f : x y /~ x < y} -> {mono f : x y /~ x <= y} := + le_nmono (f := f). +Definition nhomo_mono := ler_nmono. +Definition ler_mono_in : + {in D &, {homo f : x y / x < y}} -> {in D &, {mono f : x y / x <= y}} := + le_mono_in (f := f). +Definition homo_mono_in := ler_mono_in. +Definition ler_nmono_in : + {in D &, {homo f : x y /~ x < y}} -> {in D &, {mono f : x y /~ x <= y}} := + le_nmono_in (f := f). +Definition nhomo_mono_in := ler_nmono_in. +Definition lern_mono : + {homo f' : m n / m < n >-> (m < n)%N} -> + {mono f' : m n / m <= n >-> (m <= n)%N} := + le_mono (f := f'). +Definition lern_nmono : + {homo f' : m n / n < m >-> (m < n)%N} -> + {mono f' : m n / n <= m >-> (m <= n)%N} := + le_nmono (f := f'). +Definition lern_mono_in : + {in D &, {homo f' : m n / m < n >-> (m < n)%N}} -> + {in D &, {mono f' : m n / m <= n >-> (m <= n)%N}} := + le_mono_in (f := f'). +Definition lern_nmono_in : + {in D &, {homo f' : m n / n < m >-> (m < n)%N}} -> + {in D &, {mono f' : m n / n <= m >-> (m <= n)%N}} := + le_nmono_in (f := f'). +End RealDomainMonotony. + +Section RealDomainOperations. +Variable R : realDomainType. +Implicit Types x y z : R. +Section MinMax. +Definition minrC : @commutative R R min := @meetC _ R. +Definition minrr : @idempotent R min := @meetxx _ R. +Definition minr_l x y : x <= y -> min x y = x := @meet_l _ _ x y. +Definition minr_r x y : y <= x -> min x y = y := @meet_r _ _ x y. +Definition maxrC : @commutative R R max := @joinC _ R. +Definition maxrr : @idempotent R max := @joinxx _ R. +Definition maxr_l x y : y <= x -> max x y = x := @join_l _ _ x y. +Definition maxr_r x y : x <= y -> max x y = y := @join_r _ _ x y. +Definition minrA x y z : min x (min y z) = min (min x y) z := meetA x y z. +Definition minrCA : @left_commutative R R min := meetCA. +Definition minrAC : @right_commutative R R min := meetAC. +Definition maxrA x y z : max x (max y z) = max (max x y) z := joinA x y z. +Definition maxrCA : @left_commutative R R max := joinCA. +Definition maxrAC : @right_commutative R R max := joinAC. +Definition eqr_minl x y : (min x y == x) = (x <= y) := eq_meetl x y. +Definition eqr_minr x y : (min x y == y) = (y <= x) := eq_meetr x y. +Definition eqr_maxl x y : (max x y == x) = (y <= x) := eq_joinl x y. +Definition eqr_maxr x y : (max x y == y) = (x <= y) := eq_joinr x y. +Definition ler_minr x y z : (x <= min y z) = (x <= y) && (x <= z) := lexI x y z. +Definition ler_minl x y z : (min y z <= x) = (y <= x) || (z <= x) := leIx x y z. +Definition ler_maxr x y z : (x <= max y z) = (x <= y) || (x <= z) := lexU x y z. +Definition ler_maxl x y z : (max y z <= x) = (y <= x) && (z <= x) := leUx y z x. +Definition ltr_minr x y z : (x < min y z) = (x < y) && (x < z) := ltxI x y z. +Definition ltr_minl x y z : (min y z < x) = (y < x) || (z < x) := ltIx x y z. +Definition ltr_maxr x y z : (x < max y z) = (x < y) || (x < z) := ltxU x y z. +Definition ltr_maxl x y z : (max y z < x) = (y < x) && (z < x) := ltUx x y z. +Definition lter_minr := (ler_minr, ltr_minr). +Definition lter_minl := (ler_minl, ltr_minl). +Definition lter_maxr := (ler_maxr, ltr_maxr). +Definition lter_maxl := (ler_maxl, ltr_maxl). +Definition minrK x y : max (min x y) x = x := meetUKC y x. +Definition minKr x y : min y (max x y) = y := joinKIC x y. +Definition maxr_minl : @left_distributive R R max min := @joinIl _ R. +Definition maxr_minr : @right_distributive R R max min := @joinIr _ R. +Definition minr_maxl : @left_distributive R R min max := @meetUl _ R. +Definition minr_maxr : @right_distributive R R min max := @meetUr _ R. +Variant minr_spec x y : bool -> bool -> R -> Type := +| Minr_r of x <= y : minr_spec x y true false x +| Minr_l of y < x : minr_spec x y false true y. +Lemma minrP x y : minr_spec x y (x <= y) (y < x) (min x y). +Proof. by case: leP; constructor. Qed. +Variant maxr_spec x y : bool -> bool -> R -> Type := +| Maxr_r of y <= x : maxr_spec x y true false x +| Maxr_l of x < y : maxr_spec x y false true y. +Lemma maxrP x y : maxr_spec x y (y <= x) (x < y) (max x y). +Proof. by case: (leP y); constructor. Qed. +End MinMax. +End RealDomainOperations. + +Arguments lerifP {R x y C}. +Arguments lerif_refl {R x C}. +Arguments mono_in_lerif [R A f C]. +Arguments nmono_in_lerif [R A f C]. +Arguments mono_lerif [R f C]. +Arguments nmono_lerif [R f C]. + +Section RealDomainArgExtremum. + +Context {R : realDomainType} {I : finType} (i0 : I). +Context (P : pred I) (F : I -> R) (Pi0 : P i0). + +Definition arg_minr := extremum <=%R i0 P F. +Definition arg_maxr := extremum >=%R i0 P F. +Definition arg_minrP : extremum_spec <=%R P F arg_minr := arg_minP F Pi0. +Definition arg_maxrP : extremum_spec >=%R P F arg_maxr := arg_maxP F Pi0. + +End RealDomainArgExtremum. + +Notation "@ 'real_lerP'" := + (deprecate real_lerP real_leP) (at level 10, only parsing) : fun_scope. +Notation real_lerP := (@real_lerP _ _ _) (only parsing). +Notation "@ 'real_ltrP'" := + (deprecate real_ltrP real_ltP) (at level 10, only parsing) : fun_scope. +Notation real_ltrP := (@real_ltrP _ _ _) (only parsing). +Notation "@ 'real_ltrNge'" := + (deprecate real_ltrNge real_ltNge) (at level 10, only parsing) : fun_scope. +Notation real_ltrNge := (@real_ltrNge _ _ _) (only parsing). +Notation "@ 'real_lerNgt'" := + (deprecate real_lerNgt real_leNgt) (at level 10, only parsing) : fun_scope. +Notation real_lerNgt := (@real_lerNgt _ _ _) (only parsing). +Notation "@ 'real_ltrgtP'" := + (deprecate real_ltrgtP real_ltgtP) (at level 10, only parsing) : fun_scope. +Notation real_ltrgtP := (@real_ltrgtP _ _ _) (only parsing). +Notation "@ 'real_ger0P'" := + (deprecate real_ger0P real_ge0P) (at level 10, only parsing) : fun_scope. +Notation real_ger0P := (@real_ger0P _ _) (only parsing). +Notation "@ 'real_ltrgt0P'" := + (deprecate real_ltrgt0P real_ltgt0P) (at level 10, only parsing) : fun_scope. +Notation real_ltrgt0P := (@real_ltrgt0P _ _) (only parsing). +Notation lerif_nat := (deprecate lerif_nat leif_nat_r) (only parsing). +Notation "@ 'lerif_subLR'" := + (deprecate lerif_subLR leif_subLR) (at level 10, only parsing) : fun_scope. +Notation lerif_subLR := (@lerif_subLR _) (only parsing). +Notation "@ 'lerif_subRL'" := + (deprecate lerif_subRL leif_subRL) (at level 10, only parsing) : fun_scope. +Notation lerif_subRL := (@lerif_subRL _) (only parsing). +Notation "@ 'lerif_add'" := + (deprecate lerif_add leif_add) (at level 10, only parsing) : fun_scope. +Notation lerif_add := (@lerif_add _ _ _ _ _ _ _) (only parsing). +Notation "@ 'lerif_sum'" := + (deprecate lerif_sum leif_sum) (at level 10, only parsing) : fun_scope. +Notation lerif_sum := (@lerif_sum _ _ _ _ _ _) (only parsing). +Notation "@ 'lerif_0_sum'" := + (deprecate lerif_0_sum leif_0_sum) (at level 10, only parsing) : fun_scope. +Notation lerif_0_sum := (@lerif_0_sum _ _ _ _ _) (only parsing). +Notation "@ 'real_lerif_norm'" := + (deprecate real_lerif_norm real_leif_norm) + (at level 10, only parsing) : fun_scope. +Notation real_lerif_norm := (@real_lerif_norm _ _) (only parsing). +Notation "@ 'lerif_pmul'" := + (deprecate lerif_pmul leif_pmul) (at level 10, only parsing) : fun_scope. +Notation lerif_pmul := (@lerif_pmul _ _ _ _ _ _ _) (only parsing). +Notation "@ 'lerif_nmul'" := + (deprecate lerif_nmul leif_nmul) (at level 10, only parsing) : fun_scope. +Notation lerif_nmul := (@lerif_nmul _ _ _ _ _ _ _) (only parsing). +Notation "@ 'lerif_pprod'" := + (deprecate lerif_pprod leif_pprod) (at level 10, only parsing) : fun_scope. +Notation lerif_pprod := (@lerif_pprod _ _ _ _ _ _) (only parsing). +Notation "@ 'real_lerif_mean_square_scaled'" := + (deprecate real_lerif_mean_square_scaled real_leif_mean_square_scaled) + (at level 10, only parsing) : fun_scope. +Notation real_lerif_mean_square_scaled := + (@real_lerif_mean_square_scaled _ _ _ _ _ _) (only parsing). +Notation "@ 'real_lerif_AGM2_scaled'" := + (deprecate real_lerif_AGM2_scaled real_leif_AGM2_scaled) + (at level 10, only parsing) : fun_scope. +Notation real_lerif_AGM2_scaled := + (@real_lerif_AGM2_scaled _ _ _) (only parsing). +Notation "@ 'lerif_AGM_scaled'" := + (deprecate lerif_AGM_scaled leif_AGM2_scaled) + (at level 10, only parsing) : fun_scope. +Notation lerif_AGM_scaled := (@lerif_AGM_scaled _ _ _ _) (only parsing). +Notation "@ 'real_lerif_mean_square'" := + (deprecate real_lerif_mean_square real_leif_mean_square) + (at level 10, only parsing) : fun_scope. +Notation real_lerif_mean_square := + (@real_lerif_mean_square _ _ _) (only parsing). +Notation "@ 'real_lerif_AGM2'" := + (deprecate real_lerif_AGM2 real_leif_AGM2) + (at level 10, only parsing) : fun_scope. +Notation real_lerif_AGM2 := (@real_lerif_AGM2 _ _ _) (only parsing). +Notation "@ 'lerif_AGM'" := + (deprecate lerif_AGM leif_AGM) (at level 10, only parsing) : fun_scope. +Notation lerif_AGM := (@lerif_AGM _ _ _ _) (only parsing). +Notation "@ 'lerif_mean_square_scaled'" := + (deprecate lerif_mean_square_scaled leif_mean_square_scaled) + (at level 10, only parsing) : fun_scope. +Notation lerif_mean_square_scaled := + (@lerif_mean_square_scaled _) (only parsing). +Notation "@ 'lerif_AGM2_scaled'" := + (deprecate lerif_AGM2_scaled leif_AGM2_scaled) + (at level 10, only parsing) : fun_scope. +Notation lerif_AGM2_scaled := (@lerif_AGM2_scaled _) (only parsing). +Notation "@ 'lerif_mean_square'" := + (deprecate lerif_mean_square leif_mean_square) + (at level 10, only parsing) : fun_scope. +Notation lerif_mean_square := (@lerif_mean_square _) (only parsing). +Notation "@ 'lerif_AGM2'" := + (deprecate lerif_AGM2 leif_AGM2) (at level 10, only parsing) : fun_scope. +Notation lerif_AGM2 := (@lerif_AGM2 _) (only parsing). +Notation "@ 'lerif_normC_Re_Creal'" := + (deprecate lerif_normC_Re_Creal leif_normC_Re_Creal) + (at level 10, only parsing) : fun_scope. +Notation lerif_normC_Re_Creal := (@lerif_normC_Re_Creal _) (only parsing). +Notation "@ 'lerif_Re_Creal'" := + (deprecate lerif_Re_Creal leif_Re_Creal) + (at level 10, only parsing) : fun_scope. +Notation lerif_Re_Creal := (@lerif_Re_Creal _) (only parsing). +Notation "@ 'lerif_rootC_AGM'" := + (deprecate lerif_rootC_AGM leif_rootC_AGM) + (at level 10, only parsing) : fun_scope. +Notation lerif_rootC_AGM := (@lerif_rootC_AGM _ _ _ _) (only parsing). + +End Theory. + +Notation "[ 'arg' 'minr_' ( i < i0 | P ) F ]" := + (arg_minr i0 (fun i => P%B) (fun i => F)) + (at level 0, i, i0 at level 10, + format "[ 'arg' 'minr_' ( i < i0 | P ) F ]") : form_scope. + +Notation "[ 'arg' 'minr_' ( i < i0 'in' A ) F ]" := + [arg minr_(i < i0 | i \in A) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'minr_' ( i < i0 'in' A ) F ]") : form_scope. + +Notation "[ 'arg' 'minr_' ( i < i0 ) F ]" := [arg minr_(i < i0 | true) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'minr_' ( i < i0 ) F ]") : form_scope. + +Notation "[ 'arg' 'maxr_' ( i > i0 | P ) F ]" := + (arg_maxr i0 (fun i => P%B) (fun i => F)) + (at level 0, i, i0 at level 10, + format "[ 'arg' 'maxr_' ( i > i0 | P ) F ]") : form_scope. + +Notation "[ 'arg' 'maxr_' ( i > i0 'in' A ) F ]" := + [arg maxr_(i > i0 | i \in A) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'maxr_' ( i > i0 'in' A ) F ]") : form_scope. + +Notation "[ 'arg' 'maxr_' ( i > i0 ) F ]" := [arg maxr_(i > i0 | true) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'maxr_' ( i > i0 ) F ]") : form_scope. + +End Num. +End mc_1_10. diff --git a/mathcomp/character/character.v b/mathcomp/character/character.v index c95d03b..78c5295 100644 --- a/mathcomp/character/character.v +++ b/mathcomp/character/character.v @@ -1,9 +1,9 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun eqtype choice ssrnat seq. -From mathcomp Require Import path div fintype tuple finfun bigop prime ssralg. -From mathcomp Require Import poly finset gproduct fingroup morphism perm. -From mathcomp Require Import automorphism quotient finalg action zmodp. +From mathcomp Require Import path div fintype tuple finfun bigop prime order. +From mathcomp Require Import ssralg poly finset gproduct fingroup morphism. +From mathcomp Require Import perm automorphism quotient finalg action zmodp. From mathcomp Require Import commutator cyclic center pgroup nilpotent sylow. From mathcomp Require Import abelian matrix mxalgebra mxpoly mxrepresentation. From mathcomp Require Import vector ssrnum algC classfun. @@ -67,7 +67,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GroupScope GRing.Theory Num.Theory. +Import Order.TTheory GroupScope GRing.Theory Num.Theory. Local Open Scope ring_scope. Local Notation algCF := [fieldType of algC]. @@ -656,7 +656,7 @@ Lemma irr1_gt0 i : 0 < 'chi_i 1%g. Proof. by rewrite irr1_degree ltr0n irr_degree_gt0. Qed. Lemma irr1_neq0 i : 'chi_i 1%g != 0. -Proof. by rewrite eqr_le ltr_geF ?irr1_gt0. Qed. +Proof. by rewrite eq_le lt_geF ?irr1_gt0. Qed. Lemma irr_neq0 i : 'chi_i != 0. Proof. by apply: contraNneq (irr1_neq0 i) => ->; rewrite cfunE. Qed. @@ -880,7 +880,7 @@ Lemma char1_eq0 chi : chi \is a character -> (chi 1%g == 0) = (chi == 0). Proof. case/char_sum_irr=> r ->; apply/idP/idP=> [|/eqP->]; last by rewrite cfunE. case: r => [|i r]; rewrite ?big_nil // sum_cfunE big_cons. -rewrite paddr_eq0 ?sumr_ge0 => // [||j _]; rewrite 1?ltrW ?irr1_gt0 //. +rewrite paddr_eq0 ?sumr_ge0 => // [||j _]; rewrite 1?ltW ?irr1_gt0 //. by rewrite (negbTE (irr1_neq0 i)). Qed. @@ -1021,7 +1021,7 @@ Proof. by move=> Gx; rewrite -lin_charX // expg_order lin_char1. Qed. Lemma normC_lin_char x : x \in G -> `|xi x| = 1. Proof. -move=> Gx; apply/eqP; rewrite -(@pexpr_eq1 _ _ #[x]) ?normr_ge0 //. +move=> Gx; apply/eqP; rewrite -(@pexpr_eq1 _ _ #[x]) //. by rewrite -normrX // lin_char_unity_root ?normr1. Qed. @@ -1163,12 +1163,11 @@ have exp_e j: e 0 j ^+ #[x] = 1. rewrite expgS repr_mxM ?groupX // {1}rGx -!mulmxA mulKVmx //. by rewrite mul_diag_mx mulmxA [M in _ = M]mxE -IHn exprS {1}mxE eqxx. have norm1_e j: `|e 0 j| = 1. - apply/eqP; rewrite -(@pexpr_eq1 _ _ #[x]) ?normr_ge0 //. - by rewrite -normrX exp_e normr1. + by apply/eqP; rewrite -(@pexpr_eq1 _ _ #[x]) // -normrX exp_e normr1. exists e; split=> //; first by exists B. rewrite cfRepr1 !cfunE Gx rGx mxtrace_mulC mulKVmx // mxtrace_diag. - split=> //=; apply: (ler_trans (ler_norm_sum _ _ _)). - by rewrite (eq_bigr _ (in1W norm1_e)) sumr_const card_ord lerr. + split=> //=; apply: (le_trans (ler_norm_sum _ _ _)). + by rewrite (eq_bigr _ (in1W norm1_e)) sumr_const card_ord lexx. rewrite !cfunE groupV !mulrb Gx rGx mxtrace_mulC mulKVmx //. rewrite -trace_map_mx map_diag_mx; set d' := diag_mx _. rewrite -[d'](mulKVmx unitB) mxtrace_mulC -[_ *m _](repr_mxK rG Gx) rGx. @@ -1559,7 +1558,7 @@ Lemma constt_ortho_char (phi psi : 'CF(G)) i j : Proof. move=> _ _ /constt_charP[//|phi1 Nphi1 ->] /constt_charP[//|psi1 Npsi1 ->]. rewrite cfdot_irr; case: eqP => // -> /eqP/idPn[]. -rewrite cfdotDl !cfdotDr cfnorm_irr -addrA gtr_eqF ?ltr_paddr ?ltr01 //. +rewrite cfdotDl !cfdotDr cfnorm_irr -addrA gt_eqF ?ltr_paddr ?ltr01 //. by rewrite Cnat_ge0 ?rpredD ?Cnat_cfdot_char ?irr_char. Qed. @@ -1756,7 +1755,7 @@ Lemma constt_Res_trans j psi : psi \is a character -> j \in irr_constt psi -> {subset irr_constt ('Res[H, G] 'chi_j) <= irr_constt ('Res[H] psi)}. Proof. -move=> Npsi Cj i; apply: contraNneq; rewrite eqr_le => {1}<-. +move=> Npsi Cj i; apply: contraNneq; rewrite eq_le => {1}<-. rewrite cfdot_Res_ge_constt ?Cnat_ge0 ?Cnat_cfdot_char_irr //. by rewrite cfRes_char ?irr_char. Qed. @@ -2859,18 +2858,18 @@ by rewrite cfcenter_eq_center subsetIr. Qed. (* This is Isaacs (2.29). *) -Lemma cfnorm_Res_lerif H phi : +Lemma cfnorm_Res_leif H phi : H \subset G -> '['Res[H] phi] <= #|G : H|%:R * '[phi] ?= iff (phi \in 'CF(G, H)). Proof. move=> sHG; rewrite cfun_onE mulrCA natf_indexg // -mulrA mulKf ?neq0CG //. rewrite (big_setID H) (setIidPr sHG) /= addrC. -rewrite (mono_lerif (ler_pmul2l _)) ?invr_gt0 ?gt0CG // -lerif_subLR -sumrB. +rewrite (mono_leif (ler_pmul2l _)) ?invr_gt0 ?gt0CG // -leif_subLR -sumrB. rewrite big1 => [|x Hx]; last by rewrite !cfResE ?subrr. have ->: (support phi \subset H) = (G :\: H \subset [set x | phi x == 0]). rewrite subDset setUC -subDset; apply: eq_subset => x. by rewrite !inE (andb_idr (contraR _)) // => /cfun0->. -rewrite (sameP subsetP forall_inP); apply: lerif_0_sum => x _. +rewrite (sameP subsetP forall_inP); apply: leif_0_sum => x _. by rewrite !inE /<?=%R mul_conjC_ge0 eq_sym mul_conjC_eq0. Qed. @@ -2879,7 +2878,7 @@ Lemma irr1_bound (i : Iirr G) : ('chi_i 1%g) ^+ 2 <= #|G : 'Z('chi_i)%CF|%:R ?= iff ('chi_i \in 'CF(G, 'Z('chi_i)%CF)). Proof. -congr (_ <= _ ?= iff _): (cfnorm_Res_lerif 'chi_i (cfcenter_sub 'chi_i)). +congr (_ <= _ ?= iff _): (cfnorm_Res_leif 'chi_i (cfcenter_sub 'chi_i)). have [xi Lxi ->] := cfcenter_Res 'chi_i. have /irrP[j ->] := lin_char_irr Lxi; rewrite cfdotZl cfdotZr cfdot_irr eqxx. by rewrite mulr1 irr1_degree conjC_nat. diff --git a/mathcomp/character/classfun.v b/mathcomp/character/classfun.v index 3f461e3..295b77a 100644 --- a/mathcomp/character/classfun.v +++ b/mathcomp/character/classfun.v @@ -1,7 +1,7 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path. -From mathcomp Require Import div choice fintype tuple finfun bigop prime. +From mathcomp Require Import div choice fintype tuple finfun bigop prime order. From mathcomp Require Import ssralg poly finset fingroup morphism perm. From mathcomp Require Import automorphism quotient finalg action gproduct. From mathcomp Require Import zmodp commutator cyclic center pgroup sylow. @@ -91,7 +91,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GroupScope GRing.Theory Num.Theory. +Import Order.TTheory GroupScope GRing.Theory Num.Theory. Local Open Scope ring_scope. Delimit Scope cfun_scope with CF. @@ -900,7 +900,7 @@ by rewrite phi0 // => y _; apply: mul_conjC_ge0. Qed. Lemma cfnorm_gt0 phi : ('[phi] > 0) = (phi != 0). -Proof. by rewrite ltr_def cfnorm_ge0 cfnorm_eq0 andbT. Qed. +Proof. by rewrite lt_def cfnorm_ge0 cfnorm_eq0 andbT. Qed. Lemma sqrt_cfnorm_ge0 phi : 0 <= sqrtC '[phi]. Proof. by rewrite sqrtC_ge0 cfnorm_ge0. Qed. @@ -944,7 +944,7 @@ Lemma cfCauchySchwarz phi psi : Proof. rewrite free_cons span_seq1 seq1_free -negb_or negbK orbC. have [-> | nz_psi] /= := altP (psi =P 0). - by apply/lerifP; rewrite !cfdot0r normCK mul0r mulr0. + by apply/leifP; rewrite !cfdot0r normCK mul0r mulr0. without loss ophi: phi / '[phi, psi] = 0. move=> IHo; pose a := '[phi, psi] / '[psi]; pose phi1 := phi - a *: psi. have ophi: '[phi1, psi] = 0. @@ -952,7 +952,7 @@ without loss ophi: phi / '[phi, psi] = 0. rewrite (canRL (subrK _) (erefl phi1)) rpredDr ?rpredZ ?memv_line //. rewrite cfdotDl ophi add0r cfdotZl normrM (ger0_norm (cfnorm_ge0 _)). rewrite exprMn mulrA -cfnormZ cfnormDd; last by rewrite cfdotZr ophi mulr0. - by have:= IHo _ ophi; rewrite mulrDl -lerif_subLR subrr ophi normCK mul0r. + by have:= IHo _ ophi; rewrite mulrDl -leif_subLR subrr ophi normCK mul0r. rewrite ophi normCK mul0r; split; first by rewrite mulr_ge0 ?cfnorm_ge0. rewrite eq_sym mulf_eq0 orbC cfnorm_eq0 (negPf nz_psi) /=. apply/idP/idP=> [|/vlineP[a {2}->]]; last by rewrite cfdotZr ophi mulr0. @@ -963,20 +963,19 @@ Lemma cfCauchySchwarz_sqrt phi psi : `|'[phi, psi]| <= sqrtC '[phi] * sqrtC '[psi] ?= iff ~~ free (phi :: psi). Proof. rewrite -(sqrCK (normr_ge0 _)) -sqrtCM ?qualifE ?cfnorm_ge0 //. -rewrite (mono_in_lerif (@ler_sqrtC _)) 1?rpredM ?qualifE; -rewrite ?normr_ge0 ?cfnorm_ge0 //. +rewrite (mono_in_leif (@ler_sqrtC _)) 1?rpredM ?qualifE ?cfnorm_ge0 //. exact: cfCauchySchwarz. Qed. -Lemma cf_triangle_lerif phi psi : +Lemma cf_triangle_leif phi psi : sqrtC '[phi + psi] <= sqrtC '[phi] + sqrtC '[psi] ?= iff ~~ free (phi :: psi) && (0 <= coord [tuple psi] 0 phi). Proof. -rewrite -(mono_in_lerif ler_sqr) ?rpredD ?qualifE ?sqrtC_ge0 ?cfnorm_ge0 //. -rewrite andbC sqrrD !sqrtCK addrAC cfnormD (mono_lerif (ler_add2l _)). +rewrite -(mono_in_leif ler_sqr) ?rpredD ?qualifE ?sqrtC_ge0 ?cfnorm_ge0 //. +rewrite andbC sqrrD !sqrtCK addrAC cfnormD (mono_leif (ler_add2l _)). rewrite -mulr_natr -[_ + _](divfK (negbT (eqC_nat 2 0))) -/('Re _). -rewrite (mono_lerif (ler_pmul2r _)) ?ltr0n //. -have:= lerif_trans (lerif_Re_Creal '[phi, psi]) (cfCauchySchwarz_sqrt phi psi). +rewrite (mono_leif (ler_pmul2r _)) ?ltr0n //. +have:= leif_trans (leif_Re_Creal '[phi, psi]) (cfCauchySchwarz_sqrt phi psi). congr (_ <= _ ?= iff _); apply: andb_id2r. rewrite free_cons span_seq1 seq1_free -negb_or negbK orbC. have [-> | nz_psi] := altP (psi =P 0); first by rewrite cfdot0r coord0. @@ -2489,3 +2488,8 @@ Definition conj_cfQuo := cfAutQuo conjC. Definition conj_cfMod := cfAutMod conjC. Definition conj_cfInd := cfAutInd conjC. Definition cfconjC_eq1 := cfAut_eq1 conjC. + +Notation "@ 'cf_triangle_lerif'" := + (deprecate cf_triangle_lerif cf_triangle_leif) + (at level 10, only parsing). +Notation cf_triangle_lerif := (@cf_triangle_lerif _ _) (only parsing). diff --git a/mathcomp/character/inertia.v b/mathcomp/character/inertia.v index c644150..3809e73 100644 --- a/mathcomp/character/inertia.v +++ b/mathcomp/character/inertia.v @@ -1,7 +1,7 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path. -From mathcomp Require Import choice div fintype tuple finfun bigop prime. +From mathcomp Require Import choice fintype div tuple finfun bigop prime order. From mathcomp Require Import ssralg ssrnum finset fingroup morphism perm. From mathcomp Require Import automorphism quotient action zmodp cyclic center. From mathcomp Require Import gproduct commutator gseries nilpotent pgroup. @@ -12,7 +12,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GroupScope GRing.Theory Num.Theory. +Import Order.TTheory GroupScope GRing.Theory Num.Theory. Local Open Scope ring_scope. (******************************************************************************) @@ -938,7 +938,7 @@ have AtoB_P s (psi := 'chi_s) (chi := 'Ind[G] psi): s \in calA -> have ub_chi_r: 'chi_r 1%g <= chi 1%g ?= iff ('chi_r == chi). have Nchi: chi \is a character by rewrite cfInd_char ?irr_char. have [chi1 Nchi1->] := constt_charP _ Nchi sGr. - rewrite addrC cfunE -lerif_subLR subrr eq_sym -subr_eq0 addrK. + rewrite addrC cfunE -leif_subLR subrr eq_sym -subr_eq0 addrK. by split; rewrite ?char1_ge0 // eq_sym char1_eq0. have lb_chi_r: chi 1%g <= 'chi_r 1%g ?= iff (f == e). rewrite cfInd1 // -(cfRes1 H) DpsiH -(cfRes1 H 'chi_r) DrH !cfunE sum_cfunE. @@ -946,10 +946,10 @@ have AtoB_P s (psi := 'chi_s) (chi := 'Ind[G] psi): s \in calA -> by case/cfclassP=> y _ ->; rewrite cfConjg1. rewrite reindex_cfclass //= sumr_const -(eq_card (cfclass_IirrE _ _)). rewrite mulr_natl mulrnAr card_cfclass_Iirr //. - rewrite (mono_lerif (ler_pmuln2r (indexg_gt0 G T))). - rewrite (mono_lerif (ler_pmul2r (irr1_gt0 t))); apply: lerif_eq. + rewrite (mono_leif (ler_pmuln2r (indexg_gt0 G T))). + rewrite (mono_leif (ler_pmul2r (irr1_gt0 t))); apply: leif_eq. by rewrite /e -(cfResRes _ sHT) ?cfdot_Res_ge_constt. - have [_ /esym] := lerif_trans ub_chi_r lb_chi_r; rewrite eqxx. + have [_ /esym] := leif_trans ub_chi_r lb_chi_r; rewrite eqxx. by case/andP=> /eqP Dchi /eqP->; rewrite cfIirrE -/chi -?Dchi ?mem_irr. have part_c: {in calA, forall s (chi := 'Ind[G] 'chi_s), [predI irr_constt ('Res[T] chi) & calA] =i pred1 s}. @@ -957,14 +957,14 @@ have part_c: {in calA, forall s (chi := 'Ind[G] 'chi_s), have chiTs: s \in irr_constt ('Res[T] chi). by rewrite irr_consttE cfdot_Res_l irrWnorm ?oner_eq0. apply/andP/eqP=> [[/= chiTs1 As1] | -> //]. - apply: contraTeq Dchi_theta => s's1; rewrite ltr_eqF // -/chi. + apply: contraTeq Dchi_theta => s's1; rewrite lt_eqF // -/chi. have [|phi Nphi DchiT] := constt_charP _ _ chiTs. by rewrite cfRes_char ?cfInd_char ?irr_char. have [|phi1 Nphi1 Dphi] := constt_charP s1 Nphi _. rewrite irr_consttE -(canLR (addKr _) DchiT) addrC cfdotBl cfdot_irr. by rewrite mulrb ifN_eqC ?subr0. rewrite -(cfResRes chi sHT sTG) DchiT Dphi !rmorphD !cfdotDl /=. - rewrite -ltr_subl_addl subrr ltr_paddr ?ltr_def //; + rewrite -ltr_subl_addl subrr ltr_paddr ?lt_def //; rewrite Cnat_ge0 ?Cnat_cfdot_char ?cfRes_char ?irr_char //. by rewrite andbT -irr_consttE -constt_Ind_Res. do [split=> //; try by move=> s /AtoB_P[]] => [s1 s2 As1 As2 | r]. @@ -1039,8 +1039,8 @@ have [_]: '['Ind[G] phi] <= '['Ind[G] psi] ?= iff d_delta. rewrite DpsiG cfdot_suml; apply: eq_bigr => b _. rewrite -scalerAl cfdotZl cfdot_sumr; congr (_ * _). by apply: eq_bigr => g _; rewrite -scalerAl cfdotZr conj_Cnat. - have eMmono := mono_lerif (ler_pmul2l (egt0 _ _)). - apply: lerif_sum => b /eMmono->; apply: lerif_sum => g /eMmono->. + have eMmono := mono_leif (ler_pmul2l (egt0 _ _)). + apply: leif_sum => b /eMmono->; apply: leif_sum => g /eMmono->. split; last exact: eq_sym. have /CnatP[n Dd]: d b g \in Cnat by rewrite Cnat_cfdot_char. have [Db | _] := eqP; rewrite Dd leC_nat // -ltC_nat -Dd Db cfnorm_gt0. @@ -1059,7 +1059,7 @@ apply/idP/idP=> [|/imageP[b Sb ->]]. apply: contraR => N'i; rewrite big1 // => b Sb. rewrite cfdotZl cfdot_irr mulrb ifN_eqC ?mulr0 //. by apply: contraNneq N'i => ->; apply: image_f. -rewrite gtr_eqF // (bigD1 b) //= cfdotZl cfnorm_irr mulr1 ltr_paddr ?egt0 //. +rewrite gt_eqF // (bigD1 b) //= cfdotZl cfnorm_irr mulr1 ltr_paddr ?egt0 //. apply: sumr_ge0 => g /andP[Sg _]; rewrite cfdotZl cfdot_irr. by rewrite mulr_ge0 ?ler0n ?Cnat_ge0. Qed. @@ -1186,7 +1186,7 @@ have ltUK: U \proper K. rewrite Dmu subGcfker -irr_eq1 -Dmu cfMod_eq1 //. by rewrite (can2_eq (divrK Uj) (mulrK Uj)) mul1r (inj_eq irr_inj). suffices: theta \in 'CF(K, L). - rewrite -cfnorm_Res_lerif // DthL cfnormZ !cfnorm_irr !mulr1 normr_nat. + rewrite -cfnorm_Res_leif // DthL cfnormZ !cfnorm_irr !mulr1 normr_nat. by rewrite -natrX eqC_nat => /eqP. have <-: gcore U G = L. apply: maxL; last by rewrite sub_gcore ?cfker_mod. diff --git a/mathcomp/character/integral_char.v b/mathcomp/character/integral_char.v index 22bd171..1022afa 100644 --- a/mathcomp/character/integral_char.v +++ b/mathcomp/character/integral_char.v @@ -1,7 +1,7 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path. -From mathcomp Require Import div choice fintype tuple finfun bigop prime. +From mathcomp Require Import div choice fintype tuple finfun bigop prime order. From mathcomp Require Import ssralg poly finset fingroup morphism perm. From mathcomp Require Import automorphism quotient action finalg zmodp. From mathcomp Require Import commutator cyclic center pgroup sylow gseries. @@ -34,7 +34,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GroupScope GRing.Theory Num.Theory. +Import Order.TTheory GroupScope GRing.Theory Num.Theory. Local Open Scope ring_scope. Lemma group_num_field_exists (gT : finGroupType) (G : {group gT}) : @@ -274,7 +274,7 @@ have nz_m: m%:R != 0 :> algC by rewrite pnatr_eq0 -lt0n. pose alpha := 'chi_i g / m%:R. have a_lt1: `|alpha| < 1. rewrite normrM normfV normr_nat -{2}(divff nz_m). - rewrite ltr_def (can_eq (mulfVK nz_m)) eq_sym -{1}Dm -irr_cfcenterE // notZg. + rewrite lt_def (can_eq (mulfVK nz_m)) eq_sym -{1}Dm -irr_cfcenterE // notZg. by rewrite ler_pmul2r ?invr_gt0 ?ltr0n // -Dm char1_ge_norm ?irr_char. have Za: alpha \in Aint. have [u _ /dvdnP[v eq_uv]] := Bezoutl #|g ^: G| m_gt0. @@ -304,13 +304,13 @@ have Zbeta: beta \in Cint. have [|nz_a] := boolP (alpha == 0). by rewrite (can2_eq (divfK _) (mulfK _)) // mul0r => /eqP. have: beta != 0 by rewrite Dbeta; apply/prodf_neq0 => nu _; rewrite fmorph_eq0. -move/(norm_Cint_ge1 Zbeta); rewrite ltr_geF //; apply: ler_lt_trans a_lt1. +move/(norm_Cint_ge1 Zbeta); rewrite lt_geF //; apply: le_lt_trans a_lt1. rewrite -[`|alpha|]mulr1 Dbeta (bigD1 1%g) ?group1 //= -Da. -case: (gQnC _) => /= _ <-; rewrite gal_id normrM. -rewrite -subr_ge0 -mulrBr mulr_ge0 ?normr_ge0 // Da subr_ge0. -elim/big_rec: _ => [|nu c _]; first by rewrite normr1 lerr. -apply: ler_trans; rewrite -subr_ge0 -{1}[`|c|]mul1r normrM -mulrBl. -by rewrite mulr_ge0 ?normr_ge0 // subr_ge0 norm_a_nu. +case: (gQnC _) => /= _ <-. +rewrite gal_id normrM -subr_ge0 -mulrBr mulr_ge0 // Da subr_ge0. +elim/big_rec: _ => [|nu c _]; first by rewrite normr1 lexx. +apply: le_trans; rewrite -subr_ge0 -{1}[`|c|]mul1r normrM -mulrBl. +by rewrite mulr_ge0 // subr_ge0 norm_a_nu. Qed. End GringIrrMode. @@ -677,18 +677,18 @@ have{pi1 Zpi1} pi2_ge1: 1 <= pi2. by rewrite Cint_normK // sqr_Cint_ge1 //; apply/prodf_neq0. have Sgt0: (#|S| > 0)%N by rewrite (cardD1 g) [g \in S]Sg. rewrite -mulr_natr -ler_pdivl_mulr ?ltr0n //. -have n2chi_ge0 s: s \in S -> 0 <= `|chi s| ^+ 2 by rewrite exprn_ge0 ?normr_ge0. +have n2chi_ge0 s: s \in S -> 0 <= `|chi s| ^+ 2 by rewrite exprn_ge0. rewrite -(expr_ge1 Sgt0); last by rewrite divr_ge0 ?ler0n ?sumr_ge0. -by rewrite (ler_trans pi2_ge1) // lerif_AGM. +by rewrite (le_trans pi2_ge1) // leif_AGM. Qed. (* This is Burnside's vanishing theorem (Isaacs, Theorem (3.15)). *) Theorem nonlinear_irr_vanish gT (G : {group gT}) i : 'chi[G]_i 1%g > 1 -> exists2 x, x \in G & 'chi_i x = 0. Proof. -move=> chi1gt1; apply/exists_eq_inP; apply: contraFT (ltr_geF chi1gt1). -move/exists_inPn => -nz_chi. -rewrite -(norm_Cnat (Cnat_irr1 i)) -(@expr_le1 _ 2) ?normr_ge0 //. +move=> chi1gt1; apply/exists_eq_inP; apply: contraFT (lt_geF chi1gt1). +move=> /exists_inPn-nz_chi. +rewrite -(norm_Cnat (Cnat_irr1 i)) -(@expr_le1 _ 2)//. rewrite -(ler_add2r (#|G|%:R * '['chi_i])) {1}cfnorm_irr mulr1. rewrite (cfnormE (cfun_onG _)) mulVKf ?neq0CG // (big_setD1 1%g) //=. rewrite addrCA ler_add2l (cardsD1 1%g) group1 mulrS ler_add2l. diff --git a/mathcomp/character/vcharacter.v b/mathcomp/character/vcharacter.v index 72bacc3..faecc02 100644 --- a/mathcomp/character/vcharacter.v +++ b/mathcomp/character/vcharacter.v @@ -1,7 +1,7 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path. -From mathcomp Require Import div choice fintype tuple finfun bigop prime. +From mathcomp Require Import div choice fintype tuple finfun bigop prime order. From mathcomp Require Import ssralg poly finset fingroup morphism perm. From mathcomp Require Import automorphism quotient finalg action gproduct. From mathcomp Require Import zmodp commutator cyclic center pgroup sylow. @@ -40,7 +40,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GroupScope GRing.Theory Num.Theory. +Import Order.TTheory GroupScope GRing.Theory Num.Theory. Local Open Scope ring_scope. Section Basics. @@ -225,8 +225,8 @@ exists chi1; last exists (- nchi2); last by rewrite opprK. apply: rpred_sum => i zi_ge0; rewrite -tnth_nth rpredZ_Cnat ?irr_char //. by rewrite CnatEint Zz. rewrite -sumrN rpred_sum // => i zi_lt0; rewrite -scaleNr -tnth_nth. -rewrite rpredZ_Cnat ?irr_char // CnatEint rpredN Zz oppr_ge0 ltrW //. -by rewrite real_ltrNge ?Creal_Cint. +rewrite rpredZ_Cnat ?irr_char // CnatEint rpredN Zz oppr_ge0 ltW //. +by rewrite real_ltNge ?Creal_Cint. Qed. Lemma Aint_vchar phi x : phi \in 'Z[irr G] -> phi x \in Aint. @@ -501,8 +501,8 @@ have neq_ji: j != i. by rewrite signr_eq0. have neq_bc: b != c. apply: contraTneq phi1_0; rewrite def_phi def_chi def_xi => ->. - rewrite -scalerDr !cfunE mulf_eq0 signr_eq0 eqr_le ltr_geF //. - by rewrite ltr_paddl ?ltrW ?irr1_gt0. + rewrite -scalerDr !cfunE mulf_eq0 signr_eq0 eq_le lt_geF //. + by rewrite ltr_paddl ?ltW ?irr1_gt0. rewrite {}def_phi {}def_chi {}def_xi !scaler_sign. case: b c neq_bc => [|] [|] // _; last by exists i, j. by exists j, i; rewrite 1?eq_sym // addrC. @@ -691,8 +691,8 @@ have def_phi: {in H, phi =1 'chi_i}. have [j def_chi_j]: {j | 'chi_j = phi}. apply/sig_eqW; have [[] [j]] := vchar_norm1P Zphi n1phi; last first. by rewrite scale1r; exists j. - move/cfunP/(_ 1%g)/eqP; rewrite scaleN1r def_phi // cfunE -addr_eq0 eqr_le. - by rewrite ltr_geF // ltr_paddl ?ltrW ?irr1_gt0. + move/cfunP/(_ 1%g)/eqP; rewrite scaleN1r def_phi // cfunE -addr_eq0 eq_le. + by rewrite lt_geF // ltr_paddl ?ltW ?irr1_gt0. exists j; rewrite ?cfkerEirr def_chi_j //; apply/subsetP => x /setDP[Gx notHx]. rewrite inE cfunE def_phi // cfunE -/a cfun1E // Gx mulr1 cfIndE //. rewrite big1 ?mulr0 ?add0r // => y Gy; apply/theta0/(contra _ notHx) => Hxy. @@ -834,7 +834,7 @@ Proof. by rewrite inE. Qed. Lemma Cnat_dirr (phi : 'CF(G)) i : phi \in 'Z[irr G] -> i \in dirr_constt phi -> '[phi, dchi i] \in Cnat. Proof. -move=> PiZ; rewrite CnatEint dirr_consttE andbC => /ltrW -> /=. +move=> PiZ; rewrite CnatEint dirr_consttE andbC => /ltW -> /=. by case: i => b i; rewrite cfdotZr rmorph_sign rpredMsign Cint_cfdot_vchar_irr. Qed. @@ -846,15 +846,14 @@ Lemma dirr_constt_oppI (phi: 'CF(G)) : dirr_constt phi :&: dirr_constt (-phi) = set0. Proof. apply/setP=> i; rewrite inE !dirr_consttE cfdotNl inE. -apply/idP=> /andP [L1 L2]; have := ltr_paddl (ltrW L1) L2. -by rewrite subrr ltr_def eqxx. +apply/idP=> /andP [L1 L2]; have := ltr_paddl (ltW L1) L2. +by rewrite subrr lt_def eqxx. Qed. Lemma dirr_constt_oppl (phi: 'CF(G)) i : - i \in dirr_constt phi -> (ndirr i) \notin dirr_constt phi. + i \in dirr_constt phi -> (ndirr i) \notin dirr_constt phi. Proof. -rewrite !dirr_consttE dchi_ndirrE cfdotNr oppr_gt0. -by move/ltrW=> /ler_gtF ->. +by rewrite !dirr_consttE dchi_ndirrE cfdotNr oppr_gt0 => /ltW /le_gtF ->. Qed. Definition to_dirr (B : {set gT}) (phi : 'CF(B)) (i : Iirr B) : dIirr B := @@ -876,7 +875,7 @@ Lemma of_irrK (phi: 'CF(G)) : {in dirr_constt phi, cancel (@of_irr G) (to_dirr phi)}. Proof. case=> b i; rewrite dirr_consttE cfdotZr rmorph_sign /= /to_dirr mulr_sign. -by rewrite fun_if oppr_gt0; case: b => [|/ltrW/ler_gtF] ->. +by rewrite fun_if oppr_gt0; case: b => [|/ltW/le_gtF] ->. Qed. Lemma cfdot_todirrE (phi: 'CF(G)) i (phi_i := dchi (to_dirr phi i)) : @@ -913,7 +912,7 @@ Lemma dirr_small_norm (phi : 'CF(G)) n : Proof. move=> PiZ Pln; rewrite ltnNge -leC_nat => Nl4. suffices Fd i: i \in dirr_constt phi -> '[phi, dchi i] = 1. - split; last 2 [by apply/setP=> u; rewrite !inE cfdotNl oppr_gt0 ltr_asym]. + split; last 2 [by apply/setP=> u; rewrite !inE cfdotNl oppr_gt0 lt_asym]. apply/eqP; rewrite -eqC_nat -sumr_const -Pln (cnorm_dconstt PiZ). by apply/eqP/eq_bigr=> i Hi; rewrite Fd // expr1n. rewrite {1}[phi]cfun_sum_dconstt //. diff --git a/mathcomp/field/algC.v b/mathcomp/field/algC.v index e1fd4d1..0e3005d 100644 --- a/mathcomp/field/algC.v +++ b/mathcomp/field/algC.v @@ -1,9 +1,9 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq choice. -From mathcomp Require Import div fintype path bigop finset prime ssralg poly. -From mathcomp Require Import polydiv mxpoly generic_quotient countalg ssrnum. -From mathcomp Require Import closed_field ssrint rat intdiv. +From mathcomp Require Import div fintype path bigop finset prime order ssralg. +From mathcomp Require Import poly polydiv mxpoly generic_quotient countalg. +From mathcomp Require Import ssrnum closed_field ssrint rat intdiv. From mathcomp Require Import algebraics_fundamentals. (******************************************************************************) @@ -53,14 +53,14 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GRing.Theory Num.Theory. +Import Order.TTheory GRing.Theory Num.Theory. Local Open Scope ring_scope. (* The Num mixin for an algebraically closed field with an automorphism of *) (* order 2, making it into a field of complex numbers. *) Lemma ComplexNumMixin (L : closedFieldType) (conj : {rmorphism L -> L}) : involutive conj -> ~ conj =1 id -> - {numL | forall x : NumDomainType L numL, `|x| ^+ 2 = x * conj x}. + {numL : numMixin L | forall x : NumDomainType L numL, `|x| ^+ 2 = x * conj x}. Proof. move=> conjK conj_nt. have nz2: 2%:R != 0 :> L. @@ -184,7 +184,7 @@ have normD x y : le (norm (x + y)) (norm x + norm y). apply/posP; exists (i * (x * conj y - y * conj x)); congr (_ * _). rewrite !(rmorphM, rmorphB) iJ !conjK mulNr -mulrN opprB. by rewrite (mulrC x) (mulrC y). -by exists (Num.Mixin normD sposD norm_eq0 pos_linear normM (rrefl _) (rrefl _)). +by exists (NumMixin normD sposD norm_eq0 pos_linear normM (rrefl _) (rrefl _)). Qed. Module Algebraics. @@ -229,8 +229,10 @@ Canonical decFieldType := DecFieldType type decFieldMixin. Axiom closedFieldAxiom : GRing.ClosedField.axiom ringType. Canonical closedFieldType := ClosedFieldType type closedFieldAxiom. -Parameter numMixin : Num.mixin_of ringType. +Parameter numMixin : numMixin idomainType. +Canonical porderType := POrderType ring_display type numMixin. Canonical numDomainType := NumDomainType type numMixin. +Canonical normedZmodType := NormedZmodType type type numMixin. Canonical numFieldType := [numFieldType of type]. Parameter conjMixin : Num.ClosedField.imaginary_mixin_of numDomainType. @@ -422,13 +424,15 @@ have [i i2]: exists i : type, i ^+ 2 = -1. have [i] := @solve_monicpoly _ 2 (nth 0 [:: -1 : type]) isT. by rewrite !big_ord_recl big_ord0 /= mul0r mulr1 !addr0; exists i. move/(_ i)/(congr1 CtoL); rewrite LtoC_K => iL_J. -have/ltr_geF/idP[] := @ltr01 Lnum; rewrite -oppr_ge0 -(rmorphN1 CtoL_rmorphism). -rewrite -i2 rmorphX /= expr2 -{2}iL_J -(svalP LnumMixin). -by rewrite exprn_ge0 ?normr_ge0. +have/lt_geF/idP[] := @ltr01 Lnum; rewrite -oppr_ge0 -(rmorphN1 CtoL_rmorphism). +by rewrite -i2 rmorphX /= expr2 -{2}iL_J -(svalP LnumMixin) exprn_ge0. Qed. -Definition numMixin := sval (ComplexNumMixin conjK conj_nt). +Definition numMixin : numMixin closedFieldType := + sval (ComplexNumMixin conjK conj_nt). +Canonical porderType := POrderType ring_display type numMixin. Canonical numDomainType := NumDomainType type numMixin. +Canonical normedZmodType := NormedZmodType type type numMixin. Canonical numFieldType := [numFieldType of type]. Lemma normK u : `|u| ^+ 2 = u * conj u. @@ -480,18 +484,18 @@ Fact floorC_subproof x : {m | x \is Creal -> ZtoC m <= x < ZtoC (m + 1)}. Proof. have [Rx | _] := boolP (x \is Creal); last by exists 0. without loss x_ge0: x Rx / x >= 0. - have [x_ge0 | /ltrW x_le0] := real_ger0P Rx; first exact. + have [x_ge0 | /ltW x_le0] := real_ge0P Rx; first exact. case/(_ (- x)) => [||m /(_ isT)]; rewrite ?rpredN ?oppr_ge0 //. - rewrite ler_oppr ltr_oppl -!rmorphN opprD /= ltr_neqAle ler_eqVlt. + rewrite ler_oppr ltr_oppl -!rmorphN opprD /= lt_neqAle le_eqVlt. case: eqP => [-> _ | _ /and3P[lt_x_m _ le_m_x]]. - by exists (- m) => _; rewrite lerr rmorphD ltr_addl ltr01. + by exists (- m) => _; rewrite lexx rmorphD ltr_addl ltr01. by exists (- m - 1); rewrite le_m_x subrK. have /ex_minnP[n lt_x_n1 min_n]: exists n, x < n.+1%:R. have [n le_x_n] := rat_algebraic_archimedean algebraic x. - by exists n; rewrite -(ger0_norm x_ge0) (ltr_trans le_x_n) ?ltr_nat. + by exists n; rewrite -(ger0_norm x_ge0) (lt_trans le_x_n) ?ltr_nat. exists n%:Z => _; rewrite addrC -intS lt_x_n1 andbT. case Dn: n => // [n1]; rewrite -Dn. -have [||//|] := @real_lerP _ n%:R x; rewrite ?rpred_nat //. +have [||//|] := @real_leP _ n%:R x; rewrite ?rpred_nat //. by rewrite Dn => /min_n; rewrite Dn ltnn. Qed. @@ -535,7 +539,9 @@ Canonical unitRingType. Canonical comRingType. Canonical comUnitRingType. Canonical idomainType. +Canonical porderType. Canonical numDomainType. +Canonical normedZmodType. Canonical fieldType. Canonical numFieldType. Canonical decFieldType. @@ -663,15 +669,15 @@ Proof. by rewrite /floorC => Rx; case: (floorC_subproof x) => //= m; apply. Qed. Lemma floorC_def x m : m%:~R <= x < (m + 1)%:~R -> floorC x = m. Proof. -case/andP=> lemx ltxm1; apply/eqP; rewrite eqr_le -!ltz_addr1. +case/andP=> lemx ltxm1; apply/eqP; rewrite eq_le -!ltz_addr1. have /floorC_itv/andP[lefx ltxf1]: x \is Creal. by rewrite -[x](subrK m%:~R) rpredD ?realz ?ler_sub_real. -by rewrite -!(ltr_int [numFieldType of algC]) 2?(@ler_lt_trans _ x). +by rewrite -!(ltr_int [numFieldType of algC]) 2?(@le_lt_trans _ _ x). Qed. Lemma intCK : cancel intr floorC. Proof. -by move=> m; apply: floorC_def; rewrite ler_int ltr_int ltz_addr1 lerr. +by move=> m; apply: floorC_def; rewrite ler_int ltr_int ltz_addr1 lexx. Qed. Lemma floorCK : {in Cint, cancel floorC intr}. Proof. by move=> z /eqP. Qed. @@ -757,17 +763,17 @@ Lemma truncC_itv x : 0 <= x -> (truncC x)%:R <= x < (truncC x).+1%:R. Proof. move=> x_ge0; have /andP[lemx ltxm1] := floorC_itv (ger0_real x_ge0). rewrite /truncC x_ge0 -addn1 !pmulrn PoszD gez0_abs ?lemx //. -by rewrite -ltz_addr1 -(ltr_int [numFieldType of algC]) (ler_lt_trans x_ge0). +by rewrite -ltz_addr1 -(ltr_int [numFieldType of algC]) (le_lt_trans x_ge0). Qed. Lemma truncC_def x n : n%:R <= x < n.+1%:R -> truncC x = n. Proof. move=> ivt_n_x; have /andP[lenx _] := ivt_n_x. -by rewrite /truncC (ler_trans (ler0n _ n)) // (@floorC_def _ n) // addrC -intS. +by rewrite /truncC (le_trans (ler0n _ n)) // (@floorC_def _ n) // addrC -intS. Qed. Lemma natCK n : truncC n%:R = n. -Proof. by apply: truncC_def; rewrite lerr ltr_nat /=. Qed. +Proof. by apply: truncC_def; rewrite lexx ltr_nat /=. Qed. Lemma CnatP x : reflect (exists n, x = n%:R) (x \in Cnat). Proof. @@ -782,9 +788,9 @@ Proof. apply/idP/idP=> [m_gt0 | x_ge1]. have /truncC_itv/andP[lemx _]: 0 <= x. by move: m_gt0; rewrite /truncC; case: ifP. - by apply: ler_trans lemx; rewrite ler1n. -have /truncC_itv/andP[_ ltxm1]:= ler_trans ler01 x_ge1. -by rewrite -ltnS -ltC_nat (ler_lt_trans x_ge1). + by apply: le_trans lemx; rewrite ler1n. +have /truncC_itv/andP[_ ltxm1]:= le_trans ler01 x_ge1. +by rewrite -ltnS -ltC_nat (le_lt_trans x_ge1). Qed. Lemma truncC0Pn x : reflect (truncC x = 0%N) (~~ (1 <= x)). @@ -912,14 +918,12 @@ by rewrite pnatr_eq0 ler1n lt0n. Qed. Lemma sqr_Cint_ge1 x : x \in Cint -> x != 0 -> 1 <= x ^+ 2. -Proof. -by move=> Zx nz_x; rewrite -Cint_normK // expr_ge1 ?normr_ge0 ?norm_Cint_ge1. -Qed. +Proof. by move=> Zx nz_x; rewrite -Cint_normK // expr_ge1 ?norm_Cint_ge1. Qed. Lemma Cint_ler_sqr x : x \in Cint -> x <= x ^+ 2. Proof. move=> Zx; have [-> | nz_x] := eqVneq x 0; first by rewrite expr0n. -apply: ler_trans (_ : `|x| <= _); first by rewrite real_ler_norm ?Creal_Cint. +apply: le_trans (_ : `|x| <= _); first by rewrite real_ler_norm ?Creal_Cint. by rewrite -Cint_normK // ler_eexpr // norm_Cint_ge1. Qed. @@ -937,7 +941,7 @@ Lemma dvdCP_nat x y : 0 <= x -> 0 <= y -> (x %| y)%C -> {n | y = n%:R * x}. Proof. move=> x_ge0 y_ge0 x_dv_y; apply: sig_eqW. case/dvdCP: x_dv_y => z Zz -> in y_ge0 *; move: x_ge0 y_ge0 Zz. -rewrite ler_eqVlt => /predU1P[<- | ]; first by exists 22; rewrite !mulr0. +rewrite le_eqVlt => /predU1P[<- | ]; first by exists 22; rewrite !mulr0. by move=> /pmulr_lge0-> /CintEge0-> /CnatP[n ->]; exists n. Qed. diff --git a/mathcomp/field/algebraics_fundamentals.v b/mathcomp/field/algebraics_fundamentals.v index acafd8f..4824697 100644 --- a/mathcomp/field/algebraics_fundamentals.v +++ b/mathcomp/field/algebraics_fundamentals.v @@ -1,11 +1,11 @@ (* (c) Copyright 2006-2016 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) From mathcomp Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq choice. -From mathcomp Require Import div fintype path tuple bigop finset prime ssralg. -From mathcomp Require Import poly polydiv mxpoly countalg closed_field ssrnum. -From mathcomp Require Import ssrint rat intdiv fingroup finalg zmodp cyclic. -From mathcomp Require Import pgroup sylow vector falgebra fieldext separable. -From mathcomp Require Import galois. +From mathcomp Require Import div fintype path tuple bigop finset prime order. +From mathcomp Require Import ssralg poly polydiv mxpoly countalg closed_field. +From mathcomp Require Import ssrnum ssrint rat intdiv fingroup finalg zmodp. +From mathcomp Require Import cyclic pgroup sylow vector falgebra fieldext. +From mathcomp Require Import separable galois. (******************************************************************************) (* The main result in this file is the existence theorem that underpins the *) @@ -112,7 +112,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -Import GroupScope GRing.Theory Num.Theory. +Import Order.TTheory GroupScope GRing.Theory Num.Theory. Local Open Scope ring_scope. Local Notation "p ^ f" := (map_poly f p) : ring_scope. @@ -124,25 +124,25 @@ Lemma rat_algebraic_archimedean (C : numFieldType) (QtoC : Qmorphism C) : integralRange QtoC -> Num.archimedean_axiom C. Proof. move=> algC x. -without loss x_ge0: x / 0 <= x by rewrite -normr_id; apply; apply: normr_ge0. +without loss x_ge0: x / 0 <= x by rewrite -normr_id; apply. have [-> | nz_x] := eqVneq x 0; first by exists 1%N; rewrite normr0. have [p mon_p px0] := algC x; exists (\sum_(j < size p) `|numq p`_j|)%N. -rewrite ger0_norm // real_ltrNge ?rpred_nat ?ger0_real //. -apply: contraL px0 => lb_x; rewrite rootE gtr_eqF // horner_coef size_map_poly. -have x_gt0 k: 0 < x ^+ k by rewrite exprn_gt0 // ltr_def nz_x. +rewrite ger0_norm // real_ltNge ?rpred_nat ?ger0_real //. +apply: contraL px0 => lb_x; rewrite rootE gt_eqF // horner_coef size_map_poly. +have x_gt0 k: 0 < x ^+ k by rewrite exprn_gt0 // lt_def nz_x. move: lb_x; rewrite polySpred ?monic_neq0 // !big_ord_recr coef_map /=. rewrite -lead_coefE (monicP mon_p) natrD rmorph1 mul1r => lb_x. case: _.-1 (lb_x) => [|n]; first by rewrite !big_ord0 !add0r ltr01. rewrite -ltr_subl_addl add0r -(ler_pmul2r (x_gt0 n)) -exprS. -apply: ltr_le_trans; rewrite mulrDl mul1r ltr_spaddr // -sumrN. +apply: lt_le_trans; rewrite mulrDl mul1r ltr_spaddr // -sumrN. rewrite natr_sum mulr_suml ler_sum // => j _. -rewrite coef_map /= fmorph_eq_rat (ler_trans (real_ler_norm _)) //. +rewrite coef_map /= fmorph_eq_rat (le_trans (real_ler_norm _)) //. by rewrite rpredN rpredM ?rpred_rat ?rpredX // ger0_real. -rewrite normrN normrM ler_pmul //=. +rewrite normrN normrM ler_pmul //. rewrite normf_div -!intr_norm -!abszE ler_pimulr ?ler0n //. by rewrite invf_le1 ?ler1n ?ltr0n ?absz_gt0 ?denq_eq0. rewrite normrX ger0_norm ?(ltrW x_gt0) // ler_weexpn2l ?leq_ord //. -by rewrite (ler_trans _ lb_x) // -natrD addn1 ler1n. +by rewrite (le_trans _ lb_x) // -natrD addn1 ler1n. Qed. Definition decidable_embedding sT T (f : sT -> T) := @@ -161,8 +161,8 @@ have [n ub_n]: {n | forall y, root q y -> `|y| < n}. have /monic_Cauchy_bound[n2 ub_n2]: (-1) ^+ d *: (q \Po - 'X) \is monic. rewrite monicE lead_coefZ lead_coef_comp ?size_opp ?size_polyX // -/d. by rewrite lead_coef_opp lead_coefX (monicP mon_q) (mulrC 1) signrMK. - exists (Num.max n1 n2) => y; rewrite ltrNge ler_normr !ler_maxl rootE. - apply: contraL => /orP[]/andP[] => [/ub_n1/gtr_eqF->// | _ /ub_n2/gtr_eqF]. + exists (Num.max n1 n2) => y; rewrite ltNge ler_normr !leUx rootE. + apply: contraL => /orP[]/andP[] => [/ub_n1/gt_eqF->// | _ /ub_n2/gt_eqF]. by rewrite hornerZ horner_comp !hornerE opprK mulf_eq0 signr_eq0 => /= ->. have [p [a nz_a Dq]] := rat_poly_scale q; pose N := Num.bound `|n * a%:~R|. pose xa : seq rat := [seq (m%:R - N%:R) / a%:~R | m <- iota 0 N.*2]. @@ -193,13 +193,13 @@ have Dm: m%:R = `|y * a%:~R + N%:R|. by rewrite pmulrn abszE intr_norm Da rmorphD !rmorphM /= numqE mulrAC mulrA. have ltr_Qnat n1 n2 : (n1%:R < n2%:R :> rat = _) := ltr_nat _ n1 n2. have ub_y: `|y * a%:~R| < N%:R. - apply: ler_lt_trans (archi_boundP (normr_ge0 _)); rewrite !normrM. - by rewrite ler_pmul ?normr_ge0 // (ler_trans _ (ler_norm n)) ?ltrW ?ub_n. + apply: le_lt_trans (archi_boundP (normr_ge0 _)); rewrite !normrM. + by rewrite ler_pmul // (le_trans _ (ler_norm n)) ?ltW ?ub_n. apply/mapP; exists m. rewrite mem_iota /= add0n -addnn -ltr_Qnat Dm natrD. - by rewrite (ler_lt_trans (ler_norm_add _ _)) // normr_nat ltr_add2r. + by rewrite (le_lt_trans (ler_norm_add _ _)) // normr_nat ltr_add2r. rewrite Dm ger0_norm ?addrK ?mulfK ?intr_eq0 // -ler_subl_addl sub0r. -by rewrite (ler_trans (ler_norm _)) ?normrN ?ltrW. +by rewrite (le_trans (ler_norm _)) ?normrN ?ltW. Qed. Lemma minPoly_decidable_closure @@ -397,7 +397,7 @@ have add_Rroot xR p c: {yR | extendsR xR yR & has_Rroot xR p c -> root_in yR p}. elim: d => // d IHd in p mon_p s_p p0_le0 *; rewrite ltnS => le_p_d. have /closed_rootP/sig_eqW[y py0]: size (p ^ ofQ x) != 1%N. rewrite size_map_poly size_poly_eq1 eqp_monic ?rpred1 //. - by apply: contraTneq p0_le0 => ->; rewrite rmorph1 hornerC ltr_geF ?ltr01. + by apply: contraTneq p0_le0 => ->; rewrite rmorph1 hornerC lt_geF ?ltr01. have /s_p s_y := py0; have /z_s/sQ_inQ[u Dy] := s_y. have /pQwx[q Dq] := minPolyOver Qx u. have mon_q: q \is monic by have:= monic_minPoly Qx u; rewrite Dq map_monic. @@ -441,10 +441,10 @@ have add_Rroot xR p c: {yR | extendsR xR yR & has_Rroot xR p c -> root_in yR p}. case: (find k q _) => c d [[/= qc_le0 qd_ge0 le_cd] [/= le_ac le_db] Dcd]. have [/= le_ce le_ed] := midf_le le_cd; set e := _ / _ in le_ce le_ed. rewrite expnSr natrM invfM mulrA -{}Dcd /narrow /= -[mid _]/e. - have [qe_ge0 // | /ltrW qe_le0] := lerP 0 q.[e]. - do ?split=> //=; [exact: (ler_trans le_ed) | apply: canRL (mulfK nz2) _]. + have [qe_ge0 // | /ltW qe_le0] := lerP 0 q.[e]. + do ?split=> //=; [exact: (le_trans le_ed) | apply: canRL (mulfK nz2) _]. by rewrite mulrBl divfK // mulr_natr opprD addrACA subrr add0r. - do ?split=> //=; [exact: (ler_trans le_ac) | apply: canRL (mulfK nz2) _]. + do ?split=> //=; [exact: (le_trans le_ac) | apply: canRL (mulfK nz2) _]. by rewrite mulrBl divfK // mulr_natr opprD addrACA subrr addr0. have find_root r q ab: xup q ab -> {n | forall x, x \in itv (find n q ab) ->`|(r * q).[x]| < h2}. @@ -457,39 +457,39 @@ have add_Rroot xR p c: {yR | extendsR xR yR & has_Rroot xR p c -> root_in yR p}. have{xab} [[]] := findP n _ _ xab; case: (find n q ab) => a1 b1 /=. rewrite -/d => qa1_le0 qb1_ge0 le_ab1 [/= le_aa1 le_b1b] Dab1 le_a1c le_cb1. have /MuP lbMu: c \in itv ab. - by rewrite !inE (ler_trans le_aa1) ?(ler_trans le_cb1). - have Mu_ge0: 0 <= Mu by rewrite (ler_trans _ lbMu) ?normr_ge0. + by rewrite !inE (le_trans le_aa1) ?(le_trans le_cb1). + have Mu_ge0: 0 <= Mu by rewrite (le_trans _ lbMu). have Mdq_ge0: 0 <= Mdq. - by rewrite (ler_trans _ (MdqP 0 _)) ?normr_ge0 ?normr0. + by rewrite (le_trans _ (MdqP 0 _)) ?normr0. suffices lb1 a2 b2 (ab1 := (a1, b1)) (ab2 := (a2, b2)) : xup q ab2 /\ sub_itv ab2 ab1 -> q.[b2] - q.[a2] <= Mdq * wid ab1. - + apply: ler_lt_trans (_ : Mu * Mdq * wid (a1, b1) < h2); last first. + + apply: le_lt_trans (_ : Mu * Mdq * wid (a1, b1) < h2); last first. rewrite {}Dab1 mulrA ltr_pdivr_mulr ?ltr0n ?expn_gt0 //. - rewrite (ltr_le_trans (archi_boundP _)) ?mulr_ge0 ?ltr_nat // -/n. + rewrite (lt_le_trans (archi_boundP _)) ?mulr_ge0 ?ltr_nat // -/n. rewrite ler_pdivl_mull ?ltr0n // -natrM ler_nat. by case: n => // n; rewrite expnS leq_pmul2l // ltn_expl. - rewrite -mulrA hornerM normrM ler_pmul ?normr_ge0 //. - have [/ltrW qc_le0 | qc_ge0] := ltrP q.[c] 0. - by apply: ler_trans (lb1 c b1 _); rewrite ?ler0_norm ?ler_paddl. - by apply: ler_trans (lb1 a1 c _); rewrite ?ger0_norm ?ler_paddr ?oppr_ge0. + rewrite -mulrA hornerM normrM ler_pmul //. + have [/ltW qc_le0 | qc_ge0] := ltrP q.[c] 0. + by apply: le_trans (lb1 c b1 _); rewrite ?ler0_norm ?ler_paddl. + by apply: le_trans (lb1 a1 c _); rewrite ?ger0_norm ?ler_paddr ?oppr_ge0. case{c le_a1c le_cb1 lbMu}=> [[/=qa2_le0 qb2_ge0 le_ab2] [/=le_a12 le_b21]]. pose h := b2 - a2; have h_ge0: 0 <= h by rewrite subr_ge0. have [-> | nz_q] := eqVneq q 0. by rewrite !horner0 subrr mulr_ge0 ?subr_ge0. rewrite -(subrK a2 b2) (addrC h) (nderiv_taylor q (mulrC a2 h)). rewrite (polySpred nz_q) big_ord_recl /= mulr1 nderivn0 addrC addKr. - have [le_aa2 le_b2b] := (ler_trans le_aa1 le_a12, ler_trans le_b21 le_b1b). - have /MqP MqPx1: a2 \in itv ab by rewrite inE le_aa2 (ler_trans le_ab2). - apply: ler_trans (ler_trans (ler_norm _) (ler_norm_sum _ _ _)) _. - apply: ler_trans (_ : `|dq.[h] * h| <= _); last first. + have [le_aa2 le_b2b] := (le_trans le_aa1 le_a12, le_trans le_b21 le_b1b). + have /MqP MqPx1: a2 \in itv ab by rewrite inE le_aa2 (le_trans le_ab2). + apply: le_trans (le_trans (ler_norm _) (ler_norm_sum _ _ _)) _. + apply: le_trans (_ : `|dq.[h] * h| <= _); last first. by rewrite normrM ler_pmul ?normr_ge0 ?MdqP // ?ger0_norm ?ler_sub ?h_ge0. rewrite horner_poly ger0_norm ?mulr_ge0 ?sumr_ge0 // => [|j _]; last first. - by rewrite mulr_ge0 ?exprn_ge0 // (ler_trans _ (MqPx1 _)) ?normr_ge0. + by rewrite mulr_ge0 ?exprn_ge0 // (le_trans _ (MqPx1 _)). rewrite mulr_suml ler_sum // => j _; rewrite normrM -mulrA -exprSr. - by rewrite ler_pmul ?normr_ge0 // normrX ger0_norm. + by rewrite ler_pmul // normrX ger0_norm. have [ab0 xab0]: {ab | xup (p ^ QxR) ab}. have /monic_Cauchy_bound[b pb_gt0]: p ^ QxR \is monic by apply: monic_map. - by exists (0, `|b|); rewrite /xup normr_ge0 p0_le0 ltrW ?pb_gt0 ?ler_norm. + by exists (0, `|b|); rewrite /xup normr_ge0 p0_le0 ltW ?pb_gt0 ?ler_norm. pose ab_ n := find n (p ^ QxR) ab0; pose Iab_ n := itv (ab_ n). pose lim v a := (q_ v ^ QxR).[a]; pose nlim v n := lim v (ab_ n).2. have lim0 a: lim 0 a = 0. @@ -505,57 +505,57 @@ have add_Rroot xR p c: {yR | extendsR xR yR & has_Rroot xR p c -> root_in yR p}. have /(find_root r.1)[n ub_rp] := xab0; exists n. have [M Mgt0 ubM]: {M | 0 < M & {in Iab_ n, forall a, `|r.2.[a]| <= M}}. have [M ubM] := poly_itv_bound r.2 (ab_ n).1 (ab_ n).2. - exists (Num.max 1 M) => [|s /ubM vM]; first by rewrite ltr_maxr ltr01. - by rewrite ler_maxr orbC vM. + exists (Num.max 1 M) => [|s /ubM vM]; first by rewrite ltxU ltr01. + by rewrite lexU orbC vM. exists (h2 / M) => [|a xn_a]; first by rewrite divr_gt0 ?invr_gt0 ?ltr0n. rewrite ltr_pdivr_mulr // -(ltr_add2l h2) -mulr2n -mulr_natl divff //. rewrite -normr1 -(hornerC 1 a) -[1%:P]r_pq_1 hornerD. - rewrite ?(ler_lt_trans (ler_norm_add _ _)) ?ltr_le_add ?ub_rp //. + rewrite ?(le_lt_trans (ler_norm_add _ _)) ?ltr_le_add ?ub_rp //. by rewrite mulrC hornerM normrM ler_wpmul2l ?ubM. have ab_le m n: (m <= n)%N -> (ab_ n).2 \in Iab_ m. move/subnKC=> <-; move: {n}(n - m)%N => n; rewrite /ab_. have /(findP m)[/(findP n)[[_ _]]] := xab0. rewrite /find -iter_add -!/(find _ _) -!/(ab_ _) addnC !inE. - by move: (ab_ _) => /= ab_mn le_ab_mn [/ler_trans->]. + by move: (ab_ _) => /= ab_mn le_ab_mn [/le_trans->]. pose lt v w := 0 < nlim (w - v) (n_ (w - v)). have posN v: lt 0 (- v) = lt v 0 by rewrite /lt subr0 add0r. have posB v w: lt 0 (w - v) = lt v w by rewrite /lt subr0. have posE n v: (n_ v <= n)%N -> lt 0 v = (0 < nlim v n). rewrite /lt subr0 /nlim => /ab_le; set a := _.2; set b := _.2 => Iv_a. - have [-> | /nzP[e e_gt0]] := eqVneq v 0; first by rewrite !lim0 ltrr. + have [-> | /nzP[e e_gt0]] := eqVneq v 0; first by rewrite !lim0 ltxx. move: (n_ v) => m in Iv_a b * => v_gte. without loss lt0v: v v_gte / 0 < lim v b. - move=> IHv; apply/idP/idP => [v_gt0 | /ltrW]; first by rewrite -IHv. - rewrite ltr_def -normr_gt0 ?(ltr_trans _ (v_gte _ _)) ?ab_le //=. - rewrite !lerNgt -!oppr_gt0 -!limN; apply: contra => v_lt0. + move=> IHv; apply/idP/idP => [v_gt0 | /ltW]; first by rewrite -IHv. + rewrite lt_def -normr_gt0 ?(lt_trans _ (v_gte _ _)) ?ab_le //=. + rewrite !leNgt -!oppr_gt0 -!limN; apply: contra => v_lt0. by rewrite -IHv // => c /v_gte; rewrite limN normrN. - rewrite lt0v (ltr_trans e_gt0) ?(ltr_le_trans (v_gte a Iv_a)) //. - rewrite ger0_norm // lerNgt; apply/negP=> /ltrW lev0. + rewrite lt0v (lt_trans e_gt0) ?(lt_le_trans (v_gte a Iv_a)) //. + rewrite ger0_norm // leNgt; apply/negP=> /ltW lev0. have [le_a le_ab] : _ /\ a <= b := andP Iv_a. - have xab: xup (q_ v ^ QxR) (a, b) by move/ltrW in lt0v. + have xab: xup (q_ v ^ QxR) (a, b) by move/ltW in lt0v. have /(find_root (h2 / e)%:P)[n1] := xab; have /(findP n1)[[_ _]] := xab. case: (find _ _ _) => c d /= le_cd [/= le_ac le_db] _ /(_ c)/implyP. - rewrite inE lerr le_cd hornerM hornerC normrM ler_gtF //. - rewrite ger0_norm ?divr_ge0 ?invr_ge0 ?ler0n ?(ltrW e_gt0) // mulrAC. - rewrite ler_pdivl_mulr // ler_wpmul2l ?invr_ge0 ?ler0n // ltrW // v_gte //=. - by rewrite inE -/b (ler_trans le_a) //= (ler_trans le_cd). + rewrite inE lexx le_cd hornerM hornerC normrM le_gtF //. + rewrite ger0_norm ?divr_ge0 ?invr_ge0 ?ler0n ?(ltW e_gt0) // mulrAC. + rewrite ler_pdivl_mulr // ler_wpmul2l ?invr_ge0 ?ler0n // ltW // v_gte //=. + by rewrite inE -/b (le_trans le_a) //= (le_trans le_cd). pose lim_pos m v := exists2 e, e > 0 & forall n, (m <= n)%N -> e < nlim v n. have posP v: reflect (exists m, lim_pos m v) (lt 0 v). apply: (iffP idP) => [v_gt0|[m [e e_gt0 v_gte]]]; last first. - by rewrite (posE _ _ (leq_maxl _ m)) (ltr_trans e_gt0) ?v_gte ?leq_maxr. + by rewrite (posE _ _ (leq_maxl _ m)) (lt_trans e_gt0) ?v_gte ?leq_maxr. have [|e e_gt0 v_gte] := nzP v. - by apply: contraTneq v_gt0 => ->; rewrite /lt subr0 /nlim lim0 ltrr. + by apply: contraTneq v_gt0 => ->; rewrite /lt subr0 /nlim lim0 ltxx. exists (n_ v), e => // n le_vn; rewrite (posE n) // in v_gt0. - by rewrite -(ger0_norm (ltrW v_gt0)) v_gte ?ab_le. + by rewrite -(ger0_norm (ltW v_gt0)) v_gte ?ab_le. have posNneg v: lt 0 v -> ~~ lt v 0. case/posP=> m [d d_gt0 v_gtd]; rewrite -posN. apply: contraL d_gt0 => /posP[n [e e_gt0 nv_gte]]. - rewrite ltr_gtF // (ltr_trans (v_gtd _ (leq_maxl m n))) // -oppr_gt0. - by rewrite /nlim -limN (ltr_trans e_gt0) ?nv_gte ?leq_maxr. + rewrite lt_gtF // (lt_trans (v_gtd _ (leq_maxl m n))) // -oppr_gt0. + by rewrite /nlim -limN (lt_trans e_gt0) ?nv_gte ?leq_maxr. have posVneg v: v != 0 -> lt 0 v || lt v 0. case/nzP=> e e_gt0 v_gte; rewrite -posN; set w := - v. have [m [le_vm le_wm _]] := maxn3 (n_ v) (n_ w) 0%N; rewrite !(posE m) //. - by rewrite /nlim limN -ltr_normr (ltr_trans e_gt0) ?v_gte ?ab_le. + by rewrite /nlim limN -ltr_normr (lt_trans e_gt0) ?v_gte ?ab_le. have posD v w: lt 0 v -> lt 0 w -> lt 0 (v + w). move=> /posP[m [d d_gt0 v_gtd]] /posP[n [e e_gt0 w_gte]]. apply/posP; exists (maxn m n), (d + e) => [|k]; first exact: addr_gt0. @@ -575,20 +575,24 @@ have add_Rroot xR p c: {yR | extendsR xR yR & has_Rroot xR p c -> root_in yR p}. rewrite !geq_max => /and3P[/ab_le/ub_rp{ub_rp}ub_rp le_mk le_nk]. rewrite -(ltr_add2r f) -mulr2n -mulr_natr divfK // /nlim /lim Dqvw. rewrite rmorphD hornerD /= -addrA -ltr_subl_addl ler_lt_add //. - by rewrite rmorphM hornerM ler_pmul ?ltrW ?v_gtd ?w_gte. - rewrite -ltr_pdivr_mull ?mulr_gt0 // (ler_lt_trans _ ub_rp) //. + by rewrite rmorphM hornerM ler_pmul ?ltW ?v_gtd ?w_gte. + rewrite -ltr_pdivr_mull ?mulr_gt0 // (le_lt_trans _ ub_rp) //. by rewrite -scalerAl hornerZ -rmorphM mulrN -normrN ler_norm. - pose le v w := (w == v) || lt v w. + pose le v w := (v == w) || lt v w. pose abs v := if le 0 v then v else - v. have absN v: abs (- v) = abs v. - rewrite /abs /le oppr_eq0 opprK posN. + rewrite /abs /le !(eq_sym 0) oppr_eq0 opprK posN. have [-> | /posVneg/orP[v_gt0 | v_lt0]] := altP eqP; first by rewrite oppr0. by rewrite v_gt0 /= -if_neg posNneg. by rewrite v_lt0 /= -if_neg -(opprK v) posN posNneg ?posN. have absE v: le 0 v -> abs v = v by rewrite /abs => ->. - pose QyNum := RealLtMixin posD posM posNneg posB posVneg absN absE (rrefl _). - pose QyNumField := [numFieldType of NumDomainType (Q y) QyNum]. - pose Ry := [realFieldType of RealDomainType _ (RealLeAxiom QyNumField)]. + pose QyNum : realLtMixin (Q y) := + RealLtMixin posD posM posNneg posB posVneg absN absE (rrefl _). + pose QyOrder := + OrderType + (DistrLatticeType (POrderType ring_display (Q y) QyNum) QyNum) QyNum. + pose QyNumField := [numFieldType of NumDomainType QyOrder QyNum]. + pose Ry := [realFieldType of [realDomainType of QyNumField]]. have archiRy := @rat_algebraic_archimedean Ry _ alg_integral. by exists (ArchiFieldType Ry archiRy); apply: [rmorphism of idfun]. have some_realC: realC. @@ -626,7 +630,7 @@ have sCle m n: (m <= n)%N -> {subset sQ (z_ m) <= sQ (z_ n)}. have R'i n: i \notin sQ (x_ n). rewrite /x_; case: (xR n) => x [Rn QxR] /=. apply: contraL (@ltr01 Rn) => /sQ_inQ[v Di]. - suffices /eqP <-: - QxR v ^+ 2 == 1 by rewrite oppr_gt0 -lerNgt sqr_ge0. + suffices /eqP <-: - QxR v ^+ 2 == 1 by rewrite oppr_gt0 -leNgt sqr_ge0. rewrite -rmorphX -rmorphN fmorph_eq1 -(fmorph_eq1 (ofQ x)) rmorphN eqr_oppLR. by rewrite rmorphX Di Di2. have szX2_1: size ('X^2 + 1) = 3. diff --git a/mathcomp/field/algnum.v b/mathcomp/field/algnum.v index c191a0e..be7dd6c 100644 --- a/mathcomp/field/algnum.v +++ b/mathcomp/field/algnum.v @@ -358,7 +358,8 @@ pose Saut (mu : subAut) : {rmorphism Sdom mu -> Sdom mu} := (projT2 mu).2. have SinjZ Qr (QrC : numF_inj Qr) a x: QrC (a *: x) = QtoC a * QrC x. rewrite mulrAC; apply: canRL (mulfK _) _. by rewrite intr_eq0 denq_neq0. - by rewrite mulrzr mulrzl -!rmorphMz scalerMzl -scaler_int -mulrzr -numqE. + by rewrite mulrzr mulrzl -!rmorphMz scalerMzl -[x *~ _]scaler_int -mulrzr + -numqE. have Sinj_poly Qr (QrC : numF_inj Qr) p: map_poly QrC (map_poly (in_alg Qr) p) = pQtoC p. - rewrite -map_poly_comp; apply: eq_map_poly => a. diff --git a/mathcomp/field/finfield.v b/mathcomp/field/finfield.v index 19871cb..f0d432f 100644 --- a/mathcomp/field/finfield.v +++ b/mathcomp/field/finfield.v @@ -568,7 +568,7 @@ End FinFieldExists. Section FinDomain. -Import ssrnum ssrint algC cyclotomic Num.Theory. +Import order ssrnum ssrint algC cyclotomic Order.TTheory Num.Theory. Local Infix "%|" := dvdn. (* Hide polynomial divisibility. *) Variable R : finUnitRingType. @@ -625,18 +625,18 @@ without loss n_gt1: / (1 < n)%N by rewrite ltnNge; apply: wlog_neg. have [q_gt0 n_gt0] := (ltnW q_gt1, ltnW n_gt1). have [z z_prim] := C_prim_root_exists n_gt0. have zn1: z ^+ n = 1 by apply: prim_expr_order. -have /eqP-n1z: `|z| == 1. - by rewrite -(pexpr_eq1 n_gt0) ?normr_ge0 // -normrX zn1 normr1. -suffices /eqP/normC_sub_eq[t n1t [Dq Dz]]: `|q%:R - z| == `|q%:R| - `|z|. +have /eqP-n1z: `|z| == 1 by rewrite -(pexpr_eq1 n_gt0) // -normrX zn1 normr1. +suffices /eqP/normC_sub_eq[t n1t [Dq Dz]]: + `|q%:R - z : algC| == `|q%:R : algC| - `|z|. suffices z1: z == 1 by rewrite leq_eqVlt -dvdn1 (prim_order_dvd z_prim) z1. by rewrite Dz n1z mul1r -(eqr_pmuln2r q_gt0) Dq normr_nat mulr_natl. pose aq d : algC := (cyclotomic (z ^+ (n %/ d)) d).[q%:R]. suffices: `|aq n| <= (q - 1)%:R. - rewrite eqr_le ler_sub_dist andbT n1z normr_nat natrB //; apply: ler_trans. + rewrite eq_le ler_sub_dist andbT n1z normr_nat natrB //; apply: le_trans. rewrite {}/aq horner_prod divnn n_gt0 expr1 normr_prod. rewrite (bigD1 (Ordinal n_gt1)) ?coprime1n //= !hornerE ler_pemulr //. elim/big_ind: _ => // [|d _]; first exact: mulr_ege1. - rewrite !hornerE; apply: ler_trans (ler_sub_dist _ _). + rewrite !hornerE; apply: le_trans (ler_sub_dist _ _). by rewrite normr_nat normrX n1z expr1n ler_subr_addl (leC_nat 2). have Zaq d: d %| n -> aq d \in Cint. move/(dvdn_prim_root z_prim)=> zd_prim. diff --git a/mathcomp/ssreflect/Make b/mathcomp/ssreflect/Make index c529f21..108f545 100644 --- a/mathcomp/ssreflect/Make +++ b/mathcomp/ssreflect/Make @@ -19,6 +19,7 @@ prime.v tuple.v ssrnotations.v ssrmatching.v +order.v -I . -R . mathcomp.ssreflect diff --git a/mathcomp/ssreflect/all_ssreflect.v b/mathcomp/ssreflect/all_ssreflect.v index aae57ca..318d5ef 100644 --- a/mathcomp/ssreflect/all_ssreflect.v +++ b/mathcomp/ssreflect/all_ssreflect.v @@ -14,5 +14,6 @@ Require Export finfun. Require Export bigop. Require Export prime. Require Export finset. +Require Export order. Require Export binomial. Require Export generic_quotient. diff --git a/mathcomp/ssreflect/fintype.v b/mathcomp/ssreflect/fintype.v index b6f618d..5a42c80 100644 --- a/mathcomp/ssreflect/fintype.v +++ b/mathcomp/ssreflect/fintype.v @@ -1051,16 +1051,16 @@ End Extremum. Notation "[ 'arg[' ord ]_( i < i0 | P ) F ]" := (extremum ord i0 (fun i => P%B) (fun i => F)) (at level 0, ord, i, i0 at level 10, - format "[ 'arg[' ord ]_( i < i0 | P ) F ]") : form_scope. + format "[ 'arg[' ord ]_( i < i0 | P ) F ]") : nat_scope. Notation "[ 'arg[' ord ]_( i < i0 'in' A ) F ]" := [arg[ord]_(i < i0 | i \in A) F] (at level 0, ord, i, i0 at level 10, - format "[ 'arg[' ord ]_( i < i0 'in' A ) F ]") : form_scope. + format "[ 'arg[' ord ]_( i < i0 'in' A ) F ]") : nat_scope. Notation "[ 'arg[' ord ]_( i < i0 ) F ]" := [arg[ord]_(i < i0 | true) F] (at level 0, ord, i, i0 at level 10, - format "[ 'arg[' ord ]_( i < i0 ) F ]") : form_scope. + format "[ 'arg[' ord ]_( i < i0 ) F ]") : nat_scope. Section ArgMinMax. @@ -1086,30 +1086,30 @@ End Extrema. Notation "[ 'arg' 'min_' ( i < i0 | P ) F ]" := (arg_min i0 (fun i => P%B) (fun i => F)) (at level 0, i, i0 at level 10, - format "[ 'arg' 'min_' ( i < i0 | P ) F ]") : form_scope. + format "[ 'arg' 'min_' ( i < i0 | P ) F ]") : nat_scope. Notation "[ 'arg' 'min_' ( i < i0 'in' A ) F ]" := [arg min_(i < i0 | i \in A) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'min_' ( i < i0 'in' A ) F ]") : form_scope. + format "[ 'arg' 'min_' ( i < i0 'in' A ) F ]") : nat_scope. Notation "[ 'arg' 'min_' ( i < i0 ) F ]" := [arg min_(i < i0 | true) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'min_' ( i < i0 ) F ]") : form_scope. + format "[ 'arg' 'min_' ( i < i0 ) F ]") : nat_scope. Notation "[ 'arg' 'max_' ( i > i0 | P ) F ]" := (arg_max i0 (fun i => P%B) (fun i => F)) (at level 0, i, i0 at level 10, - format "[ 'arg' 'max_' ( i > i0 | P ) F ]") : form_scope. + format "[ 'arg' 'max_' ( i > i0 | P ) F ]") : nat_scope. Notation "[ 'arg' 'max_' ( i > i0 'in' A ) F ]" := [arg max_(i > i0 | i \in A) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'max_' ( i > i0 'in' A ) F ]") : form_scope. + format "[ 'arg' 'max_' ( i > i0 'in' A ) F ]") : nat_scope. Notation "[ 'arg' 'max_' ( i > i0 ) F ]" := [arg max_(i > i0 | true) F] (at level 0, i, i0 at level 10, - format "[ 'arg' 'max_' ( i > i0 ) F ]") : form_scope. + format "[ 'arg' 'max_' ( i > i0 ) F ]") : nat_scope. (**********************************************************************) (* *) diff --git a/mathcomp/ssreflect/order.v b/mathcomp/ssreflect/order.v new file mode 100644 index 0000000..718eea5 --- /dev/null +++ b/mathcomp/ssreflect/order.v @@ -0,0 +1,6058 @@ +(* (c) Copyright 2006-2019 Microsoft Corporation and Inria. *) +(* Distributed under the terms of CeCILL-B. *) +From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat choice seq. +From mathcomp Require Import path fintype tuple bigop finset div prime. + +(******************************************************************************) +(* This files defines types equipped with order relations. *) +(* *) +(* Use one of the following modules implementing different theories: *) +(* Order.LTheory: partially ordered types and lattices excluding complement *) +(* and totality related theorems. *) +(* Order.CTheory: complemented lattices including Order.LTheory. *) +(* Order.TTheory: totally ordered types including Order.LTheory. *) +(* Order.Theory: ordered types including all of the above theory modules *) +(* *) +(* To access the definitions, notations, and the theory from, say, *) +(* "Order.Xyz", insert "Import Order.Xyz." at the top of your scripts. *) +(* Notations are accessible by opening the scope "order_scope" bound to the *) +(* delimiting key "O". *) +(* *) +(* We provide the following structures of ordered types *) +(* porderType d == the type of partially ordered types *) +(* distrLatticeType d == the type of distributive lattices *) +(* bDistrLatticeType d == distrLatticeType with a bottom element *) +(* tbDistrLatticeType d == distrLatticeType with both a top and a bottom *) +(* cbDistrLatticeType d == the type of sectionally complemented distributive*) +(* lattices *) +(* (lattices with bottom and a difference operation)*) +(* ctbDistrLatticeType d == the type of complemented distributive lattices *) +(* (lattices with top, bottom, difference, *) +(* and complement) *) +(* orderType d == the type of totally ordered types *) +(* finPOrderType d == the type of partially ordered finite types *) +(* finDistrLatticeType d == the type of nonempty finite distributive lattices*) +(* finCDistrLatticeType d == the type of nonempty finite complemented *) +(* distributive lattices *) +(* finOrderType d == the type of nonempty totally ordered finite types*) +(* *) +(* Each generic partial order and lattice operations symbols also has a first *) +(* argument which is the display, the second which is the minimal structure *) +(* they operate on and then the operands. Here is the exhaustive list of all *) +(* such symbols for partial orders and lattices together with their default *) +(* display (as displayed by Check). We document their meaning in the *) +(* paragraph adter the next. *) +(* *) +(* For porderType T *) +(* @Order.le disp T == <=%O (in fun_scope) *) +(* @Order.lt disp T == <%O (in fun_scope) *) +(* @Order.comparable disp T == >=<%O (in fun_scope) *) +(* @Order.ge disp T == >=%O (in fun_scope) *) +(* @Order.gt disp T == >%O (in fun_scope) *) +(* @Order.leif disp T == <?=%O (in fun_scope) *) +(* For distrLatticeType T *) +(* @Order.meet disp T x y == x `&` y (in order_scope) *) +(* @Order.join disp T x y == x `|` y (in order_scope) *) +(* For bDistrLatticeType T *) +(* @Order.bottom disp T == 0 (in order_scope) *) +(* For tbDistrLatticeType T *) +(* @Order.top disp T == 1 (in order_scope) *) +(* For cbDistrLatticeType T *) +(* @Order.sub disp T x y == x `|` y (in order_scope) *) +(* For ctbDistrLatticeType T *) +(* @Order.compl disp T x == ~` x (in order_scope) *) +(* *) +(* This first argument named either d, disp or display, of type unit, *) +(* configures the printing of notations. *) +(* Instantiating d with tt or an unknown key will lead to a default *) +(* display for notations, i.e. we have: *) +(* For x, y of type T, where T is canonically a porderType d: *) +(* x <= y <-> x is less than or equal to y. *) +(* x < y <-> x is less than y (:= (y != x) && (x <= y)). *) +(* x >= y <-> x is greater than or equal to y (:= y <= x). *) +(* x > y <-> x is greater than y (:= y < x). *) +(* x <= y ?= iff C <-> x is less than y, or equal iff C is true. *) +(* x >=< y <-> x and y are comparable (:= (x <= y) || (y <= x)). *) +(* x >< y <-> x and y are incomparable (:= ~~ x >=< y). *) +(* For x, y of type T, where T is canonically a distrLatticeType d: *) +(* x `&` y == the meet of x and y. *) +(* x `|` y == the join of x and y. *) +(* In a type T, where T is canonically a bDistrLatticeType d: *) +(* 0 == the bottom element. *) +(* \join_<range> e == iterated join of a lattice with a bottom. *) +(* In a type T, where T is canonically a tbDistrLatticeType d: *) +(* 1 == the top element. *) +(* \meet_<range> e == iterated meet of a lattice with a top. *) +(* For x, y of type T, where T is canonically a cbDistrLatticeType d: *) +(* x `\` y == the (sectional) complement of y in [0, x]. *) +(* For x of type T, where T is canonically a ctbDistrLatticeType d: *) +(* ~` x == the complement of x in [0, 1]. *) +(* *) +(* There are three distinct uses of the symbols *) +(* <, <=, >, >=, _ <= _ ?= iff _, >=<, and >< *) +(* in the default display: *) +(* they can be 0-ary, unary (prefix), and binary (infix). *) +(* 0. <%O, <=%O, >%O, >=%O, <?=%O, >=<%O, and ><%O stand respectively for *) +(* lt, le, gt, ge, leif (_ <= _ ?= iff _), comparable, and incomparable. *) +(* 1. (< x), (<= x), (> x), (>= x), (>=< x), and (>< x) stand respectively *) +(* for (>%O x), (>=%O x), (<%O x), (<=%O x), (>=<%O x), and (><%O x). *) +(* So (< x) is a predicate characterizing elements smaller than x. *) +(* 2. (x < y), (x <= y), ... mean what they are expected to. *) +(* These conventions are compatible with Haskell's, *) +(* where ((< y) x) = (x < y) = ((<) x y), *) +(* except that we write <%O instead of (<). *) +(* *) +(* Alternative notation displays can be defined by : *) +(* 1. declaring a new opaque definition of type unit. Using the idiom *) +(* `Lemma my_display : unit. Proof. exact: tt. Qed.` *) +(* 2. using this symbol to tag canonical porderType structures using *) +(* `Canonical my_porderType := POrderType my_display my_type my_mixin`, *) +(* 3. declaring notations for the main operations of this library, by *) +(* setting the first argument of the definition to the display, e.g. *) +(* `Notation my_syndef_le x y := @Order.le my_display _ x y.` or *) +(* `Notation "x <<< y" := @Order.lt my_display _ x y (at level ...).` *) +(* Non overloaded notations will default to the default display. *) +(* *) +(* One may use displays either for convenience or to desambiguate between *) +(* different structures defined on "copies" of a type (as explained below.) *) +(* We provide the following "copies" of types, *) +(* the first one is a *documented example* *) +(* natdvd := nat *) +(* == a "copy" of nat which is canonically ordered using *) +(* divisibility predicate dvdn. *) +(* Notation %|, %<|, gcd, lcm are used instead of *) +(* <=, <, meet and join. *) +(* T^c := converse T, *) +(* where converse is a new definition for (fun T => T) *) +(* == a "copy" of T, such that if T is canonically ordered, *) +(* then T^c is canonically ordered with the converse *) +(* order, and displayed with an extra ^c in the notation *) +(* i.e. <=^c, <^c, >=<^c, ><^c, `&`^c, `|`^c are *) +(* used and displayed instead of *) +(* <=, <, >=<, ><, `&`, `|` *) +(* T *prod[d] T' := T * T' *) +(* == a "copy" of the cartesian product such that, *) +(* if T and T' are canonically ordered, *) +(* then T *prod[d] T' is canonically ordered in product *) +(* order. *) +(* i.e. (x1, x2) <= (y1, y2) = *) +(* (x1 <= y1) && (x2 <= y2), *) +(* and displayed in display d *) +(* T *p T' := T *prod[prod_display] T' *) +(* where prod_display adds an extra ^p to all notations *) +(* T *lexi[d] T' := T * T' *) +(* == a "copy" of the cartesian product such that, *) +(* if T and T' are canonically ordered, *) +(* then T *lexi[d] T' is canonically ordered in *) +(* lexicographic order *) +(* i.e. (x1, x2) <= (y1, y2) = *) +(* (x1 <= y1) && ((x1 >= y1) ==> (x2 <= y2)) *) +(* and (x1, x2) < (y1, y2) = *) +(* (x1 <= y1) && ((x1 >= y1) ==> (x2 < y2)) *) +(* and displayed in display d *) +(* T *l T' := T *lexi[lexi_display] T' *) +(* where lexi_display adds an extra ^l to all notations *) +(* seqprod_with d T := seq T *) +(* == a "copy" of seq, such that if T is canonically *) +(* ordered, then seqprod_with d T is canonically ordered *) +(* in product order i.e. *) +(* [:: x1, .., xn] <= [y1, .., yn] = *) +(* (x1 <= y1) && ... && (xn <= yn) *) +(* and displayed in display d *) +(* n.-tupleprod[d] T == same with n.tuple T *) +(* seqprod T := seqprod_with prod_display T *) +(* n.-tupleprod T := n.-tuple[prod_display] T *) +(* seqlexi_with d T := seq T *) +(* == a "copy" of seq, such that if T is canonically *) +(* ordered, then seqprod_with d T is canonically ordered *) +(* in lexicographic order i.e. *) +(* [:: x1, .., xn] <= [y1, .., yn] = *) +(* (x1 <= x2) && ((x1 >= y1) ==> ((x2 <= y2) && ...)) *) +(* and displayed in display d *) +(* n.-tuplelexi[d] T == same with n.tuple T *) +(* seqlexi T := lexiprod_with lexi_display T *) +(* n.-tuplelexi T := n.-tuple[lexi_display] T *) +(* *) +(* Beware that canonical structure inference will not try to find the copy of *) +(* the structures that fits the display one mentioned, but will rather *) +(* determine which canonical structure and display to use depending on the *) +(* copy of the type one provided. In this sense they are merely displays *) +(* to inform the user of what the inferrence did, rather than additional *) +(* input for the inference. *) +(* *) +(* Existing displays are either converse_display d (where d is a display), *) +(* dvd_display (both explained above), total_display (to overload meet and *) +(* join using min and max) ring_display (from algebra/ssrnum to change the *) +(* scope of the usual notations to ring_scope). We also provide lexi_display *) +(* and prod_display for lexicographic and product order respectively. *) +(* The default display is tt and users can define their own as explained *) +(* above. *) +(* *) +(* For orderType we provide the following operations (in total_display) *) +(* [arg minr_(i < i0 | P) M] == a value i : T minimizing M : R, subject to *) +(* the condition P (i may appear in P and M), and *) +(* provided P holds for i0. *) +(* [arg maxr_(i > i0 | P) M] == a value i maximizing M subject to P and *) +(* provided P holds for i0. *) +(* [arg min_(i < i0 in A) M] == an i \in A minimizing M if i0 \in A. *) +(* [arg max_(i > i0 in A) M] == an i \in A maximizing M if i0 \in A. *) +(* [arg min_(i < i0) M] == an i : T minimizing M, given i0 : T. *) +(* [arg max_(i > i0) M] == an i : T maximizing M, given i0 : T. *) +(* with head symbols Order.arg_min and Order.arg_max *) +(* *) +(* In order to build the above structures, one must provide the appropriate *) +(* mixin to the following structure constructors. The list of possible mixins *) +(* is indicated after each constructor. Each mixin is documented in the next *) +(* paragraph. *) +(* *) +(* POrderType disp T pord_mixin *) +(* == builds a porderType from a canonical choiceType *) +(* instance of T where pord_mixin can be of types *) +(* lePOrderMixin, ltPOrderMixin, meetJoinMixin, *) +(* leOrderMixin, or ltOrderMixin *) +(* or computed using PcanPOrderMixin or CanPOrderMixin. *) +(* disp is a display as explained above *) +(* *) +(* DistrLatticeType T lat_mixin *) +(* == builds a distrLatticeType from a porderType where *) +(* lat_mixin can be of types *) +(* latticeMixin, totalPOrderMixin, meetJoinMixin, *) +(* leOrderMixin, or ltOrderMixin *) +(* or computed using IsoLatticeMixin. *) +(* *) +(* BLatticeType T bot_mixin *) +(* == builds a bDistrLatticeType from a distrLatticeType and *) +(* a bottom element *) +(* *) +(* TBLatticeType T top_mixin *) +(* == builds a tbDistrLatticeType from a bDistrLatticeType *) +(* and a top element *) +(* *) +(* CBLatticeType T sub_mixin *) +(* == builds a cbDistrLatticeType from a bDistrLatticeType *) +(* and a difference operation *) +(* *) +(* CTBLatticeType T compl_mixin *) +(* == builds a ctbDistrLatticeType from a tbDistrLatticeType *) +(* and a complement operation *) +(* *) +(* OrderType T ord_mixin *) +(* == builds an orderType from a distrLatticeType where *) +(* ord_mixin can be of types *) +(* leOrderMixin, ltOrderMixin, or orderMixin, *) +(* or computed using MonoTotalMixin. *) +(* *) +(* Additionally: *) +(* - [porderType of _] ... notations are available to recover structures on *) +(* "copies" of the types, as in eqtype, choicetype, ssralg... *) +(* - [finPOrderType of _] ... notations to compute joins between finite types *) +(* and ordered types *) +(* *) +(* List of possible mixins (a.k.a. factories): *) +(* *) +(* - lePOrderMixin == on a choiceType, takes le, lt, *) +(* reflexivity, antisymmetry and transitivity of le. *) +(* (can build: porderType) *) +(* *) +(* - ltPOrderMixin == on a choiceType, takes le, lt, *) +(* irreflexivity and transitivity of lt. *) +(* (can build: porderType) *) +(* *) +(* - meetJoinMixin == on a choiceType, takes le, lt, meet, join, *) +(* commutativity and associativity of meet and join *) +(* idempotence of meet and some De Morgan laws *) +(* (can build: porderType, distrLatticeType) *) +(* *) +(* - leOrderMixin == on a choiceType, takes le, lt, meet, join *) +(* antisymmetry, transitivity and totality of le. *) +(* (can build: porderType, distrLatticeType, orderType) *) +(* *) +(* - ltOrderMixin == on a choiceType, takes le, lt, *) +(* irreflexivity, transitivity and totality of lt. *) +(* (can build: porderType, distrLatticeType, orderType) *) +(* *) +(* - totalPOrderMixin == on a porderType T, totality of the order of T *) +(* := total (<=%O : rel T) *) +(* (can build: distrLatticeType) *) +(* *) +(* - totalOrderMixin == on a distrLatticeType T, totality of the order of T *) +(* := total (<=%O : rel T) *) +(* (can build: orderType) *) +(* NB: this mixin is kept separate from totalPOrderMixin (even though it *) +(* is convertible to it), in order to avoid ambiguous coercion paths. *) +(* *) +(* - distrLatticeMixin == on a porderType T, takes meet, join *) +(* commutativity and associativity of meet and join *) +(* idempotence of meet and some De Morgan laws *) +(* (can build: distrLatticeType) *) +(* *) +(* - bDistrLatticeMixin, tbDistrLatticeMixin, cbDistrLatticeMixin, *) +(* ctbDistrLatticeMixin *) +(* == mixins with one extra operation *) +(* (respectively bottom, top, relative complement, and *) +(* total complement) *) +(* *) +(* Additionally: *) +(* - [porderMixin of T by <:] creates a porderMixin by subtyping. *) +(* - [totalOrderMixin of T by <:] creates the associated totalOrderMixin. *) +(* - PCanPOrderMixin, CanPOrderMixin create porderMixin from cancellations *) +(* - MonoTotalMixin creates a totalPOrderMixin from monotonicity *) +(* - IsoLatticeMixin creates a distrLatticeMixin from an ordered structure *) +(* isomorphism (i.e., cancel f f', cancel f' f, {mono f : x y / x <= y}) *) +(* *) +(* We provide the following canonical instances of ordered types *) +(* - all possible structures on bool *) +(* - porderType, distrLatticeType, orderType and bDistrLatticeType on nat for *) +(* the leq order *) +(* - porderType, distrLatticeType, bDistrLatticeType, cbDistrLatticeType, *) +(* ctbDistrLatticeType on nat for the dvdn order, where meet and join *) +(* are respectively gcdn and lcmn *) +(* - porderType, distrLatticeType, orderType, bDistrLatticeType, *) +(* tbDistrLatticeType, cbDistrLatticeType, ctbDistrLatticeType *) +(* on T *prod[disp] T' a "copy" of T * T' *) +(* using product order (and T *p T' its specialization to prod_display) *) +(* - porderType, distrLatticeType, and orderType, on T *lexi[disp] T' *) +(* another "copy" of T * T', with lexicographic ordering *) +(* (and T *l T' its specialization to lexi_display) *) +(* - porderType, distrLatticeType, and orderType, on {t : T & T' x} *) +(* with lexicographic ordering *) +(* - porderType, distrLatticeType, orderType, bDistrLatticeType, *) +(* cbDistrLatticeType, tbDistrLatticeType, ctbDistrLatticeType *) +(* on seqprod_with disp T a "copy" of seq T *) +(* using product order (and seqprod T' its specialization to prod_display)*) +(* - porderType, distrLatticeType, and orderType, on seqlexi_with disp T *) +(* another "copy" of seq T, with lexicographic ordering *) +(* (and seqlexi T its specialization to lexi_display) *) +(* - porderType, distrLatticeType, orderType, bDistrLatticeType, *) +(* cbDistrLatticeType, tbDistrLatticeType, ctbDistrLatticeType *) +(* on n.-tupleprod[disp] a "copy" of n.-tuple T *) +(* using product order (and n.-tupleprod T its specialization *) +(* to prod_display) *) +(* - porderType, distrLatticeType, and orderType, on n.-tuplelexi[d] T *) +(* another "copy" of n.-tuple T, with lexicographic ordering *) +(* (and n.-tuplelexi T its specialization to lexi_display) *) +(* and all possible finite type instances *) +(* *) +(* In order to get a canonical order on prod or seq, one may import modules *) +(* DefaultProdOrder or DefaultProdLexiOrder, DefaultSeqProdOrder or *) +(* DefaultSeqLexiOrder, and DefaultTupleProdOrder or DefaultTupleLexiOrder. *) +(* *) +(* On orderType, leP ltP ltgtP are the three main lemmas for case analysis. *) +(* On porderType, one may use comparableP, comparable_leP, comparable_ltP, *) +(* and comparable_ltgtP, which are the four main lemmas for case analysis. *) +(* *) +(* We also provide specialized versions of some theorems from path.v. *) +(* *) +(* This file is based on prior work by *) +(* D. Dreyer, G. Gonthier, A. Nanevski, P-Y Strub, B. Ziliani *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Delimit Scope order_scope with O. +Local Open Scope order_scope. + +Reserved Notation "<= y" (at level 35). +Reserved Notation ">= y" (at level 35). +Reserved Notation "< y" (at level 35). +Reserved Notation "> y" (at level 35). +Reserved Notation "<= y :> T" (at level 35, y at next level). +Reserved Notation ">= y :> T" (at level 35, y at next level). +Reserved Notation "< y :> T" (at level 35, y at next level). +Reserved Notation "> y :> T" (at level 35, y at next level). +Reserved Notation "x >=< y" (at level 70, no associativity). +Reserved Notation ">=< x" (at level 35). +Reserved Notation ">=< y :> T" (at level 35, y at next level). +Reserved Notation "x >< y" (at level 70, no associativity). +Reserved Notation ">< x" (at level 35). +Reserved Notation ">< y :> T" (at level 35, y at next level). + +(* Reserved notation for lattice operations. *) +Reserved Notation "A `&` B" (at level 48, left associativity). +Reserved Notation "A `|` B" (at level 52, left associativity). +Reserved Notation "A `\` B" (at level 50, left associativity). +Reserved Notation "~` A" (at level 35, right associativity). + +(* Notations for converse partial and total order *) +Reserved Notation "x <=^c y" (at level 70, y at next level). +Reserved Notation "x >=^c y" (at level 70, y at next level, only parsing). +Reserved Notation "x <^c y" (at level 70, y at next level). +Reserved Notation "x >^c y" (at level 70, y at next level, only parsing). +Reserved Notation "x <=^c y :> T" (at level 70, y at next level). +Reserved Notation "x >=^c y :> T" (at level 70, y at next level, only parsing). +Reserved Notation "x <^c y :> T" (at level 70, y at next level). +Reserved Notation "x >^c y :> T" (at level 70, y at next level, only parsing). +Reserved Notation "<=^c y" (at level 35). +Reserved Notation ">=^c y" (at level 35). +Reserved Notation "<^c y" (at level 35). +Reserved Notation ">^c y" (at level 35). +Reserved Notation "<=^c y :> T" (at level 35, y at next level). +Reserved Notation ">=^c y :> T" (at level 35, y at next level). +Reserved Notation "<^c y :> T" (at level 35, y at next level). +Reserved Notation ">^c y :> T" (at level 35, y at next level). +Reserved Notation "x >=<^c y" (at level 70, no associativity). +Reserved Notation ">=<^c x" (at level 35). +Reserved Notation ">=<^c y :> T" (at level 35, y at next level). +Reserved Notation "x ><^c y" (at level 70, no associativity). +Reserved Notation "><^c x" (at level 35). +Reserved Notation "><^c y :> T" (at level 35, y at next level). + +Reserved Notation "x <=^c y <=^c z" (at level 70, y, z at next level). +Reserved Notation "x <^c y <=^c z" (at level 70, y, z at next level). +Reserved Notation "x <=^c y <^c z" (at level 70, y, z at next level). +Reserved Notation "x <^c y <^c z" (at level 70, y, z at next level). +Reserved Notation "x <=^c y ?= 'iff' c" (at level 70, y, c at next level, + format "x '[hv' <=^c y '/' ?= 'iff' c ']'"). +Reserved Notation "x <=^c y ?= 'iff' c :> T" (at level 70, y, c at next level, + format "x '[hv' <=^c y '/' ?= 'iff' c :> T ']'"). + +(* Reserved notation for converse lattice operations. *) +Reserved Notation "A `&^c` B" (at level 48, left associativity). +Reserved Notation "A `|^c` B" (at level 52, left associativity). +Reserved Notation "A `\^c` B" (at level 50, left associativity). +Reserved Notation "~^c` A" (at level 35, right associativity). + +(* Reserved notations for product ordering of prod or seq *) +Reserved Notation "x <=^p y" (at level 70, y at next level). +Reserved Notation "x >=^p y" (at level 70, y at next level, only parsing). +Reserved Notation "x <^p y" (at level 70, y at next level). +Reserved Notation "x >^p y" (at level 70, y at next level, only parsing). +Reserved Notation "x <=^p y :> T" (at level 70, y at next level). +Reserved Notation "x >=^p y :> T" (at level 70, y at next level, only parsing). +Reserved Notation "x <^p y :> T" (at level 70, y at next level). +Reserved Notation "x >^p y :> T" (at level 70, y at next level, only parsing). +Reserved Notation "<=^p y" (at level 35). +Reserved Notation ">=^p y" (at level 35). +Reserved Notation "<^p y" (at level 35). +Reserved Notation ">^p y" (at level 35). +Reserved Notation "<=^p y :> T" (at level 35, y at next level). +Reserved Notation ">=^p y :> T" (at level 35, y at next level). +Reserved Notation "<^p y :> T" (at level 35, y at next level). +Reserved Notation ">^p y :> T" (at level 35, y at next level). +Reserved Notation "x >=<^p y" (at level 70, no associativity). +Reserved Notation ">=<^p x" (at level 35). +Reserved Notation ">=<^p y :> T" (at level 35, y at next level). +Reserved Notation "x ><^p y" (at level 70, no associativity). +Reserved Notation "><^p x" (at level 35). +Reserved Notation "><^p y :> T" (at level 35, y at next level). + +Reserved Notation "x <=^p y <=^p z" (at level 70, y, z at next level). +Reserved Notation "x <^p y <=^p z" (at level 70, y, z at next level). +Reserved Notation "x <=^p y <^p z" (at level 70, y, z at next level). +Reserved Notation "x <^p y <^p z" (at level 70, y, z at next level). +Reserved Notation "x <=^p y ?= 'iff' c" (at level 70, y, c at next level, + format "x '[hv' <=^p y '/' ?= 'iff' c ']'"). +Reserved Notation "x <=^p y ?= 'iff' c :> T" (at level 70, y, c at next level, + format "x '[hv' <=^p y '/' ?= 'iff' c :> T ']'"). + +(* Reserved notation for converse lattice operations. *) +Reserved Notation "A `&^p` B" (at level 48, left associativity). +Reserved Notation "A `|^p` B" (at level 52, left associativity). +Reserved Notation "A `\^p` B" (at level 50, left associativity). +Reserved Notation "~^p` A" (at level 35, right associativity). + +(* Reserved notations for lexicographic ordering of prod or seq *) +Reserved Notation "x <=^l y" (at level 70, y at next level). +Reserved Notation "x >=^l y" (at level 70, y at next level, only parsing). +Reserved Notation "x <^l y" (at level 70, y at next level). +Reserved Notation "x >^l y" (at level 70, y at next level, only parsing). +Reserved Notation "x <=^l y :> T" (at level 70, y at next level). +Reserved Notation "x >=^l y :> T" (at level 70, y at next level, only parsing). +Reserved Notation "x <^l y :> T" (at level 70, y at next level). +Reserved Notation "x >^l y :> T" (at level 70, y at next level, only parsing). +Reserved Notation "<=^l y" (at level 35). +Reserved Notation ">=^l y" (at level 35). +Reserved Notation "<^l y" (at level 35). +Reserved Notation ">^l y" (at level 35). +Reserved Notation "<=^l y :> T" (at level 35, y at next level). +Reserved Notation ">=^l y :> T" (at level 35, y at next level). +Reserved Notation "<^l y :> T" (at level 35, y at next level). +Reserved Notation ">^l y :> T" (at level 35, y at next level). +Reserved Notation "x >=<^l y" (at level 70, no associativity). +Reserved Notation ">=<^l x" (at level 35). +Reserved Notation ">=<^l y :> T" (at level 35, y at next level). +Reserved Notation "x ><^l y" (at level 70, no associativity). +Reserved Notation "><^l x" (at level 35). +Reserved Notation "><^l y :> T" (at level 35, y at next level). + +Reserved Notation "x <=^l y <=^l z" (at level 70, y, z at next level). +Reserved Notation "x <^l y <=^l z" (at level 70, y, z at next level). +Reserved Notation "x <=^l y <^l z" (at level 70, y, z at next level). +Reserved Notation "x <^l y <^l z" (at level 70, y, z at next level). +Reserved Notation "x <=^l y ?= 'iff' c" (at level 70, y, c at next level, + format "x '[hv' <=^l y '/' ?= 'iff' c ']'"). +Reserved Notation "x <=^l y ?= 'iff' c :> T" (at level 70, y, c at next level, + format "x '[hv' <=^l y '/' ?= 'iff' c :> T ']'"). + +(* Reserved notations for divisibility *) +Reserved Notation "x %<| y" (at level 70, no associativity). + +Reserved Notation "\gcd_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \gcd_ i '/ ' F ']'"). +Reserved Notation "\gcd_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \gcd_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \gcd_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \gcd_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \gcd_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \gcd_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\gcd_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\gcd_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \gcd_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \gcd_ ( i < n ) F ']'"). +Reserved Notation "\gcd_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \gcd_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\gcd_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \gcd_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\lcm_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \lcm_ i '/ ' F ']'"). +Reserved Notation "\lcm_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \lcm_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \lcm_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \lcm_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \lcm_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \lcm_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\lcm_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\lcm_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \lcm_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \lcm_ ( i < n ) F ']'"). +Reserved Notation "\lcm_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \lcm_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\lcm_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \lcm_ ( i 'in' A ) '/ ' F ']'"). + +(* Reserved notation for converse lattice operations. *) +Reserved Notation "A `&^l` B" (at level 48, left associativity). +Reserved Notation "A `|^l` B" (at level 52, left associativity). +Reserved Notation "A `\^l` B" (at level 50, left associativity). +Reserved Notation "~^l` A" (at level 35, right associativity). + +Reserved Notation "\meet_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \meet_ i '/ ' F ']'"). +Reserved Notation "\meet_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \meet_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\meet_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \meet_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\meet_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \meet_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\meet_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \meet_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\meet_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \meet_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\meet_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\meet_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\meet_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \meet_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\meet_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \meet_ ( i < n ) F ']'"). +Reserved Notation "\meet_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \meet_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\meet_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \meet_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\join_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \join_ i '/ ' F ']'"). +Reserved Notation "\join_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \join_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\join_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \join_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\join_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \join_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\join_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \join_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\join_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \join_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\join_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\join_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\join_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \join_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\join_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \join_ ( i < n ) F ']'"). +Reserved Notation "\join_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \join_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\join_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \join_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\min_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \min_ i '/ ' F ']'"). +Reserved Notation "\min_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \min_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\min_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \min_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\min_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \min_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\min_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \min_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\min_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \min_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\min_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\min_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\min_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \min_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\min_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \min_ ( i < n ) F ']'"). +Reserved Notation "\min_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \min_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\min_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \min_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\meet^c_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \meet^c_ i '/ ' F ']'"). +Reserved Notation "\meet^c_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \meet^c_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \meet^c_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \meet^c_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \meet^c_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \meet^c_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\meet^c_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\meet^c_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \meet^c_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \meet^c_ ( i < n ) F ']'"). +Reserved Notation "\meet^c_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \meet^c_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\meet^c_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \meet^c_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\join^c_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \join^c_ i '/ ' F ']'"). +Reserved Notation "\join^c_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \join^c_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \join^c_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \join^c_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \join^c_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \join^c_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\join^c_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\join^c_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \join^c_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \join^c_ ( i < n ) F ']'"). +Reserved Notation "\join^c_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \join^c_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\join^c_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \join^c_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\meet^p_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \meet^p_ i '/ ' F ']'"). +Reserved Notation "\meet^p_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \meet^p_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \meet^p_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \meet^p_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \meet^p_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \meet^p_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\meet^p_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\meet^p_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \meet^p_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \meet^p_ ( i < n ) F ']'"). +Reserved Notation "\meet^p_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \meet^p_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\meet^p_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \meet^p_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\join^p_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \join^p_ i '/ ' F ']'"). +Reserved Notation "\join^p_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \join^p_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \join^p_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \join^p_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \join^p_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \join^p_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\join^p_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\join^p_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \join^p_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \join^p_ ( i < n ) F ']'"). +Reserved Notation "\join^p_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \join^p_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\join^p_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \join^p_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\min^l_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \min^l_ i '/ ' F ']'"). +Reserved Notation "\min^l_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \min^l_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \min^l_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \min^l_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \min^l_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \min^l_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\min^l_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\min^l_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \min^l_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \min^l_ ( i < n ) F ']'"). +Reserved Notation "\min^l_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \min^l_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\min^l_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \min^l_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\max^l_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \max^l_ i '/ ' F ']'"). +Reserved Notation "\max^l_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \max^l_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \max^l_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \max^l_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \max^l_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \max^l_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\max^l_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\max^l_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \max^l_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \max^l_ ( i < n ) F ']'"). +Reserved Notation "\max^l_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \max^l_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\max^l_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \max^l_ ( i 'in' A ) '/ ' F ']'"). + +(* tuple extensions *) +Lemma eqEtuple n (T : eqType) (t1 t2 : n.-tuple T) : + (t1 == t2) = [forall i, tnth t1 i == tnth t2 i]. +Proof. by apply/eqP/'forall_eqP => [->|/eq_from_tnth]. Qed. + +Lemma tnth_nseq n T x (i : 'I_n) : @tnth n T [tuple of nseq n x] i = x. +Proof. +by rewrite !(tnth_nth (tnth_default (nseq_tuple n x) i)) nth_nseq ltn_ord. +Qed. + +Lemma tnthS n T x (t : n.-tuple T) i : + tnth [tuple of x :: t] (lift ord0 i) = tnth t i. +Proof. by rewrite (tnth_nth (tnth_default t i)). Qed. + +Module Order. + +(**************) +(* STRUCTURES *) +(**************) + +Module POrder. +Section ClassDef. +Record mixin_of (T : eqType) := Mixin { + le : rel T; + lt : rel T; + _ : forall x y, lt x y = (y != x) && (le x y); + _ : reflexive le; + _ : antisymmetric le; + _ : transitive le +}. + +Record class_of T := Class { + base : Choice.class_of T; + mixin : mixin_of (EqType T base) +}. + +Local Coercion base : class_of >-> Choice.class_of. + +Structure type (disp : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack disp T c. +Definition clone_with disp' c of phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (Choice.class bT) b => + fun m => Pack disp (@Class T b m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> Choice.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Canonical eqType. +Canonical choiceType. +Notation porderType := type. +Notation lePOrderMixin := mixin_of. +Notation LePOrderMixin := Mixin. +Notation POrderType disp T m := (@pack T disp _ _ id m). +Notation "[ 'porderType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'porderType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'porderType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, format "[ 'porderType' 'of' T 'for' cT 'with' disp ]") : + form_scope. +Notation "[ 'porderType' 'of' T ]" := [porderType of T for _] + (at level 0, format "[ 'porderType' 'of' T ]") : form_scope. +Notation "[ 'porderType' 'of' T 'with' disp ]" := + [porderType of T for _ with disp] + (at level 0, format "[ 'porderType' 'of' T 'with' disp ]") : form_scope. +End Exports. + +End POrder. +Import POrder.Exports. +Bind Scope cpo_sort with POrder.sort. + +Section POrderDef. + +Variable (disp : unit). +Local Notation porderType := (porderType disp). +Variable (T : porderType). + +Definition le : rel T := POrder.le (POrder.class T). +Local Notation "x <= y" := (le x y) : order_scope. + +Definition lt : rel T := POrder.lt (POrder.class T). +Local Notation "x < y" := (lt x y) : order_scope. + +Definition comparable : rel T := fun (x y : T) => (x <= y) || (y <= x). +Local Notation "x >=< y" := (comparable x y) : order_scope. +Local Notation "x >< y" := (~~ (x >=< y)) : order_scope. + +Definition ge : simpl_rel T := [rel x y | y <= x]. +Definition gt : simpl_rel T := [rel x y | y < x]. +Definition leif (x y : T) C : Prop := ((x <= y) * ((x == y) = C))%type. + +Definition le_of_leif x y C (le_xy : @leif x y C) := le_xy.1 : le x y. + +Variant le_xor_gt (x y : T) : bool -> bool -> Set := + | LeNotGt of x <= y : le_xor_gt x y true false + | GtNotLe of y < x : le_xor_gt x y false true. + +Variant lt_xor_ge (x y : T) : bool -> bool -> Set := + | LtNotGe of x < y : lt_xor_ge x y false true + | GeNotLt of y <= x : lt_xor_ge x y true false. + +Variant compare (x y : T) : + bool -> bool -> bool -> bool -> bool -> bool -> Set := + | CompareLt of x < y : compare x y + false false false true false true + | CompareGt of y < x : compare x y + false false true false true false + | CompareEq of x = y : compare x y + true true true true false false. + +Variant incompare (x y : T) : + bool -> bool -> bool -> bool -> bool -> bool -> bool -> bool -> Set := + | InCompareLt of x < y : incompare x y + false false false true false true true true + | InCompareGt of y < x : incompare x y + false false true false true false true true + | InCompare of x >< y : incompare x y + false false false false false false false false + | InCompareEq of x = y : incompare x y + true true true true false false true true. + +End POrderDef. + +Prenex Implicits lt le leif. +Arguments ge {_ _}. +Arguments gt {_ _}. + +Module Import POSyntax. + +Notation "<=%O" := le : fun_scope. +Notation ">=%O" := ge : fun_scope. +Notation "<%O" := lt : fun_scope. +Notation ">%O" := gt : fun_scope. +Notation "<?=%O" := leif : fun_scope. +Notation ">=<%O" := comparable : fun_scope. +Notation "><%O" := (fun x y => ~~ (comparable x y)) : fun_scope. + +Notation "<= y" := (ge y) : order_scope. +Notation "<= y :> T" := (<= (y : T)) (only parsing) : order_scope. +Notation ">= y" := (le y) : order_scope. +Notation ">= y :> T" := (>= (y : T)) (only parsing) : order_scope. + +Notation "< y" := (gt y) : order_scope. +Notation "< y :> T" := (< (y : T)) (only parsing) : order_scope. +Notation "> y" := (lt y) : order_scope. +Notation "> y :> T" := (> (y : T)) (only parsing) : order_scope. + +Notation ">=< y" := (comparable y) : order_scope. +Notation ">=< y :> T" := (>=< (y : T)) (only parsing) : order_scope. + +Notation "x <= y" := (le x y) : order_scope. +Notation "x <= y :> T" := ((x : T) <= (y : T)) (only parsing) : order_scope. +Notation "x >= y" := (y <= x) (only parsing) : order_scope. +Notation "x >= y :> T" := ((x : T) >= (y : T)) (only parsing) : order_scope. + +Notation "x < y" := (lt x y) : order_scope. +Notation "x < y :> T" := ((x : T) < (y : T)) (only parsing) : order_scope. +Notation "x > y" := (y < x) (only parsing) : order_scope. +Notation "x > y :> T" := ((x : T) > (y : T)) (only parsing) : order_scope. + +Notation "x <= y <= z" := ((x <= y) && (y <= z)) : order_scope. +Notation "x < y <= z" := ((x < y) && (y <= z)) : order_scope. +Notation "x <= y < z" := ((x <= y) && (y < z)) : order_scope. +Notation "x < y < z" := ((x < y) && (y < z)) : order_scope. + +Notation "x <= y ?= 'iff' C" := (leif x y C) : order_scope. +Notation "x <= y ?= 'iff' C :> T" := ((x : T) <= (y : T) ?= iff C) + (only parsing) : order_scope. + +Notation ">=< x" := (comparable x) : order_scope. +Notation ">=< x :> T" := (>=< (x : T)) (only parsing) : order_scope. +Notation "x >=< y" := (comparable x y) : order_scope. + +Notation ">< x" := (fun y => ~~ (comparable x y)) : order_scope. +Notation ">< x :> T" := (>< (x : T)) (only parsing) : order_scope. +Notation "x >< y" := (~~ (comparable x y)) : order_scope. + +End POSyntax. + +Module POCoercions. +Coercion le_of_leif : leif >-> is_true. +End POCoercions. + +Module DistrLattice. +Section ClassDef. + +Record mixin_of d (T : porderType d) := Mixin { + meet : T -> T -> T; + join : T -> T -> T; + _ : commutative meet; + _ : commutative join; + _ : associative meet; + _ : associative join; + _ : forall y x, meet x (join x y) = x; + _ : forall y x, join x (meet x y) = x; + _ : forall x y, (x <= y) = (meet x y == x); + _ : left_distributive meet join; +}. + +Record class_of (T : Type) := Class { + base : POrder.class_of T; + mixin_disp : unit; + mixin : mixin_of (POrder.Pack mixin_disp base) +}. + +Local Coercion base : class_of >-> POrder.class_of. + +Structure type (disp : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack disp T c. +Definition clone_with disp' c of phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack d0 b0 (m0 : mixin_of (@POrder.Pack d0 T b0)) := + fun bT b & phant_id (@POrder.class disp bT) b => + fun m & phant_id m0 m => Pack disp (@Class T b d0 m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> POrder.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion porderType : type >-> POrder.type. +Canonical eqType. +Canonical choiceType. +Canonical porderType. +Notation distrLatticeType := type. +Notation distrLatticeMixin := mixin_of. +Notation DistrLatticeMixin := Mixin. +Notation DistrLatticeType T m := (@pack T _ _ _ m _ _ id _ id). +Notation "[ 'distrLatticeType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'distrLatticeType' 'of' T 'for' cT ]") : + form_scope. +Notation "[ 'distrLatticeType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, + format "[ 'distrLatticeType' 'of' T 'for' cT 'with' disp ]") : + form_scope. +Notation "[ 'distrLatticeType' 'of' T ]" := [distrLatticeType of T for _] + (at level 0, format "[ 'distrLatticeType' 'of' T ]") : form_scope. +Notation "[ 'distrLatticeType' 'of' T 'with' disp ]" := + [distrLatticeType of T for _ with disp] + (at level 0, format "[ 'distrLatticeType' 'of' T 'with' disp ]") : + form_scope. +End Exports. + +End DistrLattice. +Export DistrLattice.Exports. + +Section DistrLatticeDef. +Context {disp : unit}. +Local Notation distrLatticeType := (distrLatticeType disp). +Context {T : distrLatticeType}. +Definition meet : T -> T -> T := DistrLattice.meet (DistrLattice.class T). +Definition join : T -> T -> T := DistrLattice.join (DistrLattice.class T). + +Variant lel_xor_gt (x y : T) : bool -> bool -> T -> T -> T -> T -> Set := + | LelNotGt of x <= y : lel_xor_gt x y true false x x y y + | GtlNotLe of y < x : lel_xor_gt x y false true y y x x. + +Variant ltl_xor_ge (x y : T) : bool -> bool -> T -> T -> T -> T -> Set := + | LtlNotGe of x < y : ltl_xor_ge x y false true x x y y + | GelNotLt of y <= x : ltl_xor_ge x y true false y y x x. + +Variant comparel (x y : T) : + bool -> bool -> bool -> bool -> bool -> bool -> T -> T -> T -> T -> Set := + | ComparelLt of x < y : comparel x y + false false false true false true x x y y + | ComparelGt of y < x : comparel x y + false false true false true false y y x x + | ComparelEq of x = y : comparel x y + true true true true false false x x x x. + +Variant incomparel (x y : T) : + bool -> bool -> bool -> bool -> bool -> bool -> bool -> bool -> + T -> T -> T -> T -> Set := + | InComparelLt of x < y : incomparel x y + false false false true false true true true x x y y + | InComparelGt of y < x : incomparel x y + false false true false true false true true y y x x + | InComparel of x >< y : incomparel x y + false false false false false false false false + (meet x y) (meet x y) (join x y) (join x y) + | InComparelEq of x = y : incomparel x y + true true true true false false true true x x x x. + +End DistrLatticeDef. + +Module Import DistrLatticeSyntax. + +Notation "x `&` y" := (meet x y) : order_scope. +Notation "x `|` y" := (join x y) : order_scope. + +End DistrLatticeSyntax. + +Module Total. +Definition mixin_of d (T : distrLatticeType d) := total (<=%O : rel T). +Section ClassDef. + +Record class_of (T : Type) := Class { + base : DistrLattice.class_of T; + mixin_disp : unit; + mixin : mixin_of (DistrLattice.Pack mixin_disp base) +}. + +Local Coercion base : class_of >-> DistrLattice.class_of. + +Structure type (d : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c & phant_id class c := @Pack disp T c. +Definition clone_with disp' c & phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack d0 b0 (m0 : mixin_of (@DistrLattice.Pack d0 T b0)) := + fun bT b & phant_id (@DistrLattice.class disp bT) b => + fun m & phant_id m0 m => Pack disp (@Class T b d0 m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> DistrLattice.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion porderType : type >-> POrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical porderType. +Canonical distrLatticeType. +Notation totalOrderMixin := Total.mixin_of. +Notation orderType := type. +Notation OrderType T m := (@pack T _ _ _ m _ _ id _ id). +Notation "[ 'orderType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'orderType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'orderType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, format "[ 'orderType' 'of' T 'for' cT 'with' disp ]") : + form_scope. +Notation "[ 'orderType' 'of' T ]" := [orderType of T for _] + (at level 0, format "[ 'orderType' 'of' T ]") : form_scope. +Notation "[ 'orderType' 'of' T 'with' disp ]" := + [orderType of T for _ with disp] + (at level 0, format "[ 'orderType' 'of' T 'with' disp ]") : form_scope. +End Exports. + +End Total. +Import Total.Exports. + +Module BDistrLattice. +Section ClassDef. +Record mixin_of d (T : porderType d) := Mixin { + bottom : T; + _ : forall x, bottom <= x; +}. + +Record class_of (T : Type) := Class { + base : DistrLattice.class_of T; + mixin_disp : unit; + mixin : mixin_of (POrder.Pack mixin_disp base) +}. + +Local Coercion base : class_of >-> DistrLattice.class_of. + +Structure type (d : unit) := Pack { sort; _ : class_of sort}. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack disp T c. +Definition clone_with disp' c of phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack d0 b0 (m0 : mixin_of (@DistrLattice.Pack d0 T b0)) := + fun bT b & phant_id (@DistrLattice.class disp bT) b => + fun m & phant_id m0 m => Pack disp (@Class T b d0 m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> DistrLattice.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion porderType : type >-> POrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical porderType. +Canonical distrLatticeType. +Notation bDistrLatticeType := type. +Notation bDistrLatticeMixin := mixin_of. +Notation BDistrLatticeMixin := Mixin. +Notation BDistrLatticeType T m := (@pack T _ _ _ m _ _ id _ id). +Notation "[ 'bDistrLatticeType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'bDistrLatticeType' 'of' T 'for' cT ]") : + form_scope. +Notation "[ 'bDistrLatticeType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, + format "[ 'bDistrLatticeType' 'of' T 'for' cT 'with' disp ]") : + form_scope. +Notation "[ 'bDistrLatticeType' 'of' T ]" := [bDistrLatticeType of T for _] + (at level 0, format "[ 'bDistrLatticeType' 'of' T ]") : form_scope. +Notation "[ 'bDistrLatticeType' 'of' T 'with' disp ]" := + [bDistrLatticeType of T for _ with disp] + (at level 0, format "[ 'bDistrLatticeType' 'of' T 'with' disp ]") : + form_scope. +End Exports. + +End BDistrLattice. +Export BDistrLattice.Exports. + +Definition bottom {disp : unit} {T : bDistrLatticeType disp} : T := + BDistrLattice.bottom (BDistrLattice.class T). + +Module Import BDistrLatticeSyntax. +Notation "0" := bottom : order_scope. + +Notation "\join_ ( i <- r | P ) F" := + (\big[@join _ _/0%O]_(i <- r | P%B) F%O) : order_scope. +Notation "\join_ ( i <- r ) F" := + (\big[@join _ _/0%O]_(i <- r) F%O) : order_scope. +Notation "\join_ ( i | P ) F" := + (\big[@join _ _/0%O]_(i | P%B) F%O) : order_scope. +Notation "\join_ i F" := + (\big[@join _ _/0%O]_i F%O) : order_scope. +Notation "\join_ ( i : I | P ) F" := + (\big[@join _ _/0%O]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\join_ ( i : I ) F" := + (\big[@join _ _/0%O]_(i : I) F%O) (only parsing) : order_scope. +Notation "\join_ ( m <= i < n | P ) F" := + (\big[@join _ _/0%O]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\join_ ( m <= i < n ) F" := + (\big[@join _ _/0%O]_(m <= i < n) F%O) : order_scope. +Notation "\join_ ( i < n | P ) F" := + (\big[@join _ _/0%O]_(i < n | P%B) F%O) : order_scope. +Notation "\join_ ( i < n ) F" := + (\big[@join _ _/0%O]_(i < n) F%O) : order_scope. +Notation "\join_ ( i 'in' A | P ) F" := + (\big[@join _ _/0%O]_(i in A | P%B) F%O) : order_scope. +Notation "\join_ ( i 'in' A ) F" := + (\big[@join _ _/0%O]_(i in A) F%O) : order_scope. + +End BDistrLatticeSyntax. + +Module TBDistrLattice. +Section ClassDef. +Record mixin_of d (T : porderType d) := Mixin { + top : T; + _ : forall x, x <= top; +}. + +Record class_of (T : Type) := Class { + base : BDistrLattice.class_of T; + mixin_disp : unit; + mixin : mixin_of (POrder.Pack mixin_disp base) +}. + +Local Coercion base : class_of >-> BDistrLattice.class_of. + +Structure type (d : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack disp T c. +Definition clone_with disp' c of phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack d0 b0 (m0 : mixin_of (@BDistrLattice.Pack d0 T b0)) := + fun bT b & phant_id (@BDistrLattice.class disp bT) b => + fun m & phant_id m0 m => Pack disp (@Class T b d0 m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +Definition bDistrLatticeType := @BDistrLattice.Pack disp cT xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> BDistrLattice.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion porderType : type >-> POrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Coercion bDistrLatticeType : type >-> BDistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Notation tbDistrLatticeType := type. +Notation tbDistrLatticeMixin := mixin_of. +Notation TBDistrLatticeMixin := Mixin. +Notation TBDistrLatticeType T m := (@pack T _ _ _ m _ _ id _ id). +Notation "[ 'tbDistrLatticeType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'tbDistrLatticeType' 'of' T 'for' cT ]") : + form_scope. +Notation "[ 'tbDistrLatticeType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, + format "[ 'tbDistrLatticeType' 'of' T 'for' cT 'with' disp ]") : + form_scope. +Notation "[ 'tbDistrLatticeType' 'of' T ]" := [tbDistrLatticeType of T for _] + (at level 0, format "[ 'tbDistrLatticeType' 'of' T ]") : form_scope. +Notation "[ 'tbDistrLatticeType' 'of' T 'with' disp ]" := + [tbDistrLatticeType of T for _ with disp] + (at level 0, format "[ 'tbDistrLatticeType' 'of' T 'with' disp ]") : + form_scope. +End Exports. + +End TBDistrLattice. +Export TBDistrLattice.Exports. + +Definition top disp {T : tbDistrLatticeType disp} : T := + TBDistrLattice.top (TBDistrLattice.class T). + +Module Import TBDistrLatticeSyntax. + +Notation "1" := top : order_scope. + +Notation "\meet_ ( i <- r | P ) F" := + (\big[meet/1]_(i <- r | P%B) F%O) : order_scope. +Notation "\meet_ ( i <- r ) F" := + (\big[meet/1]_(i <- r) F%O) : order_scope. +Notation "\meet_ ( i | P ) F" := + (\big[meet/1]_(i | P%B) F%O) : order_scope. +Notation "\meet_ i F" := + (\big[meet/1]_i F%O) : order_scope. +Notation "\meet_ ( i : I | P ) F" := + (\big[meet/1]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\meet_ ( i : I ) F" := + (\big[meet/1]_(i : I) F%O) (only parsing) : order_scope. +Notation "\meet_ ( m <= i < n | P ) F" := + (\big[meet/1]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\meet_ ( m <= i < n ) F" := + (\big[meet/1]_(m <= i < n) F%O) : order_scope. +Notation "\meet_ ( i < n | P ) F" := + (\big[meet/1]_(i < n | P%B) F%O) : order_scope. +Notation "\meet_ ( i < n ) F" := + (\big[meet/1]_(i < n) F%O) : order_scope. +Notation "\meet_ ( i 'in' A | P ) F" := + (\big[meet/1]_(i in A | P%B) F%O) : order_scope. +Notation "\meet_ ( i 'in' A ) F" := + (\big[meet/1]_(i in A) F%O) : order_scope. + +End TBDistrLatticeSyntax. + +Module CBDistrLattice. +Section ClassDef. +Record mixin_of d (T : bDistrLatticeType d) := Mixin { + sub : T -> T -> T; + _ : forall x y, y `&` sub x y = bottom; + _ : forall x y, (x `&` y) `|` sub x y = x +}. + +Record class_of (T : Type) := Class { + base : BDistrLattice.class_of T; + mixin_disp : unit; + mixin : mixin_of (BDistrLattice.Pack mixin_disp base) +}. + +Local Coercion base : class_of >-> BDistrLattice.class_of. + +Structure type (d : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack disp T c. +Definition clone_with disp' c of phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack d0 b0 (m0 : mixin_of (@BDistrLattice.Pack d0 T b0)) := + fun bT b & phant_id (@BDistrLattice.class disp bT) b => + fun m & phant_id m0 m => Pack disp (@Class T b d0 m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +Definition bDistrLatticeType := @BDistrLattice.Pack disp cT xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> BDistrLattice.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion porderType : type >-> POrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Coercion bDistrLatticeType : type >-> BDistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Notation cbDistrLatticeType := type. +Notation cbDistrLatticeMixin := mixin_of. +Notation CBDistrLatticeMixin := Mixin. +Notation CBDistrLatticeType T m := (@pack T _ _ _ m _ _ id _ id). +Notation "[ 'cbDistrLatticeType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'cbDistrLatticeType' 'of' T 'for' cT ]") : + form_scope. +Notation "[ 'cbDistrLatticeType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, + format "[ 'cbDistrLatticeType' 'of' T 'for' cT 'with' disp ]") : + form_scope. +Notation "[ 'cbDistrLatticeType' 'of' T ]" := [cbDistrLatticeType of T for _] + (at level 0, format "[ 'cbDistrLatticeType' 'of' T ]") : form_scope. +Notation "[ 'cbDistrLatticeType' 'of' T 'with' disp ]" := + [cbDistrLatticeType of T for _ with disp] + (at level 0, format "[ 'cbDistrLatticeType' 'of' T 'with' disp ]") : + form_scope. +End Exports. + +End CBDistrLattice. +Export CBDistrLattice.Exports. + +Definition sub {disp : unit} {T : cbDistrLatticeType disp} : T -> T -> T := + CBDistrLattice.sub (CBDistrLattice.class T). + +Module Import CBDistrLatticeSyntax. +Notation "x `\` y" := (sub x y) : order_scope. +End CBDistrLatticeSyntax. + +Module CTBDistrLattice. +Section ClassDef. +Record mixin_of d (T : tbDistrLatticeType d) (sub : T -> T -> T) := Mixin { + compl : T -> T; + _ : forall x, compl x = sub top x +}. + +Record class_of (T : Type) := Class { + base : TBDistrLattice.class_of T; + mixin1_disp : unit; + mixin1 : CBDistrLattice.mixin_of (BDistrLattice.Pack mixin1_disp base); + mixin2_disp : unit; + mixin2 : @mixin_of _ (TBDistrLattice.Pack mixin2_disp base) + (CBDistrLattice.sub mixin1) +}. + +Local Coercion base : class_of >-> TBDistrLattice.class_of. +Local Coercion base2 T (c : class_of T) : CBDistrLattice.class_of T := + CBDistrLattice.Class (mixin1 c). + +Structure type (d : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack disp T c. +Definition clone_with disp' c of phant_id class c := @Pack disp' T c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (@TBDistrLattice.class disp bT) b => + fun disp1 m1T m1 & phant_id (@CBDistrLattice.class disp1 m1T) + (@CBDistrLattice.Class _ _ _ m1) => + fun disp2 m2 => Pack disp (@Class T b disp1 m1 disp2 m2). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +Definition bDistrLatticeType := @BDistrLattice.Pack disp cT xclass. +Definition tbDistrLatticeType := @TBDistrLattice.Pack disp cT xclass. +Definition cbDistrLatticeType := @CBDistrLattice.Pack disp cT xclass. +Definition tbd_cbDistrLatticeType := + @CBDistrLattice.Pack disp tbDistrLatticeType xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> TBDistrLattice.class_of. +Coercion base2 : class_of >-> CBDistrLattice.class_of. +Coercion mixin1 : class_of >-> CBDistrLattice.mixin_of. +Coercion mixin2 : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion porderType : type >-> POrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Coercion bDistrLatticeType : type >-> BDistrLattice.type. +Coercion tbDistrLatticeType : type >-> TBDistrLattice.type. +Coercion cbDistrLatticeType : type >-> CBDistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical cbDistrLatticeType. +Canonical tbd_cbDistrLatticeType. +Notation ctbDistrLatticeType := type. +Notation ctbDistrLatticeMixin := mixin_of. +Notation CTBDistrLatticeMixin := Mixin. +Notation CTBDistrLatticeType T m := (@pack T _ _ _ id _ _ _ id _ m). +Notation "[ 'ctbDistrLatticeType' 'of' T 'for' cT ]" := (@clone T _ cT _ id) + (at level 0, format "[ 'ctbDistrLatticeType' 'of' T 'for' cT ]") : + form_scope. +Notation "[ 'ctbDistrLatticeType' 'of' T 'for' cT 'with' disp ]" := + (@clone_with T _ cT disp _ id) + (at level 0, + format "[ 'ctbDistrLatticeType' 'of' T 'for' cT 'with' disp ]") + : form_scope. +Notation "[ 'ctbDistrLatticeType' 'of' T ]" := [ctbDistrLatticeType of T for _] + (at level 0, format "[ 'ctbDistrLatticeType' 'of' T ]") : form_scope. +Notation "[ 'ctbDistrLatticeType' 'of' T 'with' disp ]" := + [ctbDistrLatticeType of T for _ with disp] + (at level 0, format "[ 'ctbDistrLatticeType' 'of' T 'with' disp ]") : + form_scope. +Notation "[ 'default_ctbDistrLatticeType' 'of' T ]" := + (@pack T _ _ _ id _ _ id (Mixin (fun=> erefl))) + (at level 0, format "[ 'default_ctbDistrLatticeType' 'of' T ]") : + form_scope. +End Exports. + +End CTBDistrLattice. +Export CTBDistrLattice.Exports. + +Definition compl {d : unit} {T : ctbDistrLatticeType d} : T -> T := + CTBDistrLattice.compl (CTBDistrLattice.class T). + +Module Import CTBDistrLatticeSyntax. +Notation "~` A" := (compl A) : order_scope. +End CTBDistrLatticeSyntax. + +Section TotalDef. +Context {disp : unit} {T : orderType disp} {I : finType}. +Definition arg_min := @extremum T I <=%O. +Definition arg_max := @extremum T I >=%O. +End TotalDef. + +Module Import TotalSyntax. + +Fact total_display : unit. Proof. exact: tt. Qed. + +Notation max := (@join total_display _). +Notation "@ 'max' T" := + (@join total_display T) (at level 10, T at level 8, only parsing) : fun_scope. +Notation min := (@meet total_display _). +Notation "@ 'min' T" := + (@meet total_display T) (at level 10, T at level 8, only parsing) : fun_scope. + +Notation "\max_ ( i <- r | P ) F" := + (\big[max/0%O]_(i <- r | P%B) F%O) : order_scope. +Notation "\max_ ( i <- r ) F" := + (\big[max/0%O]_(i <- r) F%O) : order_scope. +Notation "\max_ ( i | P ) F" := + (\big[max/0%O]_(i | P%B) F%O) : order_scope. +Notation "\max_ i F" := + (\big[max/0%O]_i F%O) : order_scope. +Notation "\max_ ( i : I | P ) F" := + (\big[max/0%O]_(i : I | P%B) F%O) (only parsing) : + order_scope. +Notation "\max_ ( i : I ) F" := + (\big[max/0%O]_(i : I) F%O) (only parsing) : order_scope. +Notation "\max_ ( m <= i < n | P ) F" := + (\big[max/0%O]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\max_ ( m <= i < n ) F" := + (\big[max/0%O]_(m <= i < n) F%O) : order_scope. +Notation "\max_ ( i < n | P ) F" := + (\big[max/0%O]_(i < n | P%B) F%O) : order_scope. +Notation "\max_ ( i < n ) F" := + (\big[max/0%O]_(i < n) F%O) : order_scope. +Notation "\max_ ( i 'in' A | P ) F" := + (\big[max/0%O]_(i in A | P%B) F%O) : order_scope. +Notation "\max_ ( i 'in' A ) F" := + (\big[max/0%O]_(i in A) F%O) : order_scope. + +Notation "\min_ ( i <- r | P ) F" := + (\big[min/1%O]_(i <- r | P%B) F%O) : order_scope. +Notation "\min_ ( i <- r ) F" := + (\big[min/1%O]_(i <- r) F%O) : order_scope. +Notation "\min_ ( i | P ) F" := + (\big[min/1%O]_(i | P%B) F%O) : order_scope. +Notation "\min_ i F" := + (\big[min/1%O]_i F%O) : order_scope. +Notation "\min_ ( i : I | P ) F" := + (\big[min/1%O]_(i : I | P%B) F%O) (only parsing) : + order_scope. +Notation "\min_ ( i : I ) F" := + (\big[min/1%O]_(i : I) F%O) (only parsing) : order_scope. +Notation "\min_ ( m <= i < n | P ) F" := + (\big[min/1%O]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\min_ ( m <= i < n ) F" := + (\big[min/1%O]_(m <= i < n) F%O) : order_scope. +Notation "\min_ ( i < n | P ) F" := + (\big[min/1%O]_(i < n | P%B) F%O) : order_scope. +Notation "\min_ ( i < n ) F" := + (\big[min/1%O]_(i < n) F%O) : order_scope. +Notation "\min_ ( i 'in' A | P ) F" := + (\big[min/1%O]_(i in A | P%B) F%O) : order_scope. +Notation "\min_ ( i 'in' A ) F" := + (\big[min/1%O]_(i in A) F%O) : order_scope. + +Notation "[ 'arg' 'min_' ( i < i0 | P ) F ]" := + (arg_min i0 (fun i => P%B) (fun i => F)) + (at level 0, i, i0 at level 10, + format "[ 'arg' 'min_' ( i < i0 | P ) F ]") : order_scope. + +Notation "[ 'arg' 'min_' ( i < i0 'in' A ) F ]" := + [arg min_(i < i0 | i \in A) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'min_' ( i < i0 'in' A ) F ]") : order_scope. + +Notation "[ 'arg' 'min_' ( i < i0 ) F ]" := [arg min_(i < i0 | true) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'min_' ( i < i0 ) F ]") : order_scope. + +Notation "[ 'arg' 'max_' ( i > i0 | P ) F ]" := + (arg_max i0 (fun i => P%B) (fun i => F)) + (at level 0, i, i0 at level 10, + format "[ 'arg' 'max_' ( i > i0 | P ) F ]") : order_scope. + +Notation "[ 'arg' 'max_' ( i > i0 'in' A ) F ]" := + [arg max_(i > i0 | i \in A) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'max_' ( i > i0 'in' A ) F ]") : order_scope. + +Notation "[ 'arg' 'max_' ( i > i0 ) F ]" := [arg max_(i > i0 | true) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'max_' ( i > i0 ) F ]") : order_scope. + +End TotalSyntax. + +(**********) +(* FINITE *) +(**********) + +Module FinPOrder. +Section ClassDef. + +Record class_of T := Class { + base : POrder.class_of T; + mixin : Finite.mixin_of (Equality.Pack base) +}. + +Local Coercion base : class_of >-> POrder.class_of. +Local Coercion base2 T (c : class_of T) : Finite.class_of T := + Finite.Class (mixin c). + +Structure type (disp : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (@POrder.class disp bT) b => + fun mT m & phant_id (@Finite.class mT) (@Finite.Class _ _ m) => + Pack disp (@Class T b m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition countType := @Countable.Pack cT xclass. +Definition finType := @Finite.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition count_porderType := @POrder.Pack disp countType xclass. +Definition fin_porderType := @POrder.Pack disp finType xclass. +End ClassDef. + +Module Exports. +Coercion base : class_of >-> POrder.class_of. +Coercion base2 : class_of >-> Finite.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion countType : type >-> Countable.type. +Coercion finType : type >-> Finite.type. +Coercion porderType : type >-> POrder.type. +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical count_porderType. +Canonical fin_porderType. +Notation finPOrderType := type. +Notation "[ 'finPOrderType' 'of' T ]" := (@pack T _ _ _ id _ _ id) + (at level 0, format "[ 'finPOrderType' 'of' T ]") : form_scope. +End Exports. + +End FinPOrder. +Import FinPOrder.Exports. +Bind Scope cpo_sort with FinPOrder.sort. + +Module FinDistrLattice. +Section ClassDef. + +Record class_of (T : Type) := Class { + base : TBDistrLattice.class_of T; + mixin : Finite.mixin_of (Equality.Pack base); +}. + +Local Coercion base : class_of >-> TBDistrLattice.class_of. +Local Coercion base2 T (c : class_of T) : Finite.class_of T := + Finite.Class (mixin c). +Local Coercion base3 T (c : class_of T) : FinPOrder.class_of T := + @FinPOrder.Class T c c. + +Structure type (disp : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (@TBDistrLattice.class disp bT) b => + fun mT m & phant_id (@Finite.class mT) (@Finite.Class _ _ m) => + Pack disp (@Class T b m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition countType := @Countable.Pack cT xclass. +Definition finType := @Finite.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition finPOrderType := @FinPOrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +Definition bDistrLatticeType := @BDistrLattice.Pack disp cT xclass. +Definition tbDistrLatticeType := @TBDistrLattice.Pack disp cT xclass. +Definition count_distrLatticeType := @DistrLattice.Pack disp countType xclass. +Definition count_bDistrLatticeType := @BDistrLattice.Pack disp countType xclass. +Definition count_tbDistrLatticeType := + @TBDistrLattice.Pack disp countType xclass. +Definition fin_distrLatticeType := @DistrLattice.Pack disp finType xclass. +Definition fin_bDistrLatticeType := @BDistrLattice.Pack disp finType xclass. +Definition fin_tbDistrLatticeType := @TBDistrLattice.Pack disp finType xclass. +Definition finPOrder_distrLatticeType := + @DistrLattice.Pack disp finPOrderType xclass. +Definition finPOrder_bDistrLatticeType := + @BDistrLattice.Pack disp finPOrderType xclass. +Definition finPOrder_tbDistrLatticeType := + @TBDistrLattice.Pack disp finPOrderType xclass. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> TBDistrLattice.class_of. +Coercion base2 : class_of >-> Finite.class_of. +Coercion base3 : class_of >-> FinPOrder.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion countType : type >-> Countable.type. +Coercion finType : type >-> Finite.type. +Coercion porderType : type >-> POrder.type. +Coercion finPOrderType : type >-> FinPOrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Coercion bDistrLatticeType : type >-> BDistrLattice.type. +Coercion tbDistrLatticeType : type >-> TBDistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical finPOrderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical count_distrLatticeType. +Canonical count_bDistrLatticeType. +Canonical count_tbDistrLatticeType. +Canonical fin_distrLatticeType. +Canonical fin_bDistrLatticeType. +Canonical fin_tbDistrLatticeType. +Canonical finPOrder_distrLatticeType. +Canonical finPOrder_bDistrLatticeType. +Canonical finPOrder_tbDistrLatticeType. +Notation finDistrLatticeType := type. +Notation "[ 'finDistrLatticeType' 'of' T ]" := (@pack T _ _ _ id _ _ id) + (at level 0, format "[ 'finDistrLatticeType' 'of' T ]") : form_scope. +End Exports. + +End FinDistrLattice. +Export FinDistrLattice.Exports. + +Module FinCDistrLattice. +Section ClassDef. + +Record class_of (T : Type) := Class { + base : CTBDistrLattice.class_of T; + mixin : Finite.mixin_of (Equality.Pack base); +}. + +Local Coercion base : class_of >-> CTBDistrLattice.class_of. +Local Coercion base2 T (c : class_of T) : Finite.class_of T := + Finite.Class (mixin c). +Local Coercion base3 T (c : class_of T) : FinDistrLattice.class_of T := + @FinDistrLattice.Class T c c. + +Structure type (disp : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (@CTBDistrLattice.class disp bT) b => + fun mT m & phant_id (@Finite.class mT) (@Finite.Class _ _ m) => + Pack disp (@Class T b m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition countType := @Countable.Pack cT xclass. +Definition finType := @Finite.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition finPOrderType := @FinPOrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +Definition bDistrLatticeType := @BDistrLattice.Pack disp cT xclass. +Definition tbDistrLatticeType := @TBDistrLattice.Pack disp cT xclass. +Definition finDistrLatticeType := @FinDistrLattice.Pack disp cT xclass. +Definition cbDistrLatticeType := @CBDistrLattice.Pack disp cT xclass. +Definition ctbDistrLatticeType := @CTBDistrLattice.Pack disp cT xclass. +Definition count_cbDistrLatticeType := + @CBDistrLattice.Pack disp countType xclass. +Definition count_ctbDistrLatticeType := + @CTBDistrLattice.Pack disp countType xclass. +Definition fin_cbDistrLatticeType := @CBDistrLattice.Pack disp finType xclass. +Definition fin_ctbDistrLatticeType := @CTBDistrLattice.Pack disp finType xclass. +Definition finPOrder_cbDistrLatticeType := + @CBDistrLattice.Pack disp finPOrderType xclass. +Definition finPOrder_ctbDistrLatticeType := + @CTBDistrLattice.Pack disp finPOrderType xclass. +Definition finDistrLattice_cbDistrLatticeType := + @CBDistrLattice.Pack disp finDistrLatticeType xclass. +Definition finDistrLattice_ctbDistrLatticeType := + @CTBDistrLattice.Pack disp finDistrLatticeType xclass. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> CTBDistrLattice.class_of. +Coercion base2 : class_of >-> Finite.class_of. +Coercion base3 : class_of >-> FinDistrLattice.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion countType : type >-> Countable.type. +Coercion finType : type >-> Finite.type. +Coercion porderType : type >-> POrder.type. +Coercion finPOrderType : type >-> FinPOrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Coercion bDistrLatticeType : type >-> BDistrLattice.type. +Coercion tbDistrLatticeType : type >-> TBDistrLattice.type. +Coercion finDistrLatticeType : type >-> FinDistrLattice.type. +Coercion cbDistrLatticeType : type >-> CBDistrLattice.type. +Coercion ctbDistrLatticeType : type >-> CTBDistrLattice.type. +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical finPOrderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical finDistrLatticeType. +Canonical cbDistrLatticeType. +Canonical ctbDistrLatticeType. +Canonical count_cbDistrLatticeType. +Canonical count_ctbDistrLatticeType. +Canonical fin_cbDistrLatticeType. +Canonical fin_ctbDistrLatticeType. +Canonical finPOrder_cbDistrLatticeType. +Canonical finPOrder_ctbDistrLatticeType. +Canonical finDistrLattice_cbDistrLatticeType. +Canonical finDistrLattice_ctbDistrLatticeType. +Notation finCDistrLatticeType := type. +Notation "[ 'finCDistrLatticeType' 'of' T ]" := (@pack T _ _ _ id _ _ id) + (at level 0, format "[ 'finCDistrLatticeType' 'of' T ]") : form_scope. +End Exports. + +End FinCDistrLattice. +Export FinCDistrLattice.Exports. + +Module FinTotal. +Section ClassDef. + +Record class_of (T : Type) := Class { + base : FinDistrLattice.class_of T; + mixin_disp : unit; + mixin : totalOrderMixin (DistrLattice.Pack mixin_disp base) +}. + +Local Coercion base : class_of >-> FinDistrLattice.class_of. +Local Coercion base2 T (c : class_of T) : Total.class_of T := + @Total.Class _ c _ (mixin (c := c)). + +Structure type (disp : unit) := Pack { sort; _ : class_of sort }. + +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (disp : unit) (cT : type disp). + +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack := + fun bT b & phant_id (@FinDistrLattice.class disp bT) b => + fun disp' mT m & phant_id (@Total.class disp mT) (@Total.Class _ _ _ m) => + Pack disp (@Class T b disp' m). + +Definition eqType := @Equality.Pack cT xclass. +Definition choiceType := @Choice.Pack cT xclass. +Definition countType := @Countable.Pack cT xclass. +Definition finType := @Finite.Pack cT xclass. +Definition porderType := @POrder.Pack disp cT xclass. +Definition finPOrderType := @FinPOrder.Pack disp cT xclass. +Definition distrLatticeType := @DistrLattice.Pack disp cT xclass. +Definition bDistrLatticeType := @BDistrLattice.Pack disp cT xclass. +Definition tbDistrLatticeType := @TBDistrLattice.Pack disp cT xclass. +Definition finDistrLatticeType := @FinDistrLattice.Pack disp cT xclass. +Definition orderType := @Total.Pack disp cT xclass. +Definition order_countType := @Countable.Pack orderType xclass. +Definition order_finType := @Finite.Pack orderType xclass. +Definition order_finPOrderType := @FinPOrder.Pack disp orderType xclass. +Definition order_bDistrLatticeType := @BDistrLattice.Pack disp orderType xclass. +Definition order_tbDistrLatticeType := + @TBDistrLattice.Pack disp orderType xclass. +Definition order_finDistrLatticeType := + @FinDistrLattice.Pack disp orderType xclass. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> FinDistrLattice.class_of. +Coercion base2 : class_of >-> Total.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Coercion choiceType : type >-> Choice.type. +Coercion countType : type >-> Countable.type. +Coercion finType : type >-> Finite.type. +Coercion porderType : type >-> POrder.type. +Coercion finPOrderType : type >-> FinPOrder.type. +Coercion distrLatticeType : type >-> DistrLattice.type. +Coercion bDistrLatticeType : type >-> BDistrLattice.type. +Coercion tbDistrLatticeType : type >-> TBDistrLattice.type. +Coercion finDistrLatticeType : type >-> FinDistrLattice.type. +Coercion orderType : type >-> Total.type. +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical finPOrderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical finDistrLatticeType. +Canonical orderType. +Canonical order_countType. +Canonical order_finType. +Canonical order_finPOrderType. +Canonical order_bDistrLatticeType. +Canonical order_tbDistrLatticeType. +Canonical order_finDistrLatticeType. +Notation finOrderType := type. +Notation "[ 'finOrderType' 'of' T ]" := (@pack T _ _ _ id _ _ _ id) + (at level 0, format "[ 'finOrderType' 'of' T ]") : form_scope. +End Exports. + +End FinTotal. +Export FinTotal.Exports. + +(************) +(* CONVERSE *) +(************) + +Definition converse T : Type := T. +Definition converse_display : unit -> unit. Proof. exact. Qed. +Local Notation "T ^c" := (converse T) (at level 2, format "T ^c") : type_scope. + +Module Import ConverseSyntax. + +Notation "<=^c%O" := (@le (converse_display _) _) : fun_scope. +Notation ">=^c%O" := (@ge (converse_display _) _) : fun_scope. +Notation ">=^c%O" := (@ge (converse_display _) _) : fun_scope. +Notation "<^c%O" := (@lt (converse_display _) _) : fun_scope. +Notation ">^c%O" := (@gt (converse_display _) _) : fun_scope. +Notation "<?=^c%O" := (@leif (converse_display _) _) : fun_scope. +Notation ">=<^c%O" := (@comparable (converse_display _) _) : fun_scope. +Notation "><^c%O" := (fun x y => ~~ (@comparable (converse_display _) _ x y)) : + fun_scope. + +Notation "<=^c y" := (>=^c%O y) : order_scope. +Notation "<=^c y :> T" := (<=^c (y : T)) (only parsing) : order_scope. +Notation ">=^c y" := (<=^c%O y) : order_scope. +Notation ">=^c y :> T" := (>=^c (y : T)) (only parsing) : order_scope. + +Notation "<^c y" := (>^c%O y) : order_scope. +Notation "<^c y :> T" := (<^c (y : T)) (only parsing) : order_scope. +Notation ">^c y" := (<^c%O y) : order_scope. +Notation ">^c y :> T" := (>^c (y : T)) (only parsing) : order_scope. + +Notation ">=<^c y" := (>=<^c%O y) : order_scope. +Notation ">=<^c y :> T" := (>=<^c (y : T)) (only parsing) : order_scope. + +Notation "x <=^c y" := (<=^c%O x y) : order_scope. +Notation "x <=^c y :> T" := ((x : T) <=^c (y : T)) (only parsing) : order_scope. +Notation "x >=^c y" := (y <=^c x) (only parsing) : order_scope. +Notation "x >=^c y :> T" := ((x : T) >=^c (y : T)) (only parsing) : order_scope. + +Notation "x <^c y" := (<^c%O x y) : order_scope. +Notation "x <^c y :> T" := ((x : T) <^c (y : T)) (only parsing) : order_scope. +Notation "x >^c y" := (y <^c x) (only parsing) : order_scope. +Notation "x >^c y :> T" := ((x : T) >^c (y : T)) (only parsing) : order_scope. + +Notation "x <=^c y <=^c z" := ((x <=^c y) && (y <=^c z)) : order_scope. +Notation "x <^c y <=^c z" := ((x <^c y) && (y <=^c z)) : order_scope. +Notation "x <=^c y <^c z" := ((x <=^c y) && (y <^c z)) : order_scope. +Notation "x <^c y <^c z" := ((x <^c y) && (y <^c z)) : order_scope. + +Notation "x <=^c y ?= 'iff' C" := (<?=^c%O x y C) : order_scope. +Notation "x <=^c y ?= 'iff' C :> T" := ((x : T) <=^c (y : T) ?= iff C) + (only parsing) : order_scope. + +Notation ">=<^c x" := (>=<^c%O x) : order_scope. +Notation ">=<^c x :> T" := (>=<^c (x : T)) (only parsing) : order_scope. +Notation "x >=<^c y" := (>=<^c%O x y) : order_scope. + +Notation "><^c x" := (fun y => ~~ (>=<^c%O x y)) : order_scope. +Notation "><^c x :> T" := (><^c (x : T)) (only parsing) : order_scope. +Notation "x ><^c y" := (~~ (><^c%O x y)) : order_scope. + +Notation "x `&^c` y" := (@meet (converse_display _) _ x y) : order_scope. +Notation "x `|^c` y" := (@join (converse_display _) _ x y) : order_scope. + +Local Notation "0" := bottom. +Local Notation "1" := top. +Local Notation join := (@join (converse_display _) _). +Local Notation meet := (@meet (converse_display _) _). + +Notation "\join^c_ ( i <- r | P ) F" := + (\big[join/0]_(i <- r | P%B) F%O) : order_scope. +Notation "\join^c_ ( i <- r ) F" := + (\big[join/0]_(i <- r) F%O) : order_scope. +Notation "\join^c_ ( i | P ) F" := + (\big[join/0]_(i | P%B) F%O) : order_scope. +Notation "\join^c_ i F" := + (\big[join/0]_i F%O) : order_scope. +Notation "\join^c_ ( i : I | P ) F" := + (\big[join/0]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\join^c_ ( i : I ) F" := + (\big[join/0]_(i : I) F%O) (only parsing) : order_scope. +Notation "\join^c_ ( m <= i < n | P ) F" := + (\big[join/0]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\join^c_ ( m <= i < n ) F" := + (\big[join/0]_(m <= i < n) F%O) : order_scope. +Notation "\join^c_ ( i < n | P ) F" := + (\big[join/0]_(i < n | P%B) F%O) : order_scope. +Notation "\join^c_ ( i < n ) F" := + (\big[join/0]_(i < n) F%O) : order_scope. +Notation "\join^c_ ( i 'in' A | P ) F" := + (\big[join/0]_(i in A | P%B) F%O) : order_scope. +Notation "\join^c_ ( i 'in' A ) F" := + (\big[join/0]_(i in A) F%O) : order_scope. + +Notation "\meet^c_ ( i <- r | P ) F" := + (\big[meet/1]_(i <- r | P%B) F%O) : order_scope. +Notation "\meet^c_ ( i <- r ) F" := + (\big[meet/1]_(i <- r) F%O) : order_scope. +Notation "\meet^c_ ( i | P ) F" := + (\big[meet/1]_(i | P%B) F%O) : order_scope. +Notation "\meet^c_ i F" := + (\big[meet/1]_i F%O) : order_scope. +Notation "\meet^c_ ( i : I | P ) F" := + (\big[meet/1]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\meet^c_ ( i : I ) F" := + (\big[meet/1]_(i : I) F%O) (only parsing) : order_scope. +Notation "\meet^c_ ( m <= i < n | P ) F" := + (\big[meet/1]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\meet^c_ ( m <= i < n ) F" := + (\big[meet/1]_(m <= i < n) F%O) : order_scope. +Notation "\meet^c_ ( i < n | P ) F" := + (\big[meet/1]_(i < n | P%B) F%O) : order_scope. +Notation "\meet^c_ ( i < n ) F" := + (\big[meet/1]_(i < n) F%O) : order_scope. +Notation "\meet^c_ ( i 'in' A | P ) F" := + (\big[meet/1]_(i in A | P%B) F%O) : order_scope. +Notation "\meet^c_ ( i 'in' A ) F" := + (\big[meet/1]_(i in A) F%O) : order_scope. + +End ConverseSyntax. + +(**********) +(* THEORY *) +(**********) + +Module Import POrderTheory. +Section POrderTheory. + +Context {disp : unit}. +Local Notation porderType := (porderType disp). +Context {T : porderType}. + +Implicit Types x y : T. + +Lemma geE x y : ge x y = (y <= x). Proof. by []. Qed. +Lemma gtE x y : gt x y = (y < x). Proof. by []. Qed. + +Lemma lexx (x : T) : x <= x. +Proof. by case: T x => ? [? []]. Qed. +Hint Resolve lexx : core. + +Definition le_refl : reflexive le := lexx. +Definition ge_refl : reflexive ge := lexx. +Hint Resolve le_refl : core. + +Lemma le_anti: antisymmetric (<=%O : rel T). +Proof. by case: T => ? [? []]. Qed. + +Lemma ge_anti: antisymmetric (>=%O : rel T). +Proof. by move=> x y /le_anti. Qed. + +Lemma le_trans: transitive (<=%O : rel T). +Proof. by case: T => ? [? []]. Qed. + +Lemma ge_trans: transitive (>=%O : rel T). +Proof. by move=> ? ? ? ? /le_trans; apply. Qed. + +Lemma lt_def x y: (x < y) = (y != x) && (x <= y). +Proof. by case: T x y => ? [? []]. Qed. + +Lemma lt_neqAle x y: (x < y) = (x != y) && (x <= y). +Proof. by rewrite lt_def eq_sym. Qed. + +Lemma ltxx x: x < x = false. +Proof. by rewrite lt_def eqxx. Qed. + +Definition lt_irreflexive : irreflexive lt := ltxx. +Hint Resolve lt_irreflexive : core. + +Definition ltexx := (lexx, ltxx). + +Lemma le_eqVlt x y: (x <= y) = (x == y) || (x < y). +Proof. by rewrite lt_neqAle; case: eqP => //= ->; rewrite lexx. Qed. + +Lemma lt_eqF x y: x < y -> x == y = false. +Proof. by rewrite lt_neqAle => /andP [/negbTE->]. Qed. + +Lemma gt_eqF x y : y < x -> x == y = false. +Proof. by apply: contraTF => /eqP ->; rewrite ltxx. Qed. + +Lemma eq_le x y: (x == y) = (x <= y <= x). +Proof. by apply/eqP/idP => [->|/le_anti]; rewrite ?lexx. Qed. + +Lemma ltW x y: x < y -> x <= y. +Proof. by rewrite le_eqVlt orbC => ->. Qed. + +Lemma lt_le_trans y x z: x < y -> y <= z -> x < z. +Proof. +rewrite !lt_neqAle => /andP [nexy lexy leyz]; rewrite (le_trans lexy) // andbT. +by apply: contraNneq nexy => eqxz; rewrite eqxz eq_le leyz andbT in lexy *. +Qed. + +Lemma lt_trans: transitive (<%O : rel T). +Proof. by move=> y x z le1 /ltW le2; apply/(@lt_le_trans y). Qed. + +Lemma le_lt_trans y x z: x <= y -> y < z -> x < z. +Proof. by rewrite le_eqVlt => /orP [/eqP ->|/lt_trans t /t]. Qed. + +Lemma lt_nsym x y : x < y -> y < x -> False. +Proof. by move=> xy /(lt_trans xy); rewrite ltxx. Qed. + +Lemma lt_asym x y : x < y < x = false. +Proof. by apply/negP => /andP []; apply: lt_nsym. Qed. + +Lemma le_gtF x y: x <= y -> y < x = false. +Proof. +by move=> le_xy; apply/negP => /lt_le_trans /(_ le_xy); rewrite ltxx. +Qed. + +Lemma lt_geF x y : (x < y) -> y <= x = false. +Proof. +by move=> le_xy; apply/negP => /le_lt_trans /(_ le_xy); rewrite ltxx. +Qed. + +Definition lt_gtF x y hxy := le_gtF (@ltW x y hxy). + +Lemma lt_leAnge x y : (x < y) = (x <= y) && ~~ (y <= x). +Proof. +apply/idP/idP => [ltxy|/andP[lexy Nleyx]]; first by rewrite ltW // lt_geF. +by rewrite lt_neqAle lexy andbT; apply: contraNneq Nleyx => ->. +Qed. + +Lemma lt_le_asym x y : x < y <= x = false. +Proof. by rewrite lt_neqAle -andbA -eq_le eq_sym andNb. Qed. + +Lemma le_lt_asym x y : x <= y < x = false. +Proof. by rewrite andbC lt_le_asym. Qed. + +Definition lte_anti := (=^~ eq_le, lt_asym, lt_le_asym, le_lt_asym). + +Lemma lt_sorted_uniq_le (s : seq T) : + sorted lt s = uniq s && sorted le s. +Proof. +case: s => //= n s; elim: s n => //= m s IHs n. +rewrite inE lt_neqAle negb_or IHs -!andbA. +case sn: (n \in s); last do !bool_congr. +rewrite andbF; apply/and5P=> [[ne_nm lenm _ _ le_ms]]; case/negP: ne_nm. +by rewrite eq_le lenm /=; apply: (allP (order_path_min le_trans le_ms)). +Qed. + +Lemma eq_sorted_lt (s1 s2 : seq T) : + sorted lt s1 -> sorted lt s2 -> s1 =i s2 -> s1 = s2. +Proof. by apply: eq_sorted_irr => //; apply: lt_trans. Qed. + +Lemma eq_sorted_le (s1 s2 : seq T) : + sorted le s1 -> sorted le s2 -> perm_eq s1 s2 -> s1 = s2. +Proof. by apply: eq_sorted; [apply: le_trans|apply: le_anti]. Qed. + +Lemma comparable_leNgt x y : x >=< y -> (x <= y) = ~~ (y < x). +Proof. +move=> c_xy; apply/idP/idP => [/le_gtF/negP/negP//|]; rewrite lt_neqAle. +by move: c_xy => /orP [] -> //; rewrite andbT negbK => /eqP ->. +Qed. + +Lemma comparable_ltNge x y : x >=< y -> (x < y) = ~~ (y <= x). +Proof. +move=> c_xy; apply/idP/idP => [/lt_geF/negP/negP//|]. +by rewrite lt_neqAle eq_le; move: c_xy => /orP [] -> //; rewrite andbT. +Qed. + +Lemma comparable_ltgtP x y : x >=< y -> + compare x y (y == x) (x == y) (x >= y) (x <= y) (x > y) (x < y). +Proof. +rewrite />=<%O !le_eqVlt [y == x]eq_sym. +have := (altP (x =P y), (boolP (x < y), boolP (y < x))). +move=> [[->//|neq_xy /=] [[] xy [] //=]] ; do ?by rewrite ?ltxx; constructor. + by rewrite ltxx in xy. +by rewrite le_gtF // ltW. +Qed. + +Lemma comparable_leP x y : x >=< y -> le_xor_gt x y (x <= y) (y < x). +Proof. by move=> /comparable_ltgtP [?|?|->]; constructor; rewrite // ltW. Qed. + +Lemma comparable_ltP x y : x >=< y -> lt_xor_ge x y (y <= x) (x < y). +Proof. by move=> /comparable_ltgtP [?|?|->]; constructor; rewrite // ltW. Qed. + +Lemma comparable_sym x y : (y >=< x) = (x >=< y). +Proof. by rewrite /comparable orbC. Qed. + +Lemma comparablexx x : x >=< x. +Proof. by rewrite /comparable lexx. Qed. + +Lemma incomparable_eqF x y : (x >< y) -> (x == y) = false. +Proof. by apply: contraNF => /eqP ->; rewrite comparablexx. Qed. + +Lemma incomparable_leF x y : (x >< y) -> (x <= y) = false. +Proof. by apply: contraNF; rewrite /comparable => ->. Qed. + +Lemma incomparable_ltF x y : (x >< y) -> (x < y) = false. +Proof. by rewrite lt_neqAle => /incomparable_leF ->; rewrite andbF. Qed. + +Lemma comparableP x y : incompare x y + (y == x) (x == y) (x >= y) (x <= y) (x > y) (x < y) + (y >=< x) (x >=< y). +Proof. +rewrite ![y >=< _]comparable_sym; have [c_xy|i_xy] := boolP (x >=< y). + by case: (comparable_ltgtP c_xy) => ?; constructor. +by rewrite ?incomparable_eqF ?incomparable_leF ?incomparable_ltF // + 1?comparable_sym //; constructor. +Qed. + +Lemma le_comparable (x y : T) : x <= y -> x >=< y. +Proof. by case: comparableP. Qed. + +Lemma lt_comparable (x y : T) : x < y -> x >=< y. +Proof. by case: comparableP. Qed. + +Lemma ge_comparable (x y : T) : y <= x -> x >=< y. +Proof. by case: comparableP. Qed. + +Lemma gt_comparable (x y : T) : y < x -> x >=< y. +Proof. by case: comparableP. Qed. + +Lemma leifP x y C : reflect (x <= y ?= iff C) (if C then x == y else x < y). +Proof. +rewrite /leif le_eqVlt; apply: (iffP idP)=> [|[]]. + by case: C => [/eqP->|lxy]; rewrite ?eqxx // lxy lt_eqF. +by move=> /orP[/eqP->|lxy] <-; rewrite ?eqxx // lt_eqF. +Qed. + +Lemma leif_refl x C : reflect (x <= x ?= iff C) C. +Proof. by apply: (iffP idP) => [-> | <-] //; split; rewrite ?eqxx. Qed. + +Lemma leif_trans x1 x2 x3 C12 C23 : + x1 <= x2 ?= iff C12 -> x2 <= x3 ?= iff C23 -> x1 <= x3 ?= iff C12 && C23. +Proof. +move=> ltx12 ltx23; apply/leifP; rewrite -ltx12. +case eqx12: (x1 == x2). + by rewrite (eqP eqx12) lt_neqAle !ltx23 andbT; case C23. +by rewrite (@lt_le_trans x2) ?ltx23 // lt_neqAle eqx12 ltx12. +Qed. + +Lemma leif_le x y : x <= y -> x <= y ?= iff (x >= y). +Proof. by move=> lexy; split=> //; rewrite eq_le lexy. Qed. + +Lemma leif_eq x y : x <= y -> x <= y ?= iff (x == y). +Proof. by []. Qed. + +Lemma ge_leif x y C : x <= y ?= iff C -> (y <= x) = C. +Proof. by case=> le_xy; rewrite eq_le le_xy. Qed. + +Lemma lt_leif x y C : x <= y ?= iff C -> (x < y) = ~~ C. +Proof. by move=> le_xy; rewrite lt_neqAle !le_xy andbT. Qed. + +Lemma ltNleif x y C : x <= y ?= iff ~~ C -> (x < y) = C. +Proof. by move=> /lt_leif; rewrite negbK. Qed. + +Lemma eq_leif x y C : x <= y ?= iff C -> (x == y) = C. +Proof. by move=> /leifP; case: C comparableP => [] []. Qed. + +Lemma eqTleif x y C : x <= y ?= iff C -> C -> x = y. +Proof. by move=> /eq_leif<-/eqP. Qed. + +Lemma mono_in_leif (A : {pred T}) (f : T -> T) C : + {in A &, {mono f : x y / x <= y}} -> + {in A &, forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C)}. +Proof. by move=> mf x y Ax Ay; rewrite /leif !eq_le !mf. Qed. + +Lemma mono_leif (f : T -> T) C : + {mono f : x y / x <= y} -> + forall x y, (f x <= f y ?= iff C) = (x <= y ?= iff C). +Proof. by move=> mf x y; rewrite /leif !eq_le !mf. Qed. + +Lemma nmono_in_leif (A : {pred T}) (f : T -> T) C : + {in A &, {mono f : x y /~ x <= y}} -> + {in A &, forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C)}. +Proof. by move=> mf x y Ax Ay; rewrite /leif !eq_le !mf. Qed. + +Lemma nmono_leif (f : T -> T) C : + {mono f : x y /~ x <= y} -> + forall x y, (f x <= f y ?= iff C) = (y <= x ?= iff C). +Proof. by move=> mf x y; rewrite /leif !eq_le !mf. Qed. + +End POrderTheory. +Section POrderMonotonyTheory. + +Context {disp disp' : unit}. +Context {T : porderType disp} {T' : porderType disp'}. +Implicit Types (m n p : nat) (x y z : T) (u v w : T'). +Variables (D D' : {pred T}) (f : T -> T'). + +Hint Resolve lexx lt_neqAle (@le_anti _ T) (@le_anti _ T') lt_def : core. + +Let ge_antiT : antisymmetric (>=%O : rel T). +Proof. by move=> ? ? /le_anti. Qed. + +Lemma ltW_homo : {homo f : x y / x < y} -> {homo f : x y / x <= y}. +Proof. exact: homoW. Qed. + +Lemma ltW_nhomo : {homo f : x y /~ x < y} -> {homo f : x y /~ x <= y}. +Proof. exact: homoW. Qed. + +Lemma inj_homo_lt : + injective f -> {homo f : x y / x <= y} -> {homo f : x y / x < y}. +Proof. exact: inj_homo. Qed. + +Lemma inj_nhomo_lt : + injective f -> {homo f : x y /~ x <= y} -> {homo f : x y /~ x < y}. +Proof. exact: inj_homo. Qed. + +Lemma inc_inj : {mono f : x y / x <= y} -> injective f. +Proof. exact: mono_inj. Qed. + +Lemma dec_inj : {mono f : x y /~ x <= y} -> injective f. +Proof. exact: mono_inj. Qed. + +Lemma leW_mono : {mono f : x y / x <= y} -> {mono f : x y / x < y}. +Proof. exact: anti_mono. Qed. + +Lemma leW_nmono : {mono f : x y /~ x <= y} -> {mono f : x y /~ x < y}. +Proof. exact: anti_mono. Qed. + +(* Monotony in D D' *) +Lemma ltW_homo_in : + {in D & D', {homo f : x y / x < y}} -> {in D & D', {homo f : x y / x <= y}}. +Proof. exact: homoW_in. Qed. + +Lemma ltW_nhomo_in : + {in D & D', {homo f : x y /~ x < y}} -> {in D & D', {homo f : x y /~ x <= y}}. +Proof. exact: homoW_in. Qed. + +Lemma inj_homo_lt_in : + {in D & D', injective f} -> {in D & D', {homo f : x y / x <= y}} -> + {in D & D', {homo f : x y / x < y}}. +Proof. exact: inj_homo_in. Qed. + +Lemma inj_nhomo_lt_in : + {in D & D', injective f} -> {in D & D', {homo f : x y /~ x <= y}} -> + {in D & D', {homo f : x y /~ x < y}}. +Proof. exact: inj_homo_in. Qed. + +Lemma inc_inj_in : {in D &, {mono f : x y / x <= y}} -> + {in D &, injective f}. +Proof. exact: mono_inj_in. Qed. + +Lemma dec_inj_in : + {in D &, {mono f : x y /~ x <= y}} -> {in D &, injective f}. +Proof. exact: mono_inj_in. Qed. + +Lemma leW_mono_in : + {in D &, {mono f : x y / x <= y}} -> {in D &, {mono f : x y / x < y}}. +Proof. exact: anti_mono_in. Qed. + +Lemma leW_nmono_in : + {in D &, {mono f : x y /~ x <= y}} -> {in D &, {mono f : x y /~ x < y}}. +Proof. exact: anti_mono_in. Qed. + +End POrderMonotonyTheory. + +End POrderTheory. + +Hint Resolve lexx le_refl ltxx lt_irreflexive ltW lt_eqF : core. + +Arguments leifP {disp T x y C}. +Arguments leif_refl {disp T x C}. +Arguments mono_in_leif [disp T A f C]. +Arguments nmono_in_leif [disp T A f C]. +Arguments mono_leif [disp T f C]. +Arguments nmono_leif [disp T f C]. + +Module Import ConversePOrder. +Section ConversePOrder. +Canonical converse_eqType (T : eqType) := EqType T [eqMixin of T^c]. +Canonical converse_choiceType (T : choiceType) := [choiceType of T^c]. + +Context {disp : unit}. +Local Notation porderType := (porderType disp). +Variable T : porderType. + +Definition converse_le (x y : T) := (y <= x). +Definition converse_lt (x y : T) := (y < x). + +Lemma converse_lt_def (x y : T) : + converse_lt x y = (y != x) && (converse_le x y). +Proof. by apply: lt_neqAle. Qed. + +Fact converse_le_anti : antisymmetric converse_le. +Proof. by move=> x y /andP [xy yx]; apply/le_anti/andP; split. Qed. + +Definition converse_porderMixin := + LePOrderMixin converse_lt_def (lexx : reflexive converse_le) converse_le_anti + (fun y z x zy yx => @le_trans _ _ y x z yx zy). +Canonical converse_porderType := + POrderType (converse_display disp) (T^c) converse_porderMixin. + +End ConversePOrder. +End ConversePOrder. + +Module Import ConverseDistrLattice. +Section ConverseDistrLattice. +Context {disp : unit}. +Local Notation distrLatticeType := (distrLatticeType disp). + +Variable L : distrLatticeType. +Implicit Types (x y : L). + +Lemma meetC : commutative (@meet _ L). Proof. by case: L => [?[? ?[]]]. Qed. +Lemma joinC : commutative (@join _ L). Proof. by case: L => [?[? ?[]]]. Qed. + +Lemma meetA : associative (@meet _ L). Proof. by case: L => [?[? ?[]]]. Qed. +Lemma joinA : associative (@join _ L). Proof. by case: L => [?[? ?[]]]. Qed. + +Lemma joinKI y x : x `&` (x `|` y) = x. +Proof. by case: L x y => [?[? ?[]]]. Qed. +Lemma meetKU y x : x `|` (x `&` y) = x. +Proof. by case: L x y => [?[? ?[]]]. Qed. + +Lemma joinKIC y x : x `&` (y `|` x) = x. Proof. by rewrite joinC joinKI. Qed. +Lemma meetKUC y x : x `|` (y `&` x) = x. Proof. by rewrite meetC meetKU. Qed. + +Lemma meetUK x y : (x `&` y) `|` y = y. +Proof. by rewrite joinC meetC meetKU. Qed. +Lemma joinIK x y : (x `|` y) `&` y = y. +Proof. by rewrite joinC meetC joinKI. Qed. + +Lemma meetUKC x y : (y `&` x) `|` y = y. Proof. by rewrite meetC meetUK. Qed. +Lemma joinIKC x y : (y `|` x) `&` y = y. Proof. by rewrite joinC joinIK. Qed. + +Lemma leEmeet x y : (x <= y) = (x `&` y == x). +Proof. by case: L x y => [?[? ?[]]]. Qed. + +Lemma leEjoin x y : (x <= y) = (x `|` y == y). +Proof. by rewrite leEmeet; apply/eqP/eqP => <-; rewrite (joinKI, meetUK). Qed. + +Lemma meetUl : left_distributive (@meet _ L) (@join _ L). +Proof. by case: L => [?[? ?[]]]. Qed. + +Lemma meetUr : right_distributive (@meet _ L) (@join _ L). +Proof. by move=> x y z; rewrite meetC meetUl ![_ `&` x]meetC. Qed. + +Lemma joinIl : left_distributive (@join _ L) (@meet _ L). +Proof. by move=> x y z; rewrite meetUr joinIK meetUl -joinA meetUKC. Qed. + +Fact converse_leEmeet (x y : L^c) : (x <= y) = (x `|` y == x). +Proof. by rewrite [LHS]leEjoin joinC. Qed. + +Definition converse_distrLatticeMixin := + @DistrLatticeMixin _ [porderType of L^c] _ _ joinC meetC + joinA meetA meetKU joinKI converse_leEmeet joinIl. +Canonical converse_distrLatticeType := + DistrLatticeType L^c converse_distrLatticeMixin. +End ConverseDistrLattice. +End ConverseDistrLattice. + +Module Import DistrLatticeTheoryMeet. +Section DistrLatticeTheoryMeet. +Context {disp : unit}. +Local Notation distrLatticeType := (distrLatticeType disp). +Context {L : distrLatticeType}. +Implicit Types (x y : L). + +(* lattice theory *) +Lemma meetAC : right_commutative (@meet _ L). +Proof. by move=> x y z; rewrite -!meetA [X in _ `&` X]meetC. Qed. +Lemma meetCA : left_commutative (@meet _ L). +Proof. by move=> x y z; rewrite !meetA [X in X `&` _]meetC. Qed. +Lemma meetACA : interchange (@meet _ L) (@meet _ L). +Proof. by move=> x y z t; rewrite !meetA [X in X `&` _]meetAC. Qed. + +Lemma meetxx x : x `&` x = x. +Proof. by rewrite -[X in _ `&` X](meetKU x) joinKI. Qed. + +Lemma meetKI y x : x `&` (x `&` y) = x `&` y. +Proof. by rewrite meetA meetxx. Qed. +Lemma meetIK y x : (x `&` y) `&` y = x `&` y. +Proof. by rewrite -meetA meetxx. Qed. +Lemma meetKIC y x : x `&` (y `&` x) = x `&` y. +Proof. by rewrite meetC meetIK meetC. Qed. +Lemma meetIKC y x : y `&` x `&` y = x `&` y. +Proof. by rewrite meetAC meetC meetxx. Qed. + +(* interaction with order *) + +Lemma lexI x y z : (x <= y `&` z) = (x <= y) && (x <= z). +Proof. +rewrite !leEmeet; apply/eqP/andP => [<-|[/eqP<- /eqP<-]]. + by rewrite meetA meetIK eqxx -meetA meetACA meetxx meetAC eqxx. +by rewrite -[X in X `&` _]meetA meetIK meetA. +Qed. + +Lemma leIxl x y z : y <= x -> y `&` z <= x. +Proof. by rewrite !leEmeet meetAC => /eqP ->. Qed. + +Lemma leIxr x y z : z <= x -> y `&` z <= x. +Proof. by rewrite !leEmeet -meetA => /eqP ->. Qed. + +Lemma leIx2 x y z : (y <= x) || (z <= x) -> y `&` z <= x. +Proof. by case/orP => [/leIxl|/leIxr]. Qed. + +Lemma leIr x y : y `&` x <= x. +Proof. by rewrite leIx2 ?lexx ?orbT. Qed. + +Lemma leIl x y : x `&` y <= x. +Proof. by rewrite leIx2 ?lexx ?orbT. Qed. + +Lemma meet_idPl {x y} : reflect (x `&` y = x) (x <= y). +Proof. by rewrite leEmeet; apply/eqP. Qed. +Lemma meet_idPr {x y} : reflect (y `&` x = x) (x <= y). +Proof. by rewrite meetC; apply/meet_idPl. Qed. + +Lemma meet_l x y : x <= y -> x `&` y = x. Proof. exact/meet_idPl. Qed. +Lemma meet_r x y : y <= x -> x `&` y = y. Proof. exact/meet_idPr. Qed. + +Lemma leIidl x y : (x <= x `&` y) = (x <= y). +Proof. by rewrite !leEmeet meetKI. Qed. +Lemma leIidr x y : (x <= y `&` x) = (x <= y). +Proof. by rewrite !leEmeet meetKIC. Qed. + +Lemma eq_meetl x y : (x `&` y == x) = (x <= y). +Proof. by apply/esym/leEmeet. Qed. + +Lemma eq_meetr x y : (x `&` y == y) = (y <= x). +Proof. by rewrite meetC eq_meetl. Qed. + +Lemma leI2 x y z t : x <= z -> y <= t -> x `&` y <= z `&` t. +Proof. by move=> xz yt; rewrite lexI !leIx2 ?xz ?yt ?orbT //. Qed. + +End DistrLatticeTheoryMeet. +End DistrLatticeTheoryMeet. + +Module Import DistrLatticeTheoryJoin. +Section DistrLatticeTheoryJoin. +Context {disp : unit}. +Local Notation distrLatticeType := (distrLatticeType disp). +Context {L : distrLatticeType}. +Implicit Types (x y : L). + +(* lattice theory *) +Lemma joinAC : right_commutative (@join _ L). +Proof. exact: (@meetAC _ [distrLatticeType of L^c]). Qed. +Lemma joinCA : left_commutative (@join _ L). +Proof. exact: (@meetCA _ [distrLatticeType of L^c]). Qed. +Lemma joinACA : interchange (@join _ L) (@join _ L). +Proof. exact: (@meetACA _ [distrLatticeType of L^c]). Qed. + +Lemma joinxx x : x `|` x = x. +Proof. exact: (@meetxx _ [distrLatticeType of L^c]). Qed. + +Lemma joinKU y x : x `|` (x `|` y) = x `|` y. +Proof. exact: (@meetKI _ [distrLatticeType of L^c]). Qed. +Lemma joinUK y x : (x `|` y) `|` y = x `|` y. +Proof. exact: (@meetIK _ [distrLatticeType of L^c]). Qed. +Lemma joinKUC y x : x `|` (y `|` x) = x `|` y. +Proof. exact: (@meetKIC _ [distrLatticeType of L^c]). Qed. +Lemma joinUKC y x : y `|` x `|` y = x `|` y. +Proof. exact: (@meetIKC _ [distrLatticeType of L^c]). Qed. + +(* interaction with order *) +Lemma leUx x y z : (x `|` y <= z) = (x <= z) && (y <= z). +Proof. exact: (@lexI _ [distrLatticeType of L^c]). Qed. +Lemma lexUl x y z : x <= y -> x <= y `|` z. +Proof. exact: (@leIxl _ [distrLatticeType of L^c]). Qed. +Lemma lexUr x y z : x <= z -> x <= y `|` z. +Proof. exact: (@leIxr _ [distrLatticeType of L^c]). Qed. +Lemma lexU2 x y z : (x <= y) || (x <= z) -> x <= y `|` z. +Proof. exact: (@leIx2 _ [distrLatticeType of L^c]). Qed. + +Lemma leUr x y : x <= y `|` x. +Proof. exact: (@leIr _ [distrLatticeType of L^c]). Qed. +Lemma leUl x y : x <= x `|` y. +Proof. exact: (@leIl _ [distrLatticeType of L^c]). Qed. + +Lemma join_idPl {x y} : reflect (x `|` y = y) (x <= y). +Proof. exact: (@meet_idPr _ [distrLatticeType of L^c]). Qed. +Lemma join_idPr {x y} : reflect (y `|` x = y) (x <= y). +Proof. exact: (@meet_idPl _ [distrLatticeType of L^c]). Qed. + +Lemma join_l x y : y <= x -> x `|` y = x. Proof. exact/join_idPr. Qed. +Lemma join_r x y : x <= y -> x `|` y = y. Proof. exact/join_idPl. Qed. + +Lemma leUidl x y : (x `|` y <= y) = (x <= y). +Proof. exact: (@leIidr _ [distrLatticeType of L^c]). Qed. +Lemma leUidr x y : (y `|` x <= y) = (x <= y). +Proof. exact: (@leIidl _ [distrLatticeType of L^c]). Qed. + +Lemma eq_joinl x y : (x `|` y == x) = (y <= x). +Proof. exact: (@eq_meetl _ [distrLatticeType of L^c]). Qed. +Lemma eq_joinr x y : (x `|` y == y) = (x <= y). +Proof. exact: (@eq_meetr _ [distrLatticeType of L^c]). Qed. + +Lemma leU2 x y z t : x <= z -> y <= t -> x `|` y <= z `|` t. +Proof. exact: (@leI2 _ [distrLatticeType of L^c]). Qed. + +(* Distributive lattice theory *) +Lemma joinIr : right_distributive (@join _ L) (@meet _ L). +Proof. exact: (@meetUr _ [distrLatticeType of L^c]). Qed. + +Lemma lcomparableP x y : incomparel x y + (y == x) (x == y) (x >= y) (x <= y) (x > y) (x < y) + (y >=< x) (x >=< y) (y `&` x) (x `&` y) (y `|` x) (x `|` y). +Proof. +by case: (comparableP x) => [hxy|hxy|hxy|->]; do 1?have hxy' := ltW hxy; + rewrite ?(meetxx, joinxx, meetC y, joinC y) + ?(meet_idPl hxy', meet_idPr hxy', join_idPl hxy', join_idPr hxy'); + constructor. +Qed. + +Lemma lcomparable_ltgtP x y : x >=< y -> + comparel x y (y == x) (x == y) (x >= y) (x <= y) (x > y) (x < y) + (y `&` x) (x `&` y) (y `|` x) (x `|` y). +Proof. by case: (lcomparableP x) => // *; constructor. Qed. + +Lemma lcomparable_leP x y : x >=< y -> + lel_xor_gt x y (x <= y) (y < x) (y `&` x) (x `&` y) (y `|` x) (x `|` y). +Proof. by move/lcomparable_ltgtP => [/ltW xy|xy|->]; constructor. Qed. + +Lemma lcomparable_ltP x y : x >=< y -> + ltl_xor_ge x y (y <= x) (x < y) (y `&` x) (x `&` y) (y `|` x) (x `|` y). +Proof. by move=> /lcomparable_ltgtP [xy|/ltW xy|->]; constructor. Qed. + +End DistrLatticeTheoryJoin. +End DistrLatticeTheoryJoin. + +Module Import TotalTheory. +Section TotalTheory. +Context {disp : unit}. +Local Notation orderType := (orderType disp). +Context {T : orderType}. +Implicit Types (x y z t : T). + +Lemma le_total : total (<=%O : rel T). Proof. by case: T => [? [?]]. Qed. +Hint Resolve le_total : core. + +Lemma ge_total : total (>=%O : rel T). +Proof. by move=> ? ?; apply: le_total. Qed. +Hint Resolve ge_total : core. + +Lemma comparableT x y : x >=< y. Proof. exact: le_total. Qed. +Hint Resolve comparableT : core. + +Lemma sort_le_sorted (s : seq T) : sorted <=%O (sort <=%O s). +Proof. exact: sort_sorted. Qed. + +Lemma sort_lt_sorted (s : seq T) : sorted lt (sort le s) = uniq s. +Proof. by rewrite lt_sorted_uniq_le sort_uniq sort_le_sorted andbT. Qed. + +Lemma sort_le_id (s : seq T) : sorted le s -> sort le s = s. +Proof. +by move=> ss; apply: eq_sorted_le; rewrite ?sort_le_sorted // perm_sort. +Qed. + +Lemma leNgt x y : (x <= y) = ~~ (y < x). Proof. exact: comparable_leNgt. Qed. + +Lemma ltNge x y : (x < y) = ~~ (y <= x). Proof. exact: comparable_ltNge. Qed. + +Definition ltgtP x y := + DistrLatticeTheoryJoin.lcomparable_ltgtP (comparableT x y). +Definition leP x y := DistrLatticeTheoryJoin.lcomparable_leP (comparableT x y). +Definition ltP x y := DistrLatticeTheoryJoin.lcomparable_ltP (comparableT x y). + +Lemma wlog_le P : + (forall x y, P y x -> P x y) -> (forall x y, x <= y -> P x y) -> + forall x y, P x y. +Proof. by move=> sP hP x y; case: (leP x y) => [| /ltW] /hP // /sP. Qed. + +Lemma wlog_lt P : + (forall x, P x x) -> + (forall x y, (P y x -> P x y)) -> (forall x y, x < y -> P x y) -> + forall x y, P x y. +Proof. by move=> rP sP hP x y; case: (ltgtP x y) => [||->] // /hP // /sP. Qed. + +Lemma neq_lt x y : (x != y) = (x < y) || (y < x). Proof. by case: ltgtP. Qed. + +Lemma lt_total x y : x != y -> (x < y) || (y < x). Proof. by case: ltgtP. Qed. + +Lemma eq_leLR x y z t : + (x <= y -> z <= t) -> (y < x -> t < z) -> (x <= y) = (z <= t). +Proof. by rewrite !ltNge => ? /contraTT ?; apply/idP/idP. Qed. + +Lemma eq_leRL x y z t : + (x <= y -> z <= t) -> (y < x -> t < z) -> (z <= t) = (x <= y). +Proof. by move=> *; symmetry; apply: eq_leLR. Qed. + +Lemma eq_ltLR x y z t : + (x < y -> z < t) -> (y <= x -> t <= z) -> (x < y) = (z < t). +Proof. by rewrite !leNgt => ? /contraTT ?; apply/idP/idP. Qed. + +Lemma eq_ltRL x y z t : + (x < y -> z < t) -> (y <= x -> t <= z) -> (z < t) = (x < y). +Proof. by move=> *; symmetry; apply: eq_ltLR. Qed. + +(* interaction with lattice operations *) + +Lemma leIx x y z : (meet y z <= x) = (y <= x) || (z <= x). +Proof. +by case: (leP y z) => hyz; case: leP => ?; + rewrite ?(orbT, orbF) //=; apply/esym/negbTE; + rewrite -ltNge ?(lt_le_trans _ hyz) ?(lt_trans _ hyz). +Qed. + +Lemma lexU x y z : (x <= join y z) = (x <= y) || (x <= z). +Proof. +by case: (leP y z) => hyz; case: leP => ?; + rewrite ?(orbT, orbF) //=; apply/esym/negbTE; + rewrite -ltNge ?(le_lt_trans hyz) ?(lt_trans hyz). +Qed. + +Lemma ltxI x y z : (x < meet y z) = (x < y) && (x < z). +Proof. by rewrite !ltNge leIx negb_or. Qed. + +Lemma ltIx x y z : (meet y z < x) = (y < x) || (z < x). +Proof. by rewrite !ltNge lexI negb_and. Qed. + +Lemma ltxU x y z : (x < join y z) = (x < y) || (x < z). +Proof. by rewrite !ltNge leUx negb_and. Qed. + +Lemma ltUx x y z : (join y z < x) = (y < x) && (z < x). +Proof. by rewrite !ltNge lexU negb_or. Qed. + +Definition ltexI := (@lexI _ T, ltxI). +Definition lteIx := (leIx, ltIx). +Definition ltexU := (lexU, ltxU). +Definition lteUx := (@leUx _ T, ltUx). + +Section ArgExtremum. + +Context (I : finType) (i0 : I) (P : {pred I}) (F : I -> T) (Pi0 : P i0). + +Definition arg_minnP := arg_minP. +Definition arg_maxnP := arg_maxP. + +Lemma arg_minP: extremum_spec <=%O P F (arg_min i0 P F). +Proof. by apply: extremumP => //; apply: le_trans. Qed. + +Lemma arg_maxP: extremum_spec >=%O P F (arg_max i0 P F). +Proof. by apply: extremumP => //; [apply: ge_refl | apply: ge_trans]. Qed. + +End ArgExtremum. + +End TotalTheory. +Section TotalMonotonyTheory. + +Context {disp : unit} {disp' : unit}. +Context {T : orderType disp} {T' : porderType disp'}. +Variables (D : {pred T}) (f : T -> T'). +Implicit Types (x y z : T) (u v w : T'). + +Hint Resolve (@le_anti _ T) (@le_anti _ T') (@lt_neqAle _ T) : core. +Hint Resolve (@lt_neqAle _ T') (@lt_def _ T) (@le_total _ T) : core. + +Lemma le_mono : {homo f : x y / x < y} -> {mono f : x y / x <= y}. +Proof. exact: total_homo_mono. Qed. + +Lemma le_nmono : {homo f : x y /~ x < y} -> {mono f : x y /~ x <= y}. +Proof. exact: total_homo_mono. Qed. + +Lemma le_mono_in : + {in D &, {homo f : x y / x < y}} -> {in D &, {mono f : x y / x <= y}}. +Proof. exact: total_homo_mono_in. Qed. + +Lemma le_nmono_in : + {in D &, {homo f : x y /~ x < y}} -> {in D &, {mono f : x y /~ x <= y}}. +Proof. exact: total_homo_mono_in. Qed. + +End TotalMonotonyTheory. +End TotalTheory. + +Module Import BDistrLatticeTheory. +Section BDistrLatticeTheory. +Context {disp : unit}. +Local Notation bDistrLatticeType := (bDistrLatticeType disp). +Context {L : bDistrLatticeType}. +Implicit Types (I : finType) (T : eqType) (x y z : L). +Local Notation "0" := bottom. + +(* Distributive lattice theory with 0 & 1*) +Lemma le0x x : 0 <= x. Proof. by case: L x => [?[? ?[]]]. Qed. +Hint Resolve le0x : core. + +Lemma lex0 x : (x <= 0) = (x == 0). +Proof. by rewrite le_eqVlt (le_gtF (le0x _)) orbF. Qed. + +Lemma ltx0 x : (x < 0) = false. +Proof. by rewrite lt_neqAle lex0 andNb. Qed. + +Lemma lt0x x : (0 < x) = (x != 0). +Proof. by rewrite lt_neqAle le0x andbT eq_sym. Qed. + +Lemma meet0x : left_zero 0 (@meet _ L). +Proof. by move=> x; apply/eqP; rewrite -leEmeet. Qed. + +Lemma meetx0 : right_zero 0 (@meet _ L). +Proof. by move=> x; rewrite meetC meet0x. Qed. + +Lemma join0x : left_id 0 (@join _ L). +Proof. by move=> x; apply/eqP; rewrite -leEjoin. Qed. + +Lemma joinx0 : right_id 0 (@join _ L). +Proof. by move=> x; rewrite joinC join0x. Qed. + +Lemma leU2l_le y t x z : x `&` t = 0 -> x `|` y <= z `|` t -> x <= z. +Proof. +by move=> xIt0 /(leI2 (lexx x)); rewrite joinKI meetUr xIt0 joinx0 leIidl. +Qed. + +Lemma leU2r_le y t x z : x `&` t = 0 -> y `|` x <= t `|` z -> x <= z. +Proof. by rewrite joinC [_ `|` z]joinC => /leU2l_le H /H. Qed. + +Lemma disjoint_lexUl z x y : x `&` z = 0 -> (x <= y `|` z) = (x <= y). +Proof. +move=> xz0; apply/idP/idP=> xy; last by rewrite lexU2 ?xy. +by apply: (@leU2l_le x z); rewrite ?joinxx. +Qed. + +Lemma disjoint_lexUr z x y : x `&` z = 0 -> (x <= z `|` y) = (x <= y). +Proof. by move=> xz0; rewrite joinC; rewrite disjoint_lexUl. Qed. + +Lemma leU2E x y z t : x `&` t = 0 -> y `&` z = 0 -> + (x `|` y <= z `|` t) = (x <= z) && (y <= t). +Proof. +move=> dxt dyz; apply/idP/andP; last by case=> ? ?; exact: leU2. +by move=> lexyzt; rewrite (leU2l_le _ lexyzt) // (leU2r_le _ lexyzt). +Qed. + +Lemma join_eq0 x y : (x `|` y == 0) = (x == 0) && (y == 0). +Proof. +apply/idP/idP; last by move=> /andP [/eqP-> /eqP->]; rewrite joinx0. +by move=> /eqP xUy0; rewrite -!lex0 -!xUy0 ?leUl ?leUr. +Qed. + +Variant eq0_xor_gt0 x : bool -> bool -> Set := + Eq0NotPOs : x = 0 -> eq0_xor_gt0 x true false + | POsNotEq0 : 0 < x -> eq0_xor_gt0 x false true. + +Lemma posxP x : eq0_xor_gt0 x (x == 0) (0 < x). +Proof. by rewrite lt0x; have [] := altP eqP; constructor; rewrite ?lt0x. Qed. + +Canonical join_monoid := Monoid.Law (@joinA _ _) join0x joinx0. +Canonical join_comoid := Monoid.ComLaw (@joinC _ _). + +Lemma join_sup I (j : I) (P : {pred I}) (F : I -> L) : + P j -> F j <= \join_(i | P i) F i. +Proof. by move=> Pj; rewrite (bigD1 j) //= lexU2 ?lexx. Qed. + +Lemma join_min I (j : I) (l : L) (P : {pred I}) (F : I -> L) : + P j -> l <= F j -> l <= \join_(i | P i) F i. +Proof. by move=> Pj /le_trans -> //; rewrite join_sup. Qed. + +Lemma joinsP I (u : L) (P : {pred I}) (F : I -> L) : + reflect (forall i : I, P i -> F i <= u) (\join_(i | P i) F i <= u). +Proof. +have -> : \join_(i | P i) F i <= u = (\big[andb/true]_(i | P i) (F i <= u)). + by elim/big_rec2: _ => [|i y b Pi <-]; rewrite ?le0x ?leUx. +rewrite big_all_cond; apply: (iffP allP) => /= H i; +have := H i _; rewrite mem_index_enum; last by move/implyP->. +by move=> /(_ isT) /implyP. +Qed. + +Lemma join_sup_seq T (r : seq T) (P : {pred T}) (F : T -> L) (x : T) : + x \in r -> P x -> F x <= \join_(i <- r | P i) F i. +Proof. by move=> /seq_tnthP[j->] Px; rewrite big_tnth join_sup. Qed. + +Lemma join_min_seq T (r : seq T) (P : {pred T}) (F : T -> L) (x : T) (l : L) : + x \in r -> P x -> l <= F x -> l <= \join_(x <- r | P x) F x. +Proof. by move=> /seq_tnthP[j->] Px; rewrite big_tnth; apply: join_min. Qed. + +Lemma joinsP_seq T (r : seq T) (P : {pred T}) (F : T -> L) (u : L) : + reflect (forall x : T, x \in r -> P x -> F x <= u) + (\join_(x <- r | P x) F x <= u). +Proof. +rewrite big_tnth; apply: (iffP (joinsP _ _ _)) => /= F_le. + by move=> x /seq_tnthP[i ->]; apply: F_le. +by move=> i /F_le->//; rewrite mem_tnth. +Qed. + +Lemma le_joins I (A B : {set I}) (F : I -> L) : + A \subset B -> \join_(i in A) F i <= \join_(i in B) F i. +Proof. +move=> AsubB; rewrite -(setID B A). +rewrite [X in _ <= X](eq_bigl [predU B :&: A & B :\: A]); last first. + by move=> i; rewrite !inE. +rewrite bigU //=; last by rewrite -setI_eq0 setDE setIACA setICr setI0. +by rewrite lexU2 // (setIidPr _) // lexx. +Qed. + +Lemma joins_setU I (A B : {set I}) (F : I -> L) : + \join_(i in (A :|: B)) F i = \join_(i in A) F i `|` \join_(i in B) F i. +Proof. +apply/eqP; rewrite eq_le leUx !le_joins ?subsetUl ?subsetUr ?andbT //. +apply/joinsP => i; rewrite inE; move=> /orP. +by case=> ?; rewrite lexU2 //; [rewrite join_sup|rewrite orbC join_sup]. +Qed. + +Lemma join_seq I (r : seq I) (F : I -> L) : + \join_(i <- r) F i = \join_(i in r) F i. +Proof. +rewrite [RHS](eq_bigl (mem [set i | i \in r])); last by move=> i; rewrite !inE. +elim: r => [|i r ihr]; first by rewrite big_nil big1 // => i; rewrite ?inE. +rewrite big_cons {}ihr; apply/eqP; rewrite eq_le set_cons. +rewrite leUx join_sup ?inE ?eqxx // le_joins //= ?subsetUr //. +apply/joinsP => j; rewrite !inE => /predU1P [->|jr]; rewrite ?lexU2 ?lexx //. +by rewrite join_sup ?orbT ?inE. +Qed. + +Lemma joins_disjoint I (d : L) (P : {pred I}) (F : I -> L) : + (forall i : I, P i -> d `&` F i = 0) -> d `&` \join_(i | P i) F i = 0. +Proof. +move=> d_Fi_disj; have : \big[andb/true]_(i | P i) (d `&` F i == 0). + rewrite big_all_cond; apply/allP => i _ /=. + by apply/implyP => /d_Fi_disj ->. +elim/big_rec2: _ => [|i y]; first by rewrite meetx0. +case; rewrite (andbF, andbT) // => Pi /(_ isT) dy /eqP dFi. +by rewrite meetUr dy dFi joinxx. +Qed. + +End BDistrLatticeTheory. +End BDistrLatticeTheory. + +Module Import ConverseTBDistrLattice. +Section ConverseTBDistrLattice. +Context {disp : unit}. +Local Notation tbDistrLatticeType := (tbDistrLatticeType disp). +Context {L : tbDistrLatticeType}. + +Lemma lex1 (x : L) : x <= top. Proof. by case: L x => [?[? ?[]]]. Qed. + +Definition converse_bDistrLatticeMixin := + @BDistrLatticeMixin _ [distrLatticeType of L^c] top lex1. +Canonical converse_bDistrLatticeType := + BDistrLatticeType L^c converse_bDistrLatticeMixin. + +Definition converse_tbDistrLatticeMixin := + @TBDistrLatticeMixin _ [distrLatticeType of L^c] (bottom : L) (@le0x _ L). +Canonical converse_tbDIstrLatticeType := + TBDistrLatticeType L^c converse_tbDistrLatticeMixin. + +End ConverseTBDistrLattice. +End ConverseTBDistrLattice. + +Module Import TBDistrLatticeTheory. +Section TBDistrLatticeTheory. +Context {disp : unit}. +Local Notation tbDistrLatticeType := (tbDistrLatticeType disp). +Context {L : tbDistrLatticeType}. +Implicit Types (I : finType) (T : eqType) (x y : L). + +Local Notation "1" := top. + +Hint Resolve le0x lex1 : core. + +Lemma meetx1 : right_id 1 (@meet _ L). +Proof. exact: (@joinx0 _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meet1x : left_id 1 (@meet _ L). +Proof. exact: (@join0x _ [tbDistrLatticeType of L^c]). Qed. + +Lemma joinx1 : right_zero 1 (@join _ L). +Proof. exact: (@meetx0 _ [tbDistrLatticeType of L^c]). Qed. + +Lemma join1x : left_zero 1 (@join _ L). +Proof. exact: (@meet0x _ [tbDistrLatticeType of L^c]). Qed. + +Lemma le1x x : (1 <= x) = (x == 1). +Proof. exact: (@lex0 _ [tbDistrLatticeType of L^c]). Qed. + +Lemma leI2l_le y t x z : y `|` z = 1 -> x `&` y <= z `&` t -> x <= z. +Proof. rewrite joinC; exact: (@leU2l_le _ [tbDistrLatticeType of L^c]). Qed. + +Lemma leI2r_le y t x z : y `|` z = 1 -> y `&` x <= t `&` z -> x <= z. +Proof. rewrite joinC; exact: (@leU2r_le _ [tbDistrLatticeType of L^c]). Qed. + +Lemma cover_leIxl z x y : z `|` y = 1 -> (x `&` z <= y) = (x <= y). +Proof. +rewrite joinC; exact: (@disjoint_lexUl _ [tbDistrLatticeType of L^c]). +Qed. + +Lemma cover_leIxr z x y : z `|` y = 1 -> (z `&` x <= y) = (x <= y). +Proof. +rewrite joinC; exact: (@disjoint_lexUr _ [tbDistrLatticeType of L^c]). +Qed. + +Lemma leI2E x y z t : x `|` t = 1 -> y `|` z = 1 -> + (x `&` y <= z `&` t) = (x <= z) && (y <= t). +Proof. +by move=> ? ?; apply: (@leU2E _ [tbDistrLatticeType of L^c]); rewrite meetC. +Qed. + +Lemma meet_eq1 x y : (x `&` y == 1) = (x == 1) && (y == 1). +Proof. exact: (@join_eq0 _ [tbDistrLatticeType of L^c]). Qed. + +Canonical meet_monoid := Monoid.Law (@meetA _ _) meet1x meetx1. +Canonical meet_comoid := Monoid.ComLaw (@meetC _ _). + +Canonical meet_muloid := Monoid.MulLaw (@meet0x _ L) (@meetx0 _ _). +Canonical join_muloid := Monoid.MulLaw join1x joinx1. +Canonical join_addoid := Monoid.AddLaw (@meetUl _ L) (@meetUr _ _). +Canonical meet_addoid := Monoid.AddLaw (@joinIl _ L) (@joinIr _ _). + +Lemma meets_inf I (j : I) (P : {pred I}) (F : I -> L) : + P j -> \meet_(i | P i) F i <= F j. +Proof. exact: (@join_sup _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meets_max I (j : I) (u : L) (P : {pred I}) (F : I -> L) : + P j -> F j <= u -> \meet_(i | P i) F i <= u. +Proof. exact: (@join_min _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meetsP I (l : L) (P : {pred I}) (F : I -> L) : + reflect (forall i : I, P i -> l <= F i) (l <= \meet_(i | P i) F i). +Proof. exact: (@joinsP _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meet_inf_seq T (r : seq T) (P : {pred T}) (F : T -> L) (x : T) : + x \in r -> P x -> \meet_(i <- r | P i) F i <= F x. +Proof. exact: (@join_sup_seq _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meet_max_seq T (r : seq T) (P : {pred T}) (F : T -> L) (x : T) (u : L) : + x \in r -> P x -> F x <= u -> \meet_(x <- r | P x) F x <= u. +Proof. exact: (@join_min_seq _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meetsP_seq T (r : seq T) (P : {pred T}) (F : T -> L) (l : L) : + reflect (forall x : T, x \in r -> P x -> l <= F x) + (l <= \meet_(x <- r | P x) F x). +Proof. exact: (@joinsP_seq _ [tbDistrLatticeType of L^c]). Qed. + +Lemma le_meets I (A B : {set I}) (F : I -> L) : + A \subset B -> \meet_(i in B) F i <= \meet_(i in A) F i. +Proof. exact: (@le_joins _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meets_setU I (A B : {set I}) (F : I -> L) : + \meet_(i in (A :|: B)) F i = \meet_(i in A) F i `&` \meet_(i in B) F i. +Proof. exact: (@joins_setU _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meet_seq I (r : seq I) (F : I -> L) : + \meet_(i <- r) F i = \meet_(i in r) F i. +Proof. exact: (@join_seq _ [tbDistrLatticeType of L^c]). Qed. + +Lemma meets_total I (d : L) (P : {pred I}) (F : I -> L) : + (forall i : I, P i -> d `|` F i = 1) -> d `|` \meet_(i | P i) F i = 1. +Proof. exact: (@joins_disjoint _ [tbDistrLatticeType of L^c]). Qed. + +End TBDistrLatticeTheory. +End TBDistrLatticeTheory. + +Module Import CBDistrLatticeTheory. +Section CBDistrLatticeTheory. +Context {disp : unit}. +Local Notation cbDistrLatticeType := (cbDistrLatticeType disp). +Context {L : cbDistrLatticeType}. +Implicit Types (x y z : L). +Local Notation "0" := bottom. + +Lemma subKI x y : y `&` (x `\` y) = 0. +Proof. by case: L x y => ? [? ?[]]. Qed. + +Lemma subIK x y : (x `\` y) `&` y = 0. +Proof. by rewrite meetC subKI. Qed. + +Lemma meetIB z x y : (z `&` y) `&` (x `\` y) = 0. +Proof. by rewrite -meetA subKI meetx0. Qed. + +Lemma meetBI z x y : (x `\` y) `&` (z `&` y) = 0. +Proof. by rewrite meetC meetIB. Qed. + +Lemma joinIB y x : (x `&` y) `|` (x `\` y) = x. +Proof. by case: L x y => ? [? ?[]]. Qed. + +Lemma joinBI y x : (x `\` y) `|` (x `&` y) = x. +Proof. by rewrite joinC joinIB. Qed. + +Lemma joinIBC y x : (y `&` x) `|` (x `\` y) = x. +Proof. by rewrite meetC joinIB. Qed. + +Lemma joinBIC y x : (x `\` y) `|` (y `&` x) = x. +Proof. by rewrite meetC joinBI. Qed. + +Lemma leBx x y : x `\` y <= x. +Proof. by rewrite -{2}[x](joinIB y) lexU2 // lexx orbT. Qed. +Hint Resolve leBx : core. + +Lemma subxx x : x `\` x = 0. +Proof. by have := subKI x x; rewrite (meet_idPr _). Qed. + +Lemma leBl z x y : x <= y -> x `\` z <= y `\` z. +Proof. +rewrite -{1}[x](joinIB z) -{1}[y](joinIB z). +by rewrite leU2E ?meetIB ?meetBI // => /andP []. +Qed. + +Lemma subKU y x : y `|` (x `\` y) = y `|` x. +Proof. +apply/eqP; rewrite eq_le leU2 //= leUx leUl. +by apply/meet_idPl; have := joinIB y x; rewrite joinIl (join_idPr _). +Qed. + +Lemma subUK y x : (x `\` y) `|` y = x `|` y. +Proof. by rewrite joinC subKU joinC. Qed. + +Lemma leBKU y x : y <= x -> y `|` (x `\` y) = x. +Proof. by move=> /join_idPl {2}<-; rewrite subKU. Qed. + +Lemma leBUK y x : y <= x -> (x `\` y) `|` y = x. +Proof. by move=> leyx; rewrite joinC leBKU. Qed. + +Lemma leBLR x y z : (x `\` y <= z) = (x <= y `|` z). +Proof. +apply/idP/idP; first by move=> /join_idPl <-; rewrite joinA subKU joinAC leUr. +by rewrite -{1}[x](joinIB y) => /(leU2r_le (subIK _ _)). +Qed. + +Lemma subUx x y z : (x `|` y) `\` z = (x `\` z) `|` (y `\` z). +Proof. +apply/eqP; rewrite eq_le leUx !leBl ?leUr ?leUl ?andbT //. +by rewrite leBLR joinA subKU joinAC subKU joinAC -joinA leUr. +Qed. + +Lemma sub_eq0 x y : (x `\` y == 0) = (x <= y). +Proof. by rewrite -lex0 leBLR joinx0. Qed. + +Lemma joinxB x y z : x `|` (y `\` z) = ((x `|` y) `\` z) `|` (x `&` z). +Proof. by rewrite subUx joinAC joinBI. Qed. + +Lemma joinBx x y z : (y `\` z) `|` x = ((y `|` x) `\` z) `|` (z `&` x). +Proof. by rewrite ![_ `|` x]joinC ![_ `&` x]meetC joinxB. Qed. + +Lemma leBr z x y : x <= y -> z `\` y <= z `\` x. +Proof. +by move=> lexy; rewrite leBLR joinxB (meet_idPr _) ?leBUK ?leUr ?lexU2 ?lexy. +Qed. + +Lemma leB2 x y z t : x <= z -> t <= y -> x `\` y <= z `\` t. +Proof. by move=> /(@leBl t) ? /(@leBr x) /le_trans ->. Qed. + +Lemma meet_eq0E_sub z x y : x <= z -> (x `&` y == 0) = (x <= z `\` y). +Proof. +move=> xz; apply/idP/idP; last by move=> /meet_idPr <-; rewrite -meetA meetBI. +by move=> /eqP xIy_eq0; rewrite -[x](joinIB y) xIy_eq0 join0x leBl. +Qed. + +Lemma leBRL x y z : (x <= z `\` y) = (x <= z) && (x `&` y == 0). +Proof. +apply/idP/idP => [xyz|]; first by rewrite (@meet_eq0E_sub z) // (le_trans xyz). +by move=> /andP [?]; rewrite -meet_eq0E_sub. +Qed. + +Lemma eq_sub x y z : (x `\` y == z) = (z <= x <= y `|` z) && (z `&` y == 0). +Proof. by rewrite eq_le leBLR leBRL andbCA andbA. Qed. + +Lemma subxU x y z : z `\` (x `|` y) = (z `\` x) `&` (z `\` y). +Proof. +apply/eqP; rewrite eq_le lexI !leBr ?leUl ?leUr //=. +rewrite leBRL leIx2 ?leBx //= meetUr meetAC subIK -meetA subIK. +by rewrite meet0x meetx0 joinx0. +Qed. + +Lemma subx0 x : x `\` 0 = x. +Proof. by apply/eqP; rewrite eq_sub join0x meetx0 lexx eqxx. Qed. + +Lemma sub0x x : 0 `\` x = 0. +Proof. by apply/eqP; rewrite eq_sub joinx0 meet0x lexx eqxx le0x. Qed. + +Lemma subIx x y z : (x `&` y) `\` z = (x `\` z) `&` (y `\` z). +Proof. +apply/eqP; rewrite eq_sub joinIr ?leI2 ?subKU ?leUr ?leBx //=. +by rewrite -meetA subIK meetx0. +Qed. + +Lemma meetxB x y z : x `&` (y `\` z) = (x `&` y) `\` z. +Proof. by rewrite subIx -{1}[x](joinBI z) meetUl meetIB joinx0. Qed. + +Lemma meetBx x y z : (x `\` y) `&` z = (x `&` z) `\` y. +Proof. by rewrite ![_ `&` z]meetC meetxB. Qed. + +Lemma subxI x y z : x `\` (y `&` z) = (x `\` y) `|` (x `\` z). +Proof. +apply/eqP; rewrite eq_sub leUx !leBx //= joinIl joinA joinCA !subKU. +rewrite joinCA -joinA [_ `|` x]joinC  //. +by rewrite -joinIl leUr /= meetUl {1}[_ `&` z]meetC ?meetBI joinx0. +Qed. + +Lemma subBx x y z : (x `\` y) `\` z = x `\` (y `|` z). +Proof. +apply/eqP; rewrite eq_sub leBr ?leUl //=. +rewrite subxU joinIr subKU -joinIr (meet_idPl _) ?leUr //=. +by rewrite -meetA subIK meetx0. +Qed. + +Lemma subxB x y z : x `\` (y `\` z) = (x `\` y) `|` (x `&` z). +Proof. +rewrite -[y in RHS](joinIB z) subxU joinIl subxI -joinA. +rewrite joinBI (join_idPl _) // joinBx [x `|` _](join_idPr _) ?leIl //. +by rewrite meetA meetAC subIK meet0x joinx0 (meet_idPr _). +Qed. + +Lemma joinBK x y : (y `|` x) `\` x = (y `\` x). +Proof. by rewrite subUx subxx joinx0. Qed. + +Lemma joinBKC x y : (x `|` y) `\` x = (y `\` x). +Proof. by rewrite subUx subxx join0x. Qed. + +Lemma disj_le x y : x `&` y == 0 -> x <= y = (x == 0). +Proof. +have [->|x_neq0] := altP (x =P 0); first by rewrite le0x. +by apply: contraTF => lexy; rewrite (meet_idPl _). +Qed. + +Lemma disj_leC x y : y `&` x == 0 -> x <= y = (x == 0). +Proof. by rewrite meetC => /disj_le. Qed. + +Lemma disj_subl x y : x `&` y == 0 -> x `\` y = x. +Proof. by move=> dxy; apply/eqP; rewrite eq_sub dxy lexx leUr. Qed. + +Lemma disj_subr x y : x `&` y == 0 -> y `\` x = y. +Proof. by rewrite meetC => /disj_subl. Qed. + +Lemma lt0B x y : x < y -> 0 < y `\` x. +Proof. +by move=> ?; rewrite lt_leAnge leBRL leBLR le0x meet0x eqxx joinx0 /= lt_geF. +Qed. + +End CBDistrLatticeTheory. +End CBDistrLatticeTheory. + +Module Import CTBDistrLatticeTheory. +Section CTBDistrLatticeTheory. +Context {disp : unit}. +Local Notation ctbDistrLatticeType := (ctbDistrLatticeType disp). +Context {L : ctbDistrLatticeType}. +Implicit Types (x y z : L). +Local Notation "0" := bottom. +Local Notation "1" := top. + +Lemma complE x : ~` x = 1 `\` x. +Proof. by case: L x => [?[? ? ? ?[]]]. Qed. + +Lemma sub1x x : 1 `\` x = ~` x. +Proof. by rewrite complE. Qed. + +Lemma subE x y : x `\` y = x `&` ~` y. +Proof. by rewrite complE meetxB meetx1. Qed. + +Lemma complK : involutive (@compl _ L). +Proof. by move=> x; rewrite !complE subxB subxx meet1x join0x. Qed. + +Lemma compl_inj : injective (@compl _ L). +Proof. exact/inv_inj/complK. Qed. + +Lemma disj_leC x y : (x `&` y == 0) = (x <= ~` y). +Proof. by rewrite -sub_eq0 subE complK. Qed. + +Lemma leC x y : (~` x <= ~` y) = (y <= x). +Proof. +gen have leC : x y / y <= x -> ~` x <= ~` y; last first. + by apply/idP/idP=> /leC; rewrite ?complK. +by move=> leyx; rewrite !complE leBr. +Qed. + +Lemma complU x y : ~` (x `|` y) = ~` x `&` ~` y. +Proof. by rewrite !complE subxU. Qed. + +Lemma complI x y : ~` (x `&` y) = ~` x `|` ~` y. +Proof. by rewrite !complE subxI. Qed. + +Lemma joinxC x : x `|` ~` x = 1. +Proof. by rewrite complE subKU joinx1. Qed. + +Lemma joinCx x : ~` x `|` x = 1. +Proof. by rewrite joinC joinxC. Qed. + +Lemma meetxC x : x `&` ~` x = 0. +Proof. by rewrite complE subKI. Qed. + +Lemma meetCx x : ~` x `&` x = 0. +Proof. by rewrite meetC meetxC. Qed. + +Lemma compl1 : ~` 1 = 0 :> L. +Proof. by rewrite complE subxx. Qed. + +Lemma compl0 : ~` 0 = 1 :> L. +Proof. by rewrite complE subx0. Qed. + +Lemma complB x y : ~` (x `\` y) = ~` x `|` y. +Proof. by rewrite !complE subxB meet1x. Qed. + +Lemma leBC x y : x `\` y <= ~` y. +Proof. by rewrite leBLR joinxC lex1. Qed. + +Lemma leCx x y : (~` x <= y) = (~` y <= x). +Proof. by rewrite !complE !leBLR joinC. Qed. + +Lemma lexC x y : (x <= ~` y) = (y <= ~` x). +Proof. by rewrite !complE !leBRL !lex1 meetC. Qed. + +Lemma compl_joins (J : Type) (r : seq J) (P : {pred J}) (F : J -> L) : + ~` (\join_(j <- r | P j) F j) = \meet_(j <- r | P j) ~` F j. +Proof. by elim/big_rec2: _=> [|i x y ? <-]; rewrite ?compl0 ?complU. Qed. + +Lemma compl_meets (J : Type) (r : seq J) (P : {pred J}) (F : J -> L) : + ~` (\meet_(j <- r | P j) F j) = \join_(j <- r | P j) ~` F j. +Proof. by elim/big_rec2: _=> [|i x y ? <-]; rewrite ?compl1 ?complI. Qed. + +End CTBDistrLatticeTheory. +End CTBDistrLatticeTheory. + +(*************) +(* FACTORIES *) +(*************) + +Module TotalPOrderMixin. +Section TotalPOrderMixin. +Variable (disp : unit) (T : porderType disp). +Definition of_ := total (<=%O : rel T). +Variable (m : of_). +Implicit Types (x y z : T). + +Let comparableT x y : x >=< y := m x y. + +Fact ltgtP x y : + compare x y (y == x) (x == y) (x >= y) (x <= y) (x > y) (x < y). +Proof. exact: comparable_ltgtP. Qed. + +Fact leP x y : le_xor_gt x y (x <= y) (y < x). +Proof. exact: comparable_leP. Qed. + +Fact ltP x y : lt_xor_ge x y (y <= x) (x < y). +Proof. exact: comparable_ltP. Qed. + +Definition meet x y := if x <= y then x else y. +Definition join x y := if y <= x then x else y. + +Fact meetC : commutative meet. +Proof. by move=> x y; rewrite /meet; have [] := ltgtP. Qed. + +Fact joinC : commutative join. +Proof. by move=> x y; rewrite /join; have [] := ltgtP. Qed. + +Fact meetA : associative meet. +Proof. +move=> x y z; rewrite /meet; case: (leP y z) => yz; case: (leP x y) => xy //=. +- by rewrite (le_trans xy). +- by rewrite yz. +by rewrite !lt_geF // (lt_trans yz). +Qed. + +Fact joinA : associative join. +Proof. +move=> x y z; rewrite /join; case: (leP z y) => yz; case: (leP y x) => xy //=. +- by rewrite (le_trans yz). +- by rewrite yz. +by rewrite !lt_geF // (lt_trans xy). +Qed. + +Fact joinKI y x : meet x (join x y) = x. +Proof. by rewrite /meet /join; case: (leP y x) => yx; rewrite ?lexx ?ltW. Qed. + +Fact meetKU y x : join x (meet x y) = x. +Proof. by rewrite /meet /join; case: (leP x y) => yx; rewrite ?lexx ?ltW. Qed. + +Fact leEmeet x y : (x <= y) = (meet x y == x). +Proof. by rewrite /meet; case: leP => ?; rewrite ?eqxx ?lt_eqF. Qed. + +Fact meetUl : left_distributive meet join. +Proof. +move=> x y z; rewrite /meet /join. +case: (leP y z) => yz; case: (leP y x) => xy //=; first 1 last. +- by rewrite yz [x <= z](le_trans _ yz) ?[x <= y]ltW // lt_geF. +- by rewrite lt_geF ?lexx // (lt_le_trans yz). +- by rewrite lt_geF //; have [/lt_geF->| |->] := (ltgtP x z); rewrite ?lexx. +- by have [] := (leP x z); rewrite ?xy ?yz. +Qed. + +Definition distrLatticeMixin := + @DistrLatticeMixin _ (@POrder.Pack disp T (POrder.class T)) _ _ + meetC joinC meetA joinA joinKI meetKU leEmeet meetUl. + +Definition orderMixin : + totalOrderMixin (DistrLatticeType _ distrLatticeMixin) := + m. + +End TotalPOrderMixin. + +Module Exports. +Notation totalPOrderMixin := of_. +Coercion distrLatticeMixin : totalPOrderMixin >-> Order.DistrLattice.mixin_of. +Coercion orderMixin : totalPOrderMixin >-> totalOrderMixin. +End Exports. + +End TotalPOrderMixin. +Import TotalPOrderMixin.Exports. + +Module LtPOrderMixin. +Section LtPOrderMixin. +Variable (T : eqType). + +Record of_ := Build { + le : rel T; + lt : rel T; + le_def : forall x y, le x y = (x == y) || lt x y; + lt_irr : irreflexive lt; + lt_trans : transitive lt; +}. + +Variable (m : of_). + +Fact lt_asym x y : (lt m x y && lt m y x) = false. +Proof. +by apply/negP => /andP [] xy /(lt_trans xy); apply/negP; rewrite (lt_irr m x). +Qed. + +Fact lt_def x y : lt m x y = (y != x) && le m x y. +Proof. by rewrite le_def eq_sym; case: eqP => //= <-; rewrite lt_irr. Qed. + +Fact le_refl : reflexive (le m). +Proof. by move=> ?; rewrite le_def eqxx. Qed. + +Fact le_anti : antisymmetric (le m). +Proof. +by move=> ? ?; rewrite !le_def eq_sym -orb_andr lt_asym orbF => /eqP. +Qed. + +Fact le_trans : transitive (le m). +Proof. +by move=> y x z; rewrite !le_def => /predU1P [-> //|ltxy] /predU1P [<-|ltyz]; + rewrite ?ltxy ?(lt_trans ltxy ltyz) // ?orbT. +Qed. + +Definition lePOrderMixin : lePOrderMixin T := + @LePOrderMixin _ (le m) (lt m) lt_def le_refl le_anti le_trans. + +End LtPOrderMixin. + +Module Exports. +Notation ltPOrderMixin := of_. +Notation LtPOrderMixin := Build. +Coercion lePOrderMixin : ltPOrderMixin >-> POrder.mixin_of. +End Exports. + +End LtPOrderMixin. +Import LtPOrderMixin.Exports. + +Module MeetJoinMixin. +Section MeetJoinMixin. + +Variable (T : choiceType). + +Record of_ := Build { + le : rel T; + lt : rel T; + meet : T -> T -> T; + join : T -> T -> T; + le_def : forall x y : T, le x y = (meet x y == x); + lt_def : forall x y : T, lt x y = (y != x) && le x y; + meetC : commutative meet; + joinC : commutative join; + meetA : associative meet; + joinA : associative join; + joinKI : forall y x : T, meet x (join x y) = x; + meetKU : forall y x : T, join x (meet x y) = x; + meetUl : left_distributive meet join; + meetxx : idempotent meet; +}. + +Variable (m : of_). + +Fact le_refl : reflexive (le m). +Proof. by move=> x; rewrite le_def meetxx. Qed. + +Fact le_anti : antisymmetric (le m). +Proof. by move=> x y; rewrite !le_def meetC => /andP [] /eqP {2}<- /eqP ->. Qed. + +Fact le_trans : transitive (le m). +Proof. +move=> y x z; rewrite !le_def => /eqP lexy /eqP leyz; apply/eqP. +by rewrite -[in LHS]lexy -meetA leyz lexy. +Qed. + +Definition porderMixin : lePOrderMixin T := + LePOrderMixin (lt_def m) le_refl le_anti le_trans. + +Let T_porderType := POrderType tt T porderMixin. + +Definition distrLatticeMixin : distrLatticeMixin T_porderType := + @DistrLatticeMixin tt (POrderType tt T porderMixin) (meet m) (join m) + (meetC m) (joinC m) (meetA m) (joinA m) + (joinKI m) (meetKU m) (le_def m) (meetUl m). + +End MeetJoinMixin. + +Module Exports. +Notation meetJoinMixin := of_. +Notation MeetJoinMixin := Build. +Coercion porderMixin : meetJoinMixin >-> lePOrderMixin. +Coercion distrLatticeMixin : meetJoinMixin >-> DistrLattice.mixin_of. +End Exports. + +End MeetJoinMixin. +Import MeetJoinMixin.Exports. + +Module LeOrderMixin. +Section LeOrderMixin. + +Variables (T : choiceType). + +Record of_ := Build { + le : rel T; + lt : rel T; + meet : T -> T -> T; + join : T -> T -> T; + lt_def : forall x y, lt x y = (y != x) && le x y; + meet_def : forall x y, meet x y = if le x y then x else y; + join_def : forall x y, join x y = if le y x then x else y; + le_anti : antisymmetric le; + le_trans : transitive le; + le_total : total le; +}. + +Variables (m : of_). + +Fact le_refl : reflexive (le m). +Proof. by move=> x; case: (le m x x) (le_total m x x). Qed. + +Definition lePOrderMixin := + LePOrderMixin (lt_def m) le_refl (@le_anti m) (@le_trans m). + +Let T_total_porderType : porderType tt := POrderType tt T lePOrderMixin. + +Let T_total_distrLatticeType : distrLatticeType tt := + DistrLatticeType T_total_porderType + (le_total m : totalPOrderMixin T_total_porderType). + +Implicit Types (x y z : T_total_distrLatticeType). + +Fact meetE x y : meet m x y = x `&` y. Proof. by rewrite meet_def. Qed. +Fact joinE x y : join m x y = x `|` y. Proof. by rewrite join_def. Qed. +Fact meetC : commutative (meet m). +Proof. by move=> *; rewrite !meetE meetC. Qed. +Fact joinC : commutative (join m). +Proof. by move=> *; rewrite !joinE joinC. Qed. +Fact meetA : associative (meet m). +Proof. by move=> *; rewrite !meetE meetA. Qed. +Fact joinA : associative (join m). +Proof. by move=> *; rewrite !joinE joinA. Qed. +Fact joinKI y x : meet m x (join m x y) = x. +Proof. by rewrite meetE joinE joinKI. Qed. +Fact meetKU y x : join m x (meet m x y) = x. +Proof. by rewrite meetE joinE meetKU. Qed. +Fact meetUl : left_distributive (meet m) (join m). +Proof. by move=> *; rewrite !meetE !joinE meetUl. Qed. +Fact meetxx : idempotent (meet m). +Proof. by move=> *; rewrite meetE meetxx. Qed. +Fact le_def x y : x <= y = (meet m x y == x). +Proof. by rewrite meetE (eq_meetl x y). Qed. + +Definition distrLatticeMixin : meetJoinMixin T := + @MeetJoinMixin _ (le m) (lt m) (meet m) (join m) le_def (lt_def m) + meetC joinC meetA joinA joinKI meetKU meetUl meetxx. + +Let T_porderType := POrderType tt T distrLatticeMixin. +Let T_distrLatticeType : distrLatticeType tt := + DistrLatticeType T_porderType distrLatticeMixin. + +Definition totalMixin : totalOrderMixin T_distrLatticeType := le_total m. + +End LeOrderMixin. + +Module Exports. +Notation leOrderMixin := of_. +Notation LeOrderMixin := Build. +Coercion distrLatticeMixin : leOrderMixin >-> meetJoinMixin. +Coercion totalMixin : leOrderMixin >-> totalOrderMixin. +End Exports. + +End LeOrderMixin. +Import LeOrderMixin.Exports. + +Module LtOrderMixin. + +Record of_ (T : choiceType) := Build { + le : rel T; + lt : rel T; + meet : T -> T -> T; + join : T -> T -> T; + le_def : forall x y, le x y = (x == y) || lt x y; + meet_def : forall x y, meet x y = if lt x y then x else y; + join_def : forall x y, join x y = if lt y x then x else y; + lt_irr : irreflexive lt; + lt_trans : transitive lt; + lt_total : forall x y, x != y -> lt x y || lt y x; +}. + +Section LtOrderMixin. + +Variables (T : choiceType) (m : of_ T). + +Let T_total_porderType : porderType tt := + POrderType tt T (LtPOrderMixin (le_def m) (lt_irr m) (@lt_trans _ m)). + +Fact le_total : total (le m). +Proof. +move=> x y; rewrite !le_def (eq_sym y). +by case: (altP eqP); last exact: lt_total. +Qed. + +Let T_total_distrLatticeType : distrLatticeType tt := + DistrLatticeType T_total_porderType + (le_total : totalPOrderMixin T_total_porderType). + +Implicit Types (x y z : T_total_distrLatticeType). + +Fact leP x y : + lel_xor_gt x y (x <= y) (y < x) (y `&` x) (x `&` y) (y `|` x) (x `|` y). +Proof. by apply/lcomparable_leP/le_total. Qed. +Fact meetE x y : meet m x y = x `&` y. +Proof. by rewrite meet_def (_ : lt m x y = (x < y)) //; case: (leP y). Qed. +Fact joinE x y : join m x y = x `|` y. +Proof. by rewrite join_def (_ : lt m y x = (y < x)) //; case: leP. Qed. +Fact meetC : commutative (meet m). +Proof. by move=> *; rewrite !meetE meetC. Qed. +Fact joinC : commutative (join m). +Proof. by move=> *; rewrite !joinE joinC. Qed. +Fact meetA : associative (meet m). +Proof. by move=> *; rewrite !meetE meetA. Qed. +Fact joinA : associative (join m). +Proof. by move=> *; rewrite !joinE joinA. Qed. +Fact joinKI y x : meet m x (join m x y) = x. +Proof. by rewrite meetE joinE joinKI. Qed. +Fact meetKU y x : join m x (meet m x y) = x. +Proof. by rewrite meetE joinE meetKU. Qed. +Fact meetUl : left_distributive (meet m) (join m). +Proof. by move=> *; rewrite !meetE !joinE meetUl. Qed. +Fact meetxx : idempotent (meet m). +Proof. by move=> *; rewrite meetE meetxx. Qed. +Fact le_def' x y : x <= y = (meet m x y == x). +Proof. by rewrite meetE (eq_meetl x y). Qed. + +Definition distrLatticeMixin : meetJoinMixin T := + @MeetJoinMixin _ (le m) (lt m) (meet m) (join m) + le_def' (@lt_def _ T_total_distrLatticeType) + meetC joinC meetA joinA joinKI meetKU meetUl meetxx. + +Let T_porderType := POrderType tt T distrLatticeMixin. +Let T_distrLatticeType : distrLatticeType tt := + DistrLatticeType T_porderType distrLatticeMixin. + +Definition totalMixin : totalOrderMixin T_distrLatticeType := le_total. + +End LtOrderMixin. + +Module Exports. +Notation ltOrderMixin := of_. +Notation LtOrderMixin := Build. +Coercion distrLatticeMixin : ltOrderMixin >-> meetJoinMixin. +Coercion totalMixin : ltOrderMixin >-> totalOrderMixin. +End Exports. + +End LtOrderMixin. +Import LtOrderMixin.Exports. + +Module CanMixin. +Section CanMixin. + +Section Total. + +Variables (disp : unit) (T : porderType disp). +Variables (disp' : unit) (T' : orderType disp) (f : T -> T'). + +Lemma MonoTotal : {mono f : x y / x <= y} -> + totalPOrderMixin T' -> totalPOrderMixin T. +Proof. by move=> f_mono T'_tot x y; rewrite -!f_mono le_total. Qed. + +End Total. + +Section Order. + +Variables (T : choiceType) (disp : unit). + +Section Partial. +Variables (T' : porderType disp) (f : T -> T'). + +Section PCan. +Variables (f' : T' -> option T) (f_can : pcancel f f'). + +Definition le (x y : T) := f x <= f y. +Definition lt (x y : T) := f x < f y. + +Fact refl : reflexive le. Proof. by move=> ?; apply: lexx. Qed. +Fact anti : antisymmetric le. +Proof. by move=> x y /le_anti /(pcan_inj f_can). Qed. +Fact trans : transitive le. Proof. by move=> y x z xy /(le_trans xy). Qed. +Fact lt_def x y : lt x y = (y != x) && le x y. +Proof. by rewrite /lt lt_def (inj_eq (pcan_inj f_can)). Qed. + +Definition PcanPOrder := LePOrderMixin lt_def refl anti trans. + +End PCan. + +Definition CanPOrder f' (f_can : cancel f f') := PcanPOrder (can_pcan f_can). + +End Partial. + +Section Total. + +Variables (T' : orderType disp) (f : T -> T'). + +Section PCan. + +Variables (f' : T' -> option T) (f_can : pcancel f f'). + +Let T_porderType := POrderType disp T (PcanPOrder f_can). + +Let total_le : total (le f). +Proof. by apply: (@MonoTotal _ T_porderType _ f) => //; apply: le_total. Qed. + +Definition PcanOrder := LeOrderMixin + (@lt_def _ _ _ f_can) (fun _ _ => erefl) (fun _ _ => erefl) + (@anti _ _ _ f_can) (@trans _ _) total_le. + +End PCan. + +Definition CanOrder f' (f_can : cancel f f') := PcanOrder (can_pcan f_can). + +End Total. +End Order. + +Section DistrLattice. + +Variables (disp : unit) (T : porderType disp). +Variables (disp' : unit) (T' : distrLatticeType disp) (f : T -> T'). + +Variables (f' : T' -> T) (f_can : cancel f f') (f'_can : cancel f' f). +Variable (f_mono : {mono f : x y / x <= y}). + +Definition meet (x y : T) := f' (meet (f x) (f y)). +Definition join (x y : T) := f' (join (f x) (f y)). + +Lemma meetC : commutative meet. Proof. by move=> x y; rewrite /meet meetC. Qed. +Lemma joinC : commutative join. Proof. by move=> x y; rewrite /join joinC. Qed. +Lemma meetA : associative meet. +Proof. by move=> y x z; rewrite /meet !f'_can meetA. Qed. +Lemma joinA : associative join. +Proof. by move=> y x z; rewrite /join !f'_can joinA. Qed. +Lemma joinKI y x : meet x (join x y) = x. +Proof. by rewrite /meet /join f'_can joinKI f_can. Qed. +Lemma meetKI y x : join x (meet x y) = x. +Proof. by rewrite /join /meet f'_can meetKU f_can. Qed. +Lemma meet_eql x y : (x <= y) = (meet x y == x). +Proof. by rewrite /meet -(can_eq f_can) f'_can eq_meetl f_mono. Qed. +Lemma meetUl : left_distributive meet join. +Proof. by move=> x y z; rewrite /meet /join !f'_can meetUl. Qed. + +Definition IsoDistrLattice := + DistrLatticeMixin meetC joinC meetA joinA joinKI meetKI meet_eql meetUl. + +End DistrLattice. + +End CanMixin. + +Module Exports. +Notation MonoTotalMixin := MonoTotal. +Notation PcanPOrderMixin := PcanPOrder. +Notation CanPOrderMixin := CanPOrder. +Notation PcanOrderMixin := PcanOrder. +Notation CanOrderMixin := CanOrder. +Notation IsoDistrLatticeMixin := IsoDistrLattice. +End Exports. +End CanMixin. +Import CanMixin.Exports. + +Module SubOrder. + +Section Partial. +Context {disp : unit} {T : porderType disp} (P : {pred T}) (sT : subType P). + +Definition sub_POrderMixin := PcanPOrderMixin (@valK _ _ sT). +Canonical sub_POrderType := Eval hnf in POrderType disp sT sub_POrderMixin. + +Lemma leEsub (x y : sT) : (x <= y) = (val x <= val y). Proof. by []. Qed. + +Lemma ltEsub (x y : sT) : (x < y) = (val x < val y). Proof. by []. Qed. + +End Partial. + +Section Total. +Context {disp : unit} {T : orderType disp} (P : {pred T}) (sT : subType P). + +Definition sub_TotalOrderMixin : totalPOrderMixin (sub_POrderType sT) := + @MonoTotalMixin _ _ _ val (fun _ _ => erefl) (@le_total _ T). +Canonical sub_DistrLatticeType := + Eval hnf in DistrLatticeType sT sub_TotalOrderMixin. +Canonical sub_OrderType := Eval hnf in OrderType sT sub_TotalOrderMixin. + +End Total. +Arguments sub_TotalOrderMixin {disp T} [P]. + +Module Exports. +Notation "[ 'porderMixin' 'of' T 'by' <: ]" := + (sub_POrderMixin _ : lePOrderMixin [eqType of T]) + (at level 0, format "[ 'porderMixin' 'of' T 'by' <: ]") : form_scope. + +Notation "[ 'totalOrderMixin' 'of' T 'by' <: ]" := + (sub_TotalOrderMixin _ : totalPOrderMixin [porderType of T]) + (at level 0, format "[ 'totalOrderMixin' 'of' T 'by' <: ]", + only parsing) : form_scope. + +Canonical sub_POrderType. +Canonical sub_DistrLatticeType. +Canonical sub_OrderType. + +Definition leEsub := @leEsub. +Definition ltEsub := @ltEsub. +End Exports. +End SubOrder. +Import SubOrder.Exports. + +(*************) +(* INSTANCES *) +(*************) + +(*******************************) +(* Canonical structures on nat *) +(*******************************) + +(******************************************************************************) +(* This is an example of creation of multiple canonical declarations on the *) +(* same type, with distinct displays, on the example of natural numbers. *) +(* We declare two distinct canonical orders: *) +(* - leq which is total, and where meet and join are minn and maxn, on nat *) +(* - dvdn which is partial, and where meet and join are gcdn and lcmn, *) +(* on a "copy" of nat we name natdiv *) +(******************************************************************************) + +(******************************************************************************) +(* The Module NatOrder defines leq as the canonical order on the type nat, *) +(* i.e. without creating a "copy". We use the predefined total_display, which *) +(* is designed to parse and print meet and join as minn and maxn. This looks *) +(* like standard canonical structure declaration, except we use a display. *) +(* We also use a single factory LeOrderMixin to instanciate three different *) +(* canonical declarations porderType, distrLatticeType, orderType *) +(* We finish by providing theorems to convert the operations of ordered and *) +(* lattice types to their definition without structure abstraction. *) +(******************************************************************************) + +Module NatOrder. +Section NatOrder. + +Lemma minnE x y : minn x y = if (x <= y)%N then x else y. +Proof. by case: leqP => [/minn_idPl|/ltnW /minn_idPr]. Qed. + +Lemma maxnE x y : maxn x y = if (y <= x)%N then x else y. +Proof. by case: leqP => [/maxn_idPl|/ltnW/maxn_idPr]. Qed. + +Lemma ltn_def x y : (x < y)%N = (y != x) && (x <= y)%N. +Proof. by rewrite ltn_neqAle eq_sym. Qed. + +Definition orderMixin := + LeOrderMixin ltn_def minnE maxnE anti_leq leq_trans leq_total. + +Canonical porderType := POrderType total_display nat orderMixin. +Canonical distrLatticeType := DistrLatticeType nat orderMixin. +Canonical orderType := OrderType nat orderMixin. +Canonical bDistrLatticeType := BDistrLatticeType nat (BDistrLatticeMixin leq0n). + +Lemma leEnat : le = leq. Proof. by []. Qed. +Lemma ltEnat (n m : nat) : (n < m) = (n < m)%N. Proof. by []. Qed. +Lemma meetEnat : meet = minn. Proof. by []. Qed. +Lemma joinEnat : join = maxn. Proof. by []. Qed. +Lemma botEnat : 0%O = 0%N :> nat. Proof. by []. Qed. + +End NatOrder. +Module Exports. +Canonical porderType. +Canonical distrLatticeType. +Canonical orderType. +Canonical bDistrLatticeType. +Definition leEnat := leEnat. +Definition ltEnat := ltEnat. +Definition meetEnat := meetEnat. +Definition joinEnat := joinEnat. +Definition botEnat := botEnat. +End Exports. +End NatOrder. + +(****************************************************************************) +(* The Module DvdSyntax introduces a new set of notations using the newly *) +(* created display dvd_display. We first define the display as an opaque *) +(* definition of type unit, and we use it as the first argument of the *) +(* operator which display we want to change from the default one (here le, *) +(* lt, dvd sdvd, meet, join, top and bottom, as well as big op notations on *) +(* gcd and lcm). This notations will now be used for any ordered type which *) +(* first parameter is set to dvd_display. *) +(****************************************************************************) + +Lemma dvd_display : unit. Proof. exact: tt. Qed. + +Module DvdSyntax. + +Notation dvd := (@le dvd_display _). +Notation "@ 'dvd' T" := + (@le dvd_display T) (at level 10, T at level 8, only parsing) : fun_scope. +Notation sdvd := (@lt dvd_display _). +Notation "@ 'sdvd' T" := + (@lt dvd_display T) (at level 10, T at level 8, only parsing) : fun_scope. + +Notation "x %| y" := (dvd x y) : order_scope. +Notation "x %<| y" := (sdvd x y) : order_scope. + +Notation gcd := (@meet dvd_display _). +Notation "@ 'gcd' T" := + (@meet dvd_display T) (at level 10, T at level 8, only parsing) : fun_scope. +Notation lcm := (@join dvd_display _). +Notation "@ 'lcm' T" := + (@join dvd_display T) (at level 10, T at level 8, only parsing) : fun_scope. + +Notation nat0 := (@top dvd_display _). +Notation nat1 := (@bottom dvd_display _). + +Notation "\gcd_ ( i <- r | P ) F" := + (\big[gcd/nat0]_(i <- r | P%B) F%O) : order_scope. +Notation "\gcd_ ( i <- r ) F" := + (\big[gcd/nat0]_(i <- r) F%O) : order_scope. +Notation "\gcd_ ( i | P ) F" := + (\big[gcd/nat0]_(i | P%B) F%O) : order_scope. +Notation "\gcd_ i F" := + (\big[gcd/nat0]_i F%O) : order_scope. +Notation "\gcd_ ( i : I | P ) F" := + (\big[gcd/nat0]_(i : I | P%B) F%O) (only parsing) : + order_scope. +Notation "\gcd_ ( i : I ) F" := + (\big[gcd/nat0]_(i : I) F%O) (only parsing) : order_scope. +Notation "\gcd_ ( m <= i < n | P ) F" := + (\big[gcd/nat0]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\gcd_ ( m <= i < n ) F" := + (\big[gcd/nat0]_(m <= i < n) F%O) : order_scope. +Notation "\gcd_ ( i < n | P ) F" := + (\big[gcd/nat0]_(i < n | P%B) F%O) : order_scope. +Notation "\gcd_ ( i < n ) F" := + (\big[gcd/nat0]_(i < n) F%O) : order_scope. +Notation "\gcd_ ( i 'in' A | P ) F" := + (\big[gcd/nat0]_(i in A | P%B) F%O) : order_scope. +Notation "\gcd_ ( i 'in' A ) F" := + (\big[gcd/nat0]_(i in A) F%O) : order_scope. + +Notation "\lcm_ ( i <- r | P ) F" := + (\big[lcm/nat1]_(i <- r | P%B) F%O) : order_scope. +Notation "\lcm_ ( i <- r ) F" := + (\big[lcm/nat1]_(i <- r) F%O) : order_scope. +Notation "\lcm_ ( i | P ) F" := + (\big[lcm/nat1]_(i | P%B) F%O) : order_scope. +Notation "\lcm_ i F" := + (\big[lcm/nat1]_i F%O) : order_scope. +Notation "\lcm_ ( i : I | P ) F" := + (\big[lcm/nat1]_(i : I | P%B) F%O) (only parsing) : + order_scope. +Notation "\lcm_ ( i : I ) F" := + (\big[lcm/nat1]_(i : I) F%O) (only parsing) : order_scope. +Notation "\lcm_ ( m <= i < n | P ) F" := + (\big[lcm/nat1]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\lcm_ ( m <= i < n ) F" := + (\big[lcm/nat1]_(m <= i < n) F%O) : order_scope. +Notation "\lcm_ ( i < n | P ) F" := + (\big[lcm/nat1]_(i < n | P%B) F%O) : order_scope. +Notation "\lcm_ ( i < n ) F" := + (\big[lcm/nat1]_(i < n) F%O) : order_scope. +Notation "\lcm_ ( i 'in' A | P ) F" := + (\big[lcm/nat1]_(i in A | P%B) F%O) : order_scope. +Notation "\lcm_ ( i 'in' A ) F" := + (\big[lcm/nat1]_(i in A) F%O) : order_scope. + +End DvdSyntax. + +(******************************************************************************) +(* The Module NatDvd defines dvdn as the canonical order on NatDvd.t, which *) +(* is abbreviated using the notation natdvd at the end of the module. *) +(* We use the newly defined dvd_display, described above. This looks *) +(* like standard canonical structure declaration, except we use a display and *) +(* we declare it on a "copy" of the type. *) +(* We first recover structures that are common to both nat and natdiv *) +(* (eqType, choiceType, countType) through the clone mechanisms, then we use *) +(* a single factory MeetJoinMixin to instanciate both porderType and *) +(* distrLatticeType canonical structures,and end with top and bottom. *) +(* We finish by providing theorems to convert the operations of ordered and *) +(* lattice types to their definition without structure abstraction. *) +(******************************************************************************) + +Module NatDvd. +Section NatDvd. + +Implicit Types m n p : nat. + +Lemma lcmnn n : lcmn n n = n. +Proof. by case: n => // n; rewrite /lcmn gcdnn mulnK. Qed. + +Lemma le_def m n : m %| n = (gcdn m n == m)%N. +Proof. by apply/gcdn_idPl/eqP. Qed. + +Lemma joinKI n m : gcdn m (lcmn m n) = m. +Proof. by rewrite (gcdn_idPl _)// dvdn_lcml. Qed. + +Lemma meetKU n m : lcmn m (gcdn m n) = m. +Proof. by rewrite (lcmn_idPl _)// dvdn_gcdl. Qed. + +Lemma meetUl : left_distributive gcdn lcmn. +Proof. +move=> [|m'] [|n'] [|p'] //=; rewrite ?lcmnn ?lcm0n ?lcmn0 ?gcd0n ?gcdn0//. +- by rewrite gcdnC meetKU. +- by rewrite lcmnC gcdnC meetKU. +apply: eqn_from_log; rewrite ?(gcdn_gt0, lcmn_gt0)//= => p. +by rewrite !(logn_gcd, logn_lcm) ?(gcdn_gt0, lcmn_gt0)// minn_maxl. +Qed. + +Definition t_distrLatticeMixin := MeetJoinMixin le_def (fun _ _ => erefl _) + gcdnC lcmnC gcdnA lcmnA joinKI meetKU meetUl gcdnn. + +Definition t := nat. + +Canonical eqType := [eqType of t]. +Canonical choiceType := [choiceType of t]. +Canonical countType := [countType of t]. +Canonical porderType := POrderType dvd_display t t_distrLatticeMixin. +Canonical distrLatticeType := DistrLatticeType t t_distrLatticeMixin. +Canonical bDistrLatticeType := BDistrLatticeType t + (BDistrLatticeMixin (dvd1n : forall m : t, 1 %| m)). +Canonical tbDistrLatticeType := TBDistrLatticeType t + (TBDistrLatticeMixin (dvdn0 : forall m : t, m %| 0)). + +Import DvdSyntax. +Lemma dvdE : dvd = dvdn :> rel t. Proof. by []. Qed. +Lemma sdvdE (m n : t) : m %<| n = (n != m) && (m %| n). Proof. by []. Qed. +Lemma gcdE : gcd = gcdn :> (t -> t -> t). Proof. by []. Qed. +Lemma lcmE : lcm = lcmn :> (t -> t -> t). Proof. by []. Qed. +Lemma nat1E : nat1 = 1%N :> t. Proof. by []. Qed. +Lemma nat0E : nat0 = 0%N :> t. Proof. by []. Qed. + +End NatDvd. +Module Exports. +Notation natdvd := t. +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Definition dvdEnat := dvdE. +Definition sdvdEnat := sdvdE. +Definition gcdEnat := gcdE. +Definition lcmEnat := lcmE. +Definition nat1E := nat1E. +Definition nat0E := nat0E. +End Exports. +End NatDvd. + +(*******************************) +(* Canonical structure on bool *) +(*******************************) + +Module BoolOrder. +Section BoolOrder. +Implicit Types (x y : bool). + +Fact andbE x y : x && y = if (x <= y)%N then x else y. +Proof. by case: x y => [] []. Qed. + +Fact orbE x y : x || y = if (y <= x)%N then x else y. +Proof. by case: x y => [] []. Qed. + +Fact ltn_def x y : (x < y)%N = (y != x) && (x <= y)%N. +Proof. by case: x y => [] []. Qed. + +Fact anti : antisymmetric (leq : rel bool). +Proof. by move=> x y /anti_leq /(congr1 odd); rewrite !oddb. Qed. + +Definition sub x y := x && ~~ y. + +Lemma subKI x y : y && sub x y = false. +Proof. by case: x y => [] []. Qed. + +Lemma joinIB x y : (x && y) || sub x y = x. +Proof. by case: x y => [] []. Qed. + +Definition orderMixin := + LeOrderMixin ltn_def andbE orbE anti leq_trans leq_total. + +Canonical porderType := POrderType total_display bool orderMixin. +Canonical distrLatticeType := DistrLatticeType bool orderMixin. +Canonical orderType := OrderType bool orderMixin. +Canonical bDistrLatticeType := BDistrLatticeType bool + (@BDistrLatticeMixin _ _ false (fun b : bool => leq0n b)). +Canonical tbDistrLatticeType := + TBDistrLatticeType bool (@TBDistrLatticeMixin _ _ true leq_b1). +Canonical cbDistrLatticeType := CBDistrLatticeType bool + (@CBDistrLatticeMixin _ _ (fun x y => x && ~~ y) subKI joinIB). +Canonical ctbDistrLatticeType := CTBDistrLatticeType bool + (@CTBDistrLatticeMixin _ _ sub negb (fun x => erefl : ~~ x = sub true x)). + +Canonical finPOrderType := [finPOrderType of bool]. +Canonical finDistrLatticeType := [finDistrLatticeType of bool]. +Canonical finCDistrLatticeType := [finCDistrLatticeType of bool]. +Canonical finOrderType := [finOrderType of bool]. + +Lemma leEbool : le = (leq : rel bool). Proof. by []. Qed. +Lemma ltEbool x y : (x < y) = (x < y)%N. Proof. by []. Qed. +Lemma andEbool : meet = andb. Proof. by []. Qed. +Lemma orEbool : meet = andb. Proof. by []. Qed. +Lemma subEbool x y : x `\` y = x && ~~ y. Proof. by []. Qed. +Lemma complEbool : compl = negb. Proof. by []. Qed. + +End BoolOrder. +Module Exports. +Canonical porderType. +Canonical distrLatticeType. +Canonical orderType. +Canonical bDistrLatticeType. +Canonical cbDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical ctbDistrLatticeType. +Canonical finPOrderType. +Canonical finDistrLatticeType. +Canonical finOrderType. +Canonical finCDistrLatticeType. +Definition leEbool := leEbool. +Definition ltEbool := ltEbool. +Definition andEbool := andEbool. +Definition orEbool := orEbool. +Definition subEbool := subEbool. +Definition complEbool := complEbool. +End Exports. +End BoolOrder. + +(*******************************) +(* Definition of prod_display. *) +(*******************************) + +Fact prod_display : unit. Proof. by []. Qed. + +Module Import ProdSyntax. + +Notation "<=^p%O" := (@le prod_display _) : fun_scope. +Notation ">=^p%O" := (@ge prod_display _) : fun_scope. +Notation ">=^p%O" := (@ge prod_display _) : fun_scope. +Notation "<^p%O" := (@lt prod_display _) : fun_scope. +Notation ">^p%O" := (@gt prod_display _) : fun_scope. +Notation "<?=^p%O" := (@leif prod_display _) : fun_scope. +Notation ">=<^p%O" := (@comparable prod_display _) : fun_scope. +Notation "><^p%O" := (fun x y => ~~ (@comparable prod_display _ x y)) : + fun_scope. + +Notation "<=^p y" := (>=^p%O y) : order_scope. +Notation "<=^p y :> T" := (<=^p (y : T)) (only parsing) : order_scope. +Notation ">=^p y" := (<=^p%O y) : order_scope. +Notation ">=^p y :> T" := (>=^p (y : T)) (only parsing) : order_scope. + +Notation "<^p y" := (>^p%O y) : order_scope. +Notation "<^p y :> T" := (<^p (y : T)) (only parsing) : order_scope. +Notation ">^p y" := (<^p%O y) : order_scope. +Notation ">^p y :> T" := (>^p (y : T)) (only parsing) : order_scope. + +Notation ">=<^p y" := (>=<^p%O y) : order_scope. +Notation ">=<^p y :> T" := (>=<^p (y : T)) (only parsing) : order_scope. + +Notation "x <=^p y" := (<=^p%O x y) : order_scope. +Notation "x <=^p y :> T" := ((x : T) <=^p (y : T)) (only parsing) : order_scope. +Notation "x >=^p y" := (y <=^p x) (only parsing) : order_scope. +Notation "x >=^p y :> T" := ((x : T) >=^p (y : T)) (only parsing) : order_scope. + +Notation "x <^p y" := (<^p%O x y) : order_scope. +Notation "x <^p y :> T" := ((x : T) <^p (y : T)) (only parsing) : order_scope. +Notation "x >^p y" := (y <^p x) (only parsing) : order_scope. +Notation "x >^p y :> T" := ((x : T) >^p (y : T)) (only parsing) : order_scope. + +Notation "x <=^p y <=^p z" := ((x <=^p y) && (y <=^p z)) : order_scope. +Notation "x <^p y <=^p z" := ((x <^p y) && (y <=^p z)) : order_scope. +Notation "x <=^p y <^p z" := ((x <=^p y) && (y <^p z)) : order_scope. +Notation "x <^p y <^p z" := ((x <^p y) && (y <^p z)) : order_scope. + +Notation "x <=^p y ?= 'iff' C" := (<?=^p%O x y C) : order_scope. +Notation "x <=^p y ?= 'iff' C :> T" := ((x : T) <=^p (y : T) ?= iff C) + (only parsing) : order_scope. + +Notation ">=<^p x" := (>=<^p%O x) : order_scope. +Notation ">=<^p x :> T" := (>=<^p (x : T)) (only parsing) : order_scope. +Notation "x >=<^p y" := (>=<^p%O x y) : order_scope. + +Notation "><^p x" := (fun y => ~~ (>=<^p%O x y)) : order_scope. +Notation "><^p x :> T" := (><^p (x : T)) (only parsing) : order_scope. +Notation "x ><^p y" := (~~ (><^p%O x y)) : order_scope. + +Notation "x `&^p` y" := (@meet prod_display _ x y) : order_scope. +Notation "x `|^p` y" := (@join prod_display _ x y) : order_scope. + +Notation "\join^p_ ( i <- r | P ) F" := + (\big[join/0]_(i <- r | P%B) F%O) : order_scope. +Notation "\join^p_ ( i <- r ) F" := + (\big[join/0]_(i <- r) F%O) : order_scope. +Notation "\join^p_ ( i | P ) F" := + (\big[join/0]_(i | P%B) F%O) : order_scope. +Notation "\join^p_ i F" := + (\big[join/0]_i F%O) : order_scope. +Notation "\join^p_ ( i : I | P ) F" := + (\big[join/0]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\join^p_ ( i : I ) F" := + (\big[join/0]_(i : I) F%O) (only parsing) : order_scope. +Notation "\join^p_ ( m <= i < n | P ) F" := + (\big[join/0]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\join^p_ ( m <= i < n ) F" := + (\big[join/0]_(m <= i < n) F%O) : order_scope. +Notation "\join^p_ ( i < n | P ) F" := + (\big[join/0]_(i < n | P%B) F%O) : order_scope. +Notation "\join^p_ ( i < n ) F" := + (\big[join/0]_(i < n) F%O) : order_scope. +Notation "\join^p_ ( i 'in' A | P ) F" := + (\big[join/0]_(i in A | P%B) F%O) : order_scope. +Notation "\join^p_ ( i 'in' A ) F" := + (\big[join/0]_(i in A) F%O) : order_scope. + +Notation "\meet^p_ ( i <- r | P ) F" := + (\big[meet/1]_(i <- r | P%B) F%O) : order_scope. +Notation "\meet^p_ ( i <- r ) F" := + (\big[meet/1]_(i <- r) F%O) : order_scope. +Notation "\meet^p_ ( i | P ) F" := + (\big[meet/1]_(i | P%B) F%O) : order_scope. +Notation "\meet^p_ i F" := + (\big[meet/1]_i F%O) : order_scope. +Notation "\meet^p_ ( i : I | P ) F" := + (\big[meet/1]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\meet^p_ ( i : I ) F" := + (\big[meet/1]_(i : I) F%O) (only parsing) : order_scope. +Notation "\meet^p_ ( m <= i < n | P ) F" := + (\big[meet/1]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\meet^p_ ( m <= i < n ) F" := + (\big[meet/1]_(m <= i < n) F%O) : order_scope. +Notation "\meet^p_ ( i < n | P ) F" := + (\big[meet/1]_(i < n | P%B) F%O) : order_scope. +Notation "\meet^p_ ( i < n ) F" := + (\big[meet/1]_(i < n) F%O) : order_scope. +Notation "\meet^p_ ( i 'in' A | P ) F" := + (\big[meet/1]_(i in A | P%B) F%O) : order_scope. +Notation "\meet^p_ ( i 'in' A ) F" := + (\big[meet/1]_(i in A) F%O) : order_scope. + +End ProdSyntax. + +(*******************************) +(* Definition of lexi_display. *) +(*******************************) + +Fact lexi_display : unit. Proof. by []. Qed. + +Module Import LexiSyntax. + +Notation "<=^l%O" := (@le lexi_display _) : fun_scope. +Notation ">=^l%O" := (@ge lexi_display _) : fun_scope. +Notation ">=^l%O" := (@ge lexi_display _) : fun_scope. +Notation "<^l%O" := (@lt lexi_display _) : fun_scope. +Notation ">^l%O" := (@gt lexi_display _) : fun_scope. +Notation "<?=^l%O" := (@leif lexi_display _) : fun_scope. +Notation ">=<^l%O" := (@comparable lexi_display _) : fun_scope. +Notation "><^l%O" := (fun x y => ~~ (@comparable lexi_display _ x y)) : + fun_scope. + +Notation "<=^l y" := (>=^l%O y) : order_scope. +Notation "<=^l y :> T" := (<=^l (y : T)) (only parsing) : order_scope. +Notation ">=^l y" := (<=^l%O y) : order_scope. +Notation ">=^l y :> T" := (>=^l (y : T)) (only parsing) : order_scope. + +Notation "<^l y" := (>^l%O y) : order_scope. +Notation "<^l y :> T" := (<^l (y : T)) (only parsing) : order_scope. +Notation ">^l y" := (<^l%O y) : order_scope. +Notation ">^l y :> T" := (>^l (y : T)) (only parsing) : order_scope. + +Notation ">=<^l y" := (>=<^l%O y) : order_scope. +Notation ">=<^l y :> T" := (>=<^l (y : T)) (only parsing) : order_scope. + +Notation "x <=^l y" := (<=^l%O x y) : order_scope. +Notation "x <=^l y :> T" := ((x : T) <=^l (y : T)) (only parsing) : order_scope. +Notation "x >=^l y" := (y <=^l x) (only parsing) : order_scope. +Notation "x >=^l y :> T" := ((x : T) >=^l (y : T)) (only parsing) : order_scope. + +Notation "x <^l y" := (<^l%O x y) : order_scope. +Notation "x <^l y :> T" := ((x : T) <^l (y : T)) (only parsing) : order_scope. +Notation "x >^l y" := (y <^l x) (only parsing) : order_scope. +Notation "x >^l y :> T" := ((x : T) >^l (y : T)) (only parsing) : order_scope. + +Notation "x <=^l y <=^l z" := ((x <=^l y) && (y <=^l z)) : order_scope. +Notation "x <^l y <=^l z" := ((x <^l y) && (y <=^l z)) : order_scope. +Notation "x <=^l y <^l z" := ((x <=^l y) && (y <^l z)) : order_scope. +Notation "x <^l y <^l z" := ((x <^l y) && (y <^l z)) : order_scope. + +Notation "x <=^l y ?= 'iff' C" := (<?=^l%O x y C) : order_scope. +Notation "x <=^l y ?= 'iff' C :> T" := ((x : T) <=^l (y : T) ?= iff C) + (only parsing) : order_scope. + +Notation ">=<^l x" := (>=<^l%O x) : order_scope. +Notation ">=<^l x :> T" := (>=<^l (x : T)) (only parsing) : order_scope. +Notation "x >=<^l y" := (>=<^l%O x y) : order_scope. + +Notation "><^l x" := (fun y => ~~ (>=<^l%O x y)) : order_scope. +Notation "><^l x :> T" := (><^l (x : T)) (only parsing) : order_scope. +Notation "x ><^l y" := (~~ (><^l%O x y)) : order_scope. + +Notation minlexi := (@meet lexi_display _). +Notation maxlexi := (@join lexi_display _). + +Notation "x `&^l` y" := (minlexi x y) : order_scope. +Notation "x `|^l` y" := (maxlexi x y) : order_scope. + +Notation "\max^l_ ( i <- r | P ) F" := + (\big[maxlexi/0]_(i <- r | P%B) F%O) : order_scope. +Notation "\max^l_ ( i <- r ) F" := + (\big[maxlexi/0]_(i <- r) F%O) : order_scope. +Notation "\max^l_ ( i | P ) F" := + (\big[maxlexi/0]_(i | P%B) F%O) : order_scope. +Notation "\max^l_ i F" := + (\big[maxlexi/0]_i F%O) : order_scope. +Notation "\max^l_ ( i : I | P ) F" := + (\big[maxlexi/0]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\max^l_ ( i : I ) F" := + (\big[maxlexi/0]_(i : I) F%O) (only parsing) : order_scope. +Notation "\max^l_ ( m <= i < n | P ) F" := + (\big[maxlexi/0]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\max^l_ ( m <= i < n ) F" := + (\big[maxlexi/0]_(m <= i < n) F%O) : order_scope. +Notation "\max^l_ ( i < n | P ) F" := + (\big[maxlexi/0]_(i < n | P%B) F%O) : order_scope. +Notation "\max^l_ ( i < n ) F" := + (\big[maxlexi/0]_(i < n) F%O) : order_scope. +Notation "\max^l_ ( i 'in' A | P ) F" := + (\big[maxlexi/0]_(i in A | P%B) F%O) : order_scope. +Notation "\max^l_ ( i 'in' A ) F" := + (\big[maxlexi/0]_(i in A) F%O) : order_scope. + +Notation "\min^l_ ( i <- r | P ) F" := + (\big[minlexi/1]_(i <- r | P%B) F%O) : order_scope. +Notation "\min^l_ ( i <- r ) F" := + (\big[minlexi/1]_(i <- r) F%O) : order_scope. +Notation "\min^l_ ( i | P ) F" := + (\big[minlexi/1]_(i | P%B) F%O) : order_scope. +Notation "\min^l_ i F" := + (\big[minlexi/1]_i F%O) : order_scope. +Notation "\min^l_ ( i : I | P ) F" := + (\big[minlexi/1]_(i : I | P%B) F%O) (only parsing) : order_scope. +Notation "\min^l_ ( i : I ) F" := + (\big[minlexi/1]_(i : I) F%O) (only parsing) : order_scope. +Notation "\min^l_ ( m <= i < n | P ) F" := + (\big[minlexi/1]_(m <= i < n | P%B) F%O) : order_scope. +Notation "\min^l_ ( m <= i < n ) F" := + (\big[minlexi/1]_(m <= i < n) F%O) : order_scope. +Notation "\min^l_ ( i < n | P ) F" := + (\big[minlexi/1]_(i < n | P%B) F%O) : order_scope. +Notation "\min^l_ ( i < n ) F" := + (\big[minlexi/1]_(i < n) F%O) : order_scope. +Notation "\min^l_ ( i 'in' A | P ) F" := + (\big[minlexi/1]_(i in A | P%B) F%O) : order_scope. +Notation "\min^l_ ( i 'in' A ) F" := + (\big[minlexi/1]_(i in A) F%O) : order_scope. + +End LexiSyntax. + +(*************************************************) +(* We declare a "copy" of the cartesian product, *) +(* which has canonical product order. *) +(*************************************************) + +Module ProdOrder. +Section ProdOrder. + +Definition type (disp : unit) (T T' : Type) := (T * T')%type. + +Context {disp1 disp2 disp3 : unit}. + +Local Notation "T * T'" := (type disp3 T T') : type_scope. + +Canonical eqType (T T' : eqType):= Eval hnf in [eqType of T * T']. +Canonical choiceType (T T' : choiceType):= Eval hnf in [choiceType of T * T']. +Canonical countType (T T' : countType):= Eval hnf in [countType of T * T']. +Canonical finType (T T' : finType):= Eval hnf in [finType of T * T']. + +Section POrder. +Variable (T : porderType disp1) (T' : porderType disp2). +Implicit Types (x y : T * T'). + +Definition le x y := (x.1 <= y.1) && (x.2 <= y.2). + +Fact refl : reflexive le. +Proof. by move=> ?; rewrite /le !lexx. Qed. + +Fact anti : antisymmetric le. +Proof. +case=> [? ?] [? ?]. +by rewrite andbAC andbA andbAC -andbA => /= /andP [] /le_anti -> /le_anti ->. +Qed. + +Fact trans : transitive le. +Proof. +rewrite /le => y x z /andP [] hxy ? /andP [] /(le_trans hxy) ->. +by apply: le_trans. +Qed. + +Definition porderMixin := LePOrderMixin (rrefl _) refl anti trans. +Canonical porderType := POrderType disp3 (T * T') porderMixin. + +Lemma leEprod x y : (x <= y) = (x.1 <= y.1) && (x.2 <= y.2). +Proof. by []. Qed. + +Lemma ltEprod x y : (x < y) = [&& x != y, x.1 <= y.1 & x.2 <= y.2]. +Proof. by rewrite lt_neqAle. Qed. + +Lemma le_pair (x1 y1 : T) (x2 y2 : T') : + (x1, x2) <= (y1, y2) :> T * T' = (x1 <= y1) && (x2 <= y2). +Proof. by []. Qed. + +Lemma lt_pair (x1 y1 : T) (x2 y2 : T') : (x1, x2) < (y1, y2) :> T * T' = + [&& (x1 != y1) || (x2 != y2), x1 <= y1 & x2 <= y2]. +Proof. by rewrite ltEprod negb_and. Qed. + +End POrder. + +Section DistrLattice. +Variable (T : distrLatticeType disp1) (T' : distrLatticeType disp2). +Implicit Types (x y : T * T'). + +Definition meet x y := (x.1 `&` y.1, x.2 `&` y.2). +Definition join x y := (x.1 `|` y.1, x.2 `|` y.2). + +Fact meetC : commutative meet. +Proof. by move=> ? ?; congr pair; rewrite meetC. Qed. + +Fact joinC : commutative join. +Proof. by move=> ? ?; congr pair; rewrite joinC. Qed. + +Fact meetA : associative meet. +Proof. by move=> ? ? ?; congr pair; rewrite meetA. Qed. + +Fact joinA : associative join. +Proof. by move=> ? ? ?; congr pair; rewrite joinA. Qed. + +Fact joinKI y x : meet x (join x y) = x. +Proof. by case: x => ? ?; congr pair; rewrite joinKI. Qed. + +Fact meetKU y x : join x (meet x y) = x. +Proof. by case: x => ? ?; congr pair; rewrite meetKU. Qed. + +Fact leEmeet x y : (x <= y) = (meet x y == x). +Proof. by rewrite eqE /= -!leEmeet. Qed. + +Fact meetUl : left_distributive meet join. +Proof. by move=> ? ? ?; congr pair; rewrite meetUl. Qed. + +Definition distrLatticeMixin := + DistrLattice.Mixin meetC joinC meetA joinA joinKI meetKU leEmeet meetUl. +Canonical distrLatticeType := DistrLatticeType (T * T') distrLatticeMixin. + +Lemma meetEprod x y : x `&` y = (x.1 `&` y.1, x.2 `&` y.2). +Proof. by []. Qed. + +Lemma joinEprod x y : x `|` y = (x.1 `|` y.1, x.2 `|` y.2). +Proof. by []. Qed. + +End DistrLattice. + +Section BDistrLattice. +Variable (T : bDistrLatticeType disp1) (T' : bDistrLatticeType disp2). + +Fact le0x (x : T * T') : (0, 0) <= x :> T * T'. +Proof. by rewrite /<=%O /= /le !le0x. Qed. + +Canonical bDistrLatticeType := + BDistrLatticeType (T * T') (BDistrLattice.Mixin le0x). + +Lemma botEprod : 0 = (0, 0) :> T * T'. Proof. by []. Qed. + +End BDistrLattice. + +Section TBDistrLattice. +Variable (T : tbDistrLatticeType disp1) (T' : tbDistrLatticeType disp2). + +Fact lex1 (x : T * T') : x <= (top, top). +Proof. by rewrite /<=%O /= /le !lex1. Qed. + +Canonical tbDistrLatticeType := + TBDistrLatticeType (T * T') (TBDistrLattice.Mixin lex1). + +Lemma topEprod : 1 = (1, 1) :> T * T'. Proof. by []. Qed. + +End TBDistrLattice. + +Section CBDistrLattice. +Variable (T : cbDistrLatticeType disp1) (T' : cbDistrLatticeType disp2). +Implicit Types (x y : T * T'). + +Definition sub x y := (x.1 `\` y.1, x.2 `\` y.2). + +Lemma subKI x y : y `&` sub x y = 0. +Proof. by congr pair; rewrite subKI. Qed. + +Lemma joinIB x y : x `&` y `|` sub x y = x. +Proof. by case: x => ? ?; congr pair; rewrite joinIB. Qed. + +Definition cbDistrLatticeMixin := CBDistrLattice.Mixin subKI joinIB. +Canonical cbDistrLatticeType := CBDistrLatticeType (T * T') cbDistrLatticeMixin. + +Lemma subEprod x y : x `\` y = (x.1 `\` y.1, x.2 `\` y.2). +Proof. by []. Qed. + +End CBDistrLattice. + +Section CTBDistrLattice. +Variable (T : ctbDistrLatticeType disp1) (T' : ctbDistrLatticeType disp2). +Implicit Types (x y : T * T'). + +Definition compl x : T * T' := (~` x.1, ~` x.2). + +Lemma complE x : compl x = sub 1 x. +Proof. by congr pair; rewrite complE. Qed. + +Definition ctbDistrLatticeMixin := CTBDistrLattice.Mixin complE. +Canonical ctbDistrLatticeType := + CTBDistrLatticeType (T * T') ctbDistrLatticeMixin. + +Lemma complEprod x : ~` x = (~` x.1, ~` x.2). Proof. by []. Qed. + +End CTBDistrLattice. + +Canonical finPOrderType (T : finPOrderType disp1) + (T' : finPOrderType disp2) := [finPOrderType of T * T']. + +Canonical finDistrLatticeType (T : finDistrLatticeType disp1) + (T' : finDistrLatticeType disp2) := [finDistrLatticeType of T * T']. + +Canonical finCDistrLatticeType (T : finCDistrLatticeType disp1) + (T' : finCDistrLatticeType disp2) := [finCDistrLatticeType of T * T']. + +End ProdOrder. + +Module Exports. + +Notation "T *prod[ d ] T'" := (type d T T') + (at level 70, d at next level, format "T *prod[ d ] T'") : type_scope. +Notation "T *p T'" := (type prod_display T T') + (at level 70, format "T *p T'") : type_scope. + +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical cbDistrLatticeType. +Canonical ctbDistrLatticeType. +Canonical finPOrderType. +Canonical finDistrLatticeType. +Canonical finCDistrLatticeType. + +Definition leEprod := @leEprod. +Definition ltEprod := @ltEprod. +Definition le_pair := @le_pair. +Definition lt_pair := @lt_pair. +Definition meetEprod := @meetEprod. +Definition joinEprod := @joinEprod. +Definition botEprod := @botEprod. +Definition topEprod := @topEprod. +Definition subEprod := @subEprod. +Definition complEprod := @complEprod. + +End Exports. +End ProdOrder. +Import ProdOrder.Exports. + +Module DefaultProdOrder. +Section DefaultProdOrder. +Context {disp1 disp2 : unit}. + +Canonical prod_porderType (T : porderType disp1) (T' : porderType disp2) := + [porderType of T * T' for [porderType of T *p T']]. +Canonical prod_distrLatticeType + (T : distrLatticeType disp1) (T' : distrLatticeType disp2) := + [distrLatticeType of T * T' for [distrLatticeType of T *p T']]. +Canonical prod_bDistrLatticeType + (T : bDistrLatticeType disp1) (T' : bDistrLatticeType disp2) := + [bDistrLatticeType of T * T' for [bDistrLatticeType of T *p T']]. +Canonical prod_tbDistrLatticeType + (T : tbDistrLatticeType disp1) (T' : tbDistrLatticeType disp2) := + [tbDistrLatticeType of T * T' for [tbDistrLatticeType of T *p T']]. +Canonical prod_cbDistrLatticeType + (T : cbDistrLatticeType disp1) (T' : cbDistrLatticeType disp2) := + [cbDistrLatticeType of T * T' for [cbDistrLatticeType of T *p T']]. +Canonical prod_ctbDistrLatticeType + (T : ctbDistrLatticeType disp1) (T' : ctbDistrLatticeType disp2) := + [ctbDistrLatticeType of T * T' for [ctbDistrLatticeType of T *p T']]. +Canonical prod_finPOrderType (T : finPOrderType disp1) + (T' : finPOrderType disp2) := [finPOrderType of T * T']. +Canonical prod_finDistrLatticeType (T : finDistrLatticeType disp1) + (T' : finDistrLatticeType disp2) := [finDistrLatticeType of T * T']. +Canonical prod_finCDistrLatticeType (T : finCDistrLatticeType disp1) + (T' : finCDistrLatticeType disp2) := [finCDistrLatticeType of T * T']. + +End DefaultProdOrder. +End DefaultProdOrder. + +(********************************************************) +(* We declare lexicographic ordering on dependent pairs *) +(********************************************************) + +Module SigmaOrder. +Section SigmaOrder. + +Context {disp1 disp2 : unit}. + +Section POrder. + +Variable (T : porderType disp1) (T' : T -> porderType disp2). +Implicit Types (x y : {t : T & T' t}). + +Definition le x y := (tag x <= tag y) && + ((tag x >= tag y) ==> (tagged x <= tagged_as x y)). +Definition lt x y := (tag x <= tag y) && + ((tag x >= tag y) ==> (tagged x < tagged_as x y)). + +Fact refl : reflexive le. +Proof. by move=> [x x']; rewrite /le tagged_asE/= !lexx. Qed. + +Fact anti : antisymmetric le. +Proof. +rewrite /le => -[x x'] [y y']/=; case: comparableP => //= eq_xy. +by case: _ / eq_xy in y' *; rewrite !tagged_asE => /le_anti ->. +Qed. + +Fact trans : transitive le. +Proof. +move=> [y y'] [x x'] [z z'] /andP[/= lexy lexy'] /andP[/= leyz leyz']. +rewrite /= /le (le_trans lexy) //=; apply/implyP => lezx. +elim: _ / (@le_anti _ _ x y) in y' z' lexy' leyz' *; last first. + by rewrite lexy (le_trans leyz). +elim: _ / (@le_anti _ _ x z) in z' leyz' *; last by rewrite (le_trans lexy). +by rewrite lexx !tagged_asE/= in lexy' leyz' *; rewrite (le_trans lexy'). +Qed. + +Fact lt_def x y : lt x y = (y != x) && le x y. +Proof. +rewrite /lt /le; case: x y => [x x'] [y y']//=; rewrite andbCA. +case: (comparableP x y) => //= xy; last first. + by case: _ / xy in y' *; rewrite !tagged_asE eq_Tagged/= lt_def. +by rewrite andbT; symmetry; apply: contraTneq xy => -[yx _]; rewrite yx ltxx. +Qed. + +Definition porderMixin := LePOrderMixin lt_def refl anti trans. +Canonical porderType := POrderType disp2 {t : T & T' t} porderMixin. + +Lemma leEsig x y : x <= y = + (tag x <= tag y) && ((tag x >= tag y) ==> (tagged x <= tagged_as x y)). +Proof. by []. Qed. + +Lemma ltEsig x y : x < y = + (tag x <= tag y) && ((tag x >= tag y) ==> (tagged x < tagged_as x y)). +Proof. by []. Qed. + +Lemma le_Taggedl x (u : T' (tag x)) : (Tagged T' u <= x) = (u <= tagged x). +Proof. by case: x => [t v]/= in u *; rewrite leEsig/= lexx/= tagged_asE. Qed. + +Lemma le_Taggedr x (u : T' (tag x)) : (x <= Tagged T' u) = (tagged x <= u). +Proof. by case: x => [t v]/= in u *; rewrite leEsig/= lexx/= tagged_asE. Qed. + +Lemma lt_Taggedl x (u : T' (tag x)) : (Tagged T' u < x) = (u < tagged x). +Proof. by case: x => [t v]/= in u *; rewrite ltEsig/= lexx/= tagged_asE. Qed. + +Lemma lt_Taggedr x (u : T' (tag x)) : (x < Tagged T' u) = (tagged x < u). +Proof. by case: x => [t v]/= in u *; rewrite ltEsig/= lexx/= tagged_asE. Qed. + +End POrder. + +Section Total. +Variable (T : orderType disp1) (T' : T -> orderType disp2). +Implicit Types (x y : {t : T & T' t}). + +Fact total : totalPOrderMixin [porderType of {t : T & T' t}]. +Proof. +move=> x y; rewrite !leEsig; case: (ltgtP (tag x) (tag y)) => //=. +case: x y => [x x'] [y y']/= eqxy; elim: _ /eqxy in y' *. +by rewrite !tagged_asE le_total. +Qed. + +Canonical distrLatticeType := DistrLatticeType {t : T & T' t} total. +Canonical orderType := OrderType {t : T & T' t} total. + +End Total. + +Section FinDistrLattice. +Variable (T : finOrderType disp1) (T' : T -> finOrderType disp2). + +Fact le0x (x : {t : T & T' t}) : Tagged T' (0 : T' 0) <= x. +Proof. +rewrite leEsig /=; case: comparableP (le0x (tag x)) => //=. +by case: x => //= x px x0; rewrite x0 in px *; rewrite tagged_asE le0x. +Qed. +Canonical bDistrLatticeType := + BDistrLatticeType {t : T & T' t} (BDistrLattice.Mixin le0x). + +Lemma botEsig : 0 = Tagged T' (0 : T' 0). Proof. by []. Qed. + +Fact lex1 (x : {t : T & T' t}) : x <= Tagged T' (1 : T' 1). +Proof. +rewrite leEsig /=; case: comparableP (lex1 (tag x)) => //=. +by case: x => //= x px x0; rewrite x0 in px *; rewrite tagged_asE lex1. +Qed. +Canonical tbDistrLatticeType := + TBDistrLatticeType {t : T & T' t} (TBDistrLattice.Mixin lex1). + +Lemma topEsig : 1 = Tagged T' (1 : T' 1). Proof. by []. Qed. + +End FinDistrLattice. + +Canonical finPOrderType (T : finPOrderType disp1) + (T' : T -> finPOrderType disp2) := [finPOrderType of {t : T & T' t}]. +Canonical finDistrLatticeType (T : finOrderType disp1) + (T' : T -> finOrderType disp2) := [finDistrLatticeType of {t : T & T' t}]. +Canonical finOrderType (T : finOrderType disp1) + (T' : T -> finOrderType disp2) := [finOrderType of {t : T & T' t}]. + +End SigmaOrder. + +Module Exports. + +Canonical porderType. +Canonical distrLatticeType. +Canonical orderType. +Canonical finPOrderType. +Canonical finDistrLatticeType. +Canonical finOrderType. + +Definition leEsig := @leEsig. +Definition ltEsig := @ltEsig. +Definition le_Taggedl := @le_Taggedl. +Definition lt_Taggedl := @lt_Taggedl. +Definition le_Taggedr := @le_Taggedr. +Definition lt_Taggedr := @lt_Taggedr. +Definition topEsig := @topEsig. +Definition botEsig := @botEsig. + +End Exports. +End SigmaOrder. +Import SigmaOrder.Exports. + +(*************************************************) +(* We declare a "copy" of the cartesian product, *) +(* which has canonical lexicographic order. *) +(*************************************************) + +Module ProdLexiOrder. +Section ProdLexiOrder. + +Definition type (disp : unit) (T T' : Type) := (T * T')%type. + +Context {disp1 disp2 disp3 : unit}. + +Local Notation "T * T'" := (type disp3 T T') : type_scope. + +Canonical eqType (T T' : eqType):= Eval hnf in [eqType of T * T']. +Canonical choiceType (T T' : choiceType):= Eval hnf in [choiceType of T * T']. +Canonical countType (T T' : countType):= Eval hnf in [countType of T * T']. +Canonical finType (T T' : finType):= Eval hnf in [finType of T * T']. + +Section POrder. +Variable (T : porderType disp1) (T' : porderType disp2). + +Implicit Types (x y : T * T'). + +Definition le x y := (x.1 <= y.1) && ((x.1 >= y.1) ==> (x.2 <= y.2)). +Definition lt x y := (x.1 <= y.1) && ((x.1 >= y.1) ==> (x.2 < y.2)). + +Fact refl : reflexive le. +Proof. by move=> ?; rewrite /le !lexx. Qed. + +Fact anti : antisymmetric le. +Proof. +by rewrite /le => -[x x'] [y y'] /=; case: comparableP => //= -> /le_anti->. +Qed. + +Fact trans : transitive le. +Proof. +move=> y x z /andP [hxy /implyP hxy'] /andP [hyz /implyP hyz']. +rewrite /le (le_trans hxy) //=; apply/implyP => hzx. +by apply/le_trans/hxy'/(le_trans hyz): (hyz' (le_trans hzx hxy)). +Qed. + +Fact lt_def x y : lt x y = (y != x) && le x y. +Proof. +rewrite /lt /le; case: x y => [x1 x2] [y1 y2]//=; rewrite xpair_eqE. +by case: (comparableP x1 y1); rewrite lt_def. +Qed. + +Definition porderMixin := LePOrderMixin lt_def refl anti trans. +Canonical porderType := POrderType disp3 (T * T') porderMixin. + +Lemma leEprodlexi x y : + (x <= y) = (x.1 <= y.1) && ((x.1 >= y.1) ==> (x.2 <= y.2)). +Proof. by []. Qed. + +Lemma ltEprodlexi x y : + (x < y) = (x.1 <= y.1) && ((x.1 >= y.1) ==> (x.2 < y.2)). +Proof. by []. Qed. + +Lemma lexi_pair (x1 y1 : T) (x2 y2 : T') : + (x1, x2) <= (y1, y2) :> T * T' = (x1 <= y1) && ((x1 >= y1) ==> (x2 <= y2)). +Proof. by []. Qed. + +Lemma ltxi_pair (x1 y1 : T) (x2 y2 : T') : + (x1, x2) < (y1, y2) :> T * T' = (x1 <= y1) && ((x1 >= y1) ==> (x2 < y2)). +Proof. by []. Qed. + +End POrder. + +Section Total. +Variable (T : orderType disp1) (T' : orderType disp2). +Implicit Types (x y : T * T'). + +Fact total : totalPOrderMixin [porderType of T * T']. +Proof. +move=> x y; rewrite /<=%O /= /le; case: ltgtP => //= _; exact: le_total. +Qed. + +Canonical distrLatticeType := DistrLatticeType (T * T') total. +Canonical orderType := OrderType (T * T') total. + +End Total. + +Section FinDistrLattice. +Variable (T : finOrderType disp1) (T' : finOrderType disp2). + +Fact le0x (x : T * T') : (0, 0) <= x :> T * T'. +Proof. by case: x => // x1 x2; rewrite leEprodlexi/= !le0x implybT. Qed. +Canonical bDistrLatticeType := + BDistrLatticeType (T * T') (BDistrLattice.Mixin le0x). + +Lemma botEprodlexi : 0 = (0, 0) :> T * T'. Proof. by []. Qed. + +Fact lex1 (x : T * T') : x <= (1, 1) :> T * T'. +Proof. by case: x => // x1 x2; rewrite leEprodlexi/= !lex1 implybT. Qed. +Canonical tbDistrLatticeType := + TBDistrLatticeType (T * T') (TBDistrLattice.Mixin lex1). + +Lemma topEprodlexi : 1 = (1, 1) :> T * T'. Proof. by []. Qed. + +End FinDistrLattice. + +Canonical finPOrderType (T : finPOrderType disp1) + (T' : finPOrderType disp2) := [finPOrderType of T * T']. +Canonical finDistrLatticeType (T : finOrderType disp1) + (T' : finOrderType disp2) := [finDistrLatticeType of T * T']. +Canonical finOrderType (T : finOrderType disp1) + (T' : finOrderType disp2) := [finOrderType of T * T']. + +Lemma sub_prod_lexi d (T : POrder.Exports.porderType disp1) + (T' : POrder.Exports.porderType disp2) : + subrel (<=%O : rel (T *prod[d] T')) (<=%O : rel (T * T')). +Proof. +by case=> [x1 x2] [y1 y2]; rewrite leEprod leEprodlexi /=; case: comparableP. +Qed. + +End ProdLexiOrder. + +Module Exports. + +Notation "T *lexi[ d ] T'" := (type d T T') + (at level 70, d at next level, format "T *lexi[ d ] T'") : type_scope. +Notation "T *l T'" := (type lexi_display T T') + (at level 70, format "T *l T'") : type_scope. + +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical distrLatticeType. +Canonical orderType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical finPOrderType. +Canonical finDistrLatticeType. +Canonical finOrderType. + +Definition leEprodlexi := @leEprodlexi. +Definition ltEprodlexi := @ltEprodlexi. +Definition lexi_pair := @lexi_pair. +Definition ltxi_pair := @ltxi_pair. +Definition topEprodlexi := @topEprodlexi. +Definition botEprodlexi := @botEprodlexi. +Definition sub_prod_lexi := @sub_prod_lexi. + +End Exports. +End ProdLexiOrder. +Import ProdLexiOrder.Exports. + +Module DefaultProdLexiOrder. +Section DefaultProdLexiOrder. +Context {disp1 disp2 : unit}. + +Canonical prodlexi_porderType + (T : porderType disp1) (T' : porderType disp2) := + [porderType of T * T' for [porderType of T *l T']]. +Canonical prodlexi_distrLatticeType + (T : orderType disp1) (T' : orderType disp2) := + [distrLatticeType of T * T' for [distrLatticeType of T *l T']]. +Canonical prodlexi_orderType + (T : orderType disp1) (T' : orderType disp2) := + [orderType of T * T' for [orderType of T *l T']]. +Canonical prodlexi_bDistrLatticeType + (T : finOrderType disp1) (T' : finOrderType disp2) := + [bDistrLatticeType of T * T' for [bDistrLatticeType of T *l T']]. +Canonical prodlexi_tbDistrLatticeType + (T : finOrderType disp1) (T' : finOrderType disp2) := + [tbDistrLatticeType of T * T' for [tbDistrLatticeType of T *l T']]. +Canonical prodlexi_finPOrderType (T : finPOrderType disp1) + (T' : finPOrderType disp2) := [finPOrderType of T * T']. +Canonical prodlexi_finDistrLatticeType (T : finOrderType disp1) + (T' : finOrderType disp2) := [finDistrLatticeType of T * T']. +Canonical prodlexi_finOrderType (T : finOrderType disp1) + (T' : finOrderType disp2) := [finOrderType of T * T']. + +End DefaultProdLexiOrder. +End DefaultProdLexiOrder. + +(*****************************************) +(* We declare a "copy" of the sequences, *) +(* which has canonical product order. *) +(*****************************************) + +Module SeqProdOrder. +Section SeqProdOrder. + +Definition type (disp : unit) T := seq T. + +Context {disp disp' : unit}. + +Local Notation seq := (type disp'). + +Canonical eqType (T : eqType):= Eval hnf in [eqType of seq T]. +Canonical choiceType (T : choiceType):= Eval hnf in [choiceType of seq T]. +Canonical countType (T : countType):= Eval hnf in [countType of seq T]. + +Section POrder. +Variable T : porderType disp. +Implicit Types s : seq T. + +Fixpoint le s1 s2 := if s1 isn't x1 :: s1' then true else + if s2 isn't x2 :: s2' then false else + (x1 <= x2) && le s1' s2'. + +Fact refl : reflexive le. Proof. by elim=> //= ? ? ?; rewrite !lexx. Qed. + +Fact anti : antisymmetric le. +Proof. +by elim=> [|x s ihs] [|y s'] //=; rewrite andbACA => /andP[/le_anti-> /ihs->]. +Qed. + +Fact trans : transitive le. +Proof. +elim=> [|y ys ihs] [|x xs] [|z zs] //= /andP[xy xys] /andP[yz yzs]. +by rewrite (le_trans xy)// ihs. +Qed. + +Definition porderMixin := LePOrderMixin (rrefl _) refl anti trans. +Canonical porderType := POrderType disp' (seq T) porderMixin. + +Lemma leEseq s1 s2 : s1 <= s2 = if s1 isn't x1 :: s1' then true else + if s2 isn't x2 :: s2' then false else + (x1 <= x2) && (s1' <= s2' :> seq _). +Proof. by case: s1. Qed. + +Lemma le0s s : [::] <= s :> seq _. Proof. by []. Qed. + +Lemma les0 s : s <= [::] = (s == [::]). Proof. by rewrite leEseq. Qed. + +Lemma le_cons x1 s1 x2 s2 : + x1 :: s1 <= x2 :: s2 :> seq _ = (x1 <= x2) && (s1 <= s2). +Proof. by []. Qed. + +End POrder. + +Section BDistrLattice. +Variable T : distrLatticeType disp. +Implicit Types s : seq T. + +Fixpoint meet s1 s2 := + match s1, s2 with + | x1 :: s1', x2 :: s2' => (x1 `&` x2) :: meet s1' s2' + | _, _ => [::] + end. + +Fixpoint join s1 s2 := + match s1, s2 with + | [::], _ => s2 | _, [::] => s1 + | x1 :: s1', x2 :: s2' => (x1 `|` x2) :: join s1' s2' + end. + +Fact meetC : commutative meet. +Proof. by elim=> [|? ? ih] [|? ?] //=; rewrite meetC ih. Qed. + +Fact joinC : commutative join. +Proof. by elim=> [|? ? ih] [|? ?] //=; rewrite joinC ih. Qed. + +Fact meetA : associative meet. +Proof. by elim=> [|? ? ih] [|? ?] [|? ?] //=; rewrite meetA ih. Qed. + +Fact joinA : associative join. +Proof. by elim=> [|? ? ih] [|? ?] [|? ?] //=; rewrite joinA ih. Qed. + +Fact meetss s : meet s s = s. +Proof. by elim: s => [|? ? ih] //=; rewrite meetxx ih. Qed. + +Fact joinKI y x : meet x (join x y) = x. +Proof. +elim: x y => [|? ? ih] [|? ?] //=; rewrite (meetxx, joinKI) ?ih //. +by congr cons; rewrite meetss. +Qed. + +Fact meetKU y x : join x (meet x y) = x. +Proof. by elim: x y => [|? ? ih] [|? ?] //=; rewrite meetKU ih. Qed. + +Fact leEmeet x y : (x <= y) = (meet x y == x). +Proof. +by rewrite /<=%O /=; elim: x y => [|? ? ih] [|? ?] //=; rewrite eqE leEmeet ih. +Qed. + +Fact meetUl : left_distributive meet join. +Proof. by elim=> [|? ? ih] [|? ?] [|? ?] //=; rewrite meetUl ih. Qed. + +Definition distrLatticeMixin := + DistrLattice.Mixin meetC joinC meetA joinA joinKI meetKU leEmeet meetUl. +Canonical distrLatticeType := DistrLatticeType (seq T) distrLatticeMixin. +Canonical bDistrLatticeType := + BDistrLatticeType (seq T) (BDistrLattice.Mixin (@le0s _)). + +Lemma botEseq : 0 = [::] :> seq T. +Proof. by []. Qed. + +Lemma meetEseq s1 s2 : s1 `&` s2 = [seq x.1 `&` x.2 | x <- zip s1 s2]. +Proof. by elim: s1 s2 => [|x s1 ihs1] [|y s2]//=; rewrite -ihs1. Qed. + +Lemma meet_cons x1 s1 x2 s2 : + (x1 :: s1 : seq T) `&` (x2 :: s2) = (x1 `&` x2) :: s1 `&` s2. +Proof. by []. Qed. + +Lemma joinEseq s1 s2 : s1 `|` s2 = + match s1, s2 with + | [::], _ => s2 | _, [::] => s1 + | x1 :: s1', x2 :: s2' => (x1 `|` x2) :: ((s1' : seq _) `|` s2') + end. +Proof. by case: s1. Qed. + +Lemma join_cons x1 s1 x2 s2 : + (x1 :: s1 : seq T) `|` (x2 :: s2) = (x1 `|` x2) :: s1 `|` s2. +Proof. by []. Qed. + +End BDistrLattice. + +End SeqProdOrder. + +Module Exports. + +Notation seqprod_with := type. +Notation seqprod := (type prod_display). + +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. + +Definition leEseq := @leEseq. +Definition le0s := @le0s. +Definition les0 := @les0. +Definition le_cons := @le_cons. +Definition botEseq := @botEseq. +Definition meetEseq := @meetEseq. +Definition meet_cons := @meet_cons. +Definition joinEseq := @joinEseq. + +End Exports. +End SeqProdOrder. +Import SeqProdOrder.Exports. + +Module DefaultSeqProdOrder. +Section DefaultSeqProdOrder. +Context {disp : unit}. + +Canonical seqprod_porderType (T : porderType disp) := + [porderType of seq T for [porderType of seqprod T]]. +Canonical seqprod_distrLatticeType (T : distrLatticeType disp) := + [distrLatticeType of seq T for [distrLatticeType of seqprod T]]. +Canonical seqprod_bDistrLatticeType (T : bDistrLatticeType disp) := + [bDistrLatticeType of seq T for [bDistrLatticeType of seqprod T]]. + +End DefaultSeqProdOrder. +End DefaultSeqProdOrder. + +(*********************************************) +(* We declare a "copy" of the sequences, *) +(* which has canonical lexicographic order. *) +(*********************************************) + +Module SeqLexiOrder. +Section SeqLexiOrder. + +Definition type (disp : unit) T := seq T. + +Context {disp disp' : unit}. + +Local Notation seq := (type disp'). + +Canonical eqType (T : eqType):= Eval hnf in [eqType of seq T]. +Canonical choiceType (T : choiceType):= Eval hnf in [choiceType of seq T]. +Canonical countType (T : countType):= Eval hnf in [countType of seq T]. + +Section POrder. +Variable T : porderType disp. +Implicit Types s : seq T. + +Fixpoint le s1 s2 := if s1 isn't x1 :: s1' then true else + if s2 isn't x2 :: s2' then false else + (x1 <= x2) && ((x1 >= x2) ==> le s1' s2'). +Fixpoint lt s1 s2 := if s2 isn't x2 :: s2' then false else + if s1 isn't x1 :: s1' then true else + (x1 <= x2) && ((x1 >= x2) ==> lt s1' s2'). + +Fact refl: reflexive le. +Proof. by elim => [|x s ih] //=; rewrite lexx. Qed. + +Fact anti: antisymmetric le. +Proof. +move=> x y /andP []; elim: x y => [|x sx ih] [|y sy] //=. +by case: comparableP => //= -> lesxsy /(ih _ lesxsy) ->. +Qed. + +Fact trans: transitive le. +Proof. +elim=> [|y sy ihs] [|x sx] [|z sz] //=; case: (comparableP x y) => //= [xy|->]. + by move=> _ /andP[/(lt_le_trans xy) xz _]; rewrite (ltW xz)// lt_geF. +by case: comparableP => //= _; apply: ihs. +Qed. + +Lemma lt_def s1 s2 : lt s1 s2 = (s2 != s1) && le s1 s2. +Proof. +elim: s1 s2 => [|x s1 ihs1] [|y s2]//=. +by rewrite eqseq_cons ihs1; case: comparableP. +Qed. + +Definition porderMixin := LePOrderMixin lt_def refl anti trans. +Canonical porderType := POrderType disp' (seq T) porderMixin. + +Lemma leEseqlexi s1 s2 : + s1 <= s2 = if s1 isn't x1 :: s1' then true else + if s2 isn't x2 :: s2' then false else + (x1 <= x2) && ((x1 >= x2) ==> (s1' <= s2' :> seq T)). +Proof. by case: s1. Qed. + +Lemma ltEseqlexi s1 s2 : + s1 < s2 = if s2 isn't x2 :: s2' then false else + if s1 isn't x1 :: s1' then true else + (x1 <= x2) && ((x1 >= x2) ==> (s1' < s2' :> seq T)). +Proof. by case: s1. Qed. + +Lemma lexi0s s : [::] <= s :> seq T. Proof. by []. Qed. + +Lemma lexis0 s : s <= [::] = (s == [::]). Proof. by rewrite leEseqlexi. Qed. + +Lemma ltxi0s s : ([::] < s :> seq T) = (s != [::]). Proof. by case: s. Qed. + +Lemma ltxis0 s : s < [::] = false. Proof. by rewrite ltEseqlexi. Qed. + +Lemma lexi_cons x1 s1 x2 s2 : + x1 :: s1 <= x2 :: s2 :> seq T = (x1 <= x2) && ((x1 >= x2) ==> (s1 <= s2)). +Proof. by []. Qed. + +Lemma ltxi_cons x1 s1 x2 s2 : + x1 :: s1 < x2 :: s2 :> seq T = (x1 <= x2) && ((x1 >= x2) ==> (s1 < s2)). +Proof. by []. Qed. + +Lemma lexi_lehead x s1 y s2 : x :: s1 <= y :: s2 :> seq T -> x <= y. +Proof. by rewrite lexi_cons => /andP[]. Qed. + +Lemma ltxi_lehead x s1 y s2 : x :: s1 < y :: s2 :> seq T -> x <= y. +Proof. by rewrite ltxi_cons => /andP[]. Qed. + +Lemma eqhead_lexiE (x : T) s1 s2 : (x :: s1 <= x :: s2 :> seq _) = (s1 <= s2). +Proof. by rewrite lexi_cons lexx. Qed. + +Lemma eqhead_ltxiE (x : T) s1 s2 : (x :: s1 < x :: s2 :> seq _) = (s1 < s2). +Proof. by rewrite ltxi_cons lexx. Qed. + +Lemma neqhead_lexiE (x y : T) s1 s2 : x != y -> + (x :: s1 <= y :: s2 :> seq _) = (x < y). +Proof. by rewrite lexi_cons; case: comparableP. Qed. + +Lemma neqhead_ltxiE (x y : T) s1 s2 : x != y -> + (x :: s1 < y :: s2 :> seq _) = (x < y). +Proof. by rewrite ltxi_cons; case: (comparableP x y). Qed. + +End POrder. + +Section Total. +Variable T : orderType disp. +Implicit Types s : seq T. + +Fact total : totalPOrderMixin [porderType of seq T]. +Proof. +suff: total (<=%O : rel (seq T)) by []. +by elim=> [|x1 s1 ihs1] [|x2 s2]//=; rewrite !lexi_cons; case: ltgtP => /=. +Qed. + +Canonical distrLatticeType := DistrLatticeType (seq T) total. +Canonical bDistrLatticeType := + BDistrLatticeType (seq T) (BDistrLattice.Mixin (@lexi0s _)). +Canonical orderType := OrderType (seq T) total. + +End Total. + +Lemma sub_seqprod_lexi d (T : POrder.Exports.porderType disp) : + subrel (<=%O : rel (seqprod_with d T)) (<=%O : rel (seq T)). +Proof. +elim=> [|x1 s1 ihs1] [|x2 s2]//=; rewrite le_cons lexi_cons /=. +by move=> /andP[-> /ihs1->]; rewrite implybT. +Qed. + +End SeqLexiOrder. + +Module Exports. + +Notation seqlexi_with := type. +Notation seqlexi := (type lexi_display). + +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical orderType. + +Definition leEseqlexi := @leEseqlexi. +Definition lexi0s := @lexi0s. +Definition lexis0 := @lexis0. +Definition lexi_cons := @lexi_cons. +Definition lexi_lehead := @lexi_lehead. +Definition eqhead_lexiE := @eqhead_lexiE. +Definition neqhead_lexiE := @neqhead_lexiE. + +Definition ltEseqltxi := @ltEseqlexi. +Definition ltxi0s := @ltxi0s. +Definition ltxis0 := @ltxis0. +Definition ltxi_cons := @ltxi_cons. +Definition ltxi_lehead := @ltxi_lehead. +Definition eqhead_ltxiE := @eqhead_ltxiE. +Definition neqhead_ltxiE := @neqhead_ltxiE. +Definition sub_seqprod_lexi := @sub_seqprod_lexi. + +End Exports. +End SeqLexiOrder. +Import SeqLexiOrder.Exports. + +Module DefaultSeqLexiOrder. +Section DefaultSeqLexiOrder. +Context {disp : unit}. + +Canonical seqlexi_porderType (T : porderType disp) := + [porderType of seq T for [porderType of seqlexi T]]. +Canonical seqlexi_distrLatticeType (T : orderType disp) := + [distrLatticeType of seq T for [distrLatticeType of seqlexi T]]. +Canonical seqlexi_bDistrLatticeType (T : orderType disp) := + [bDistrLatticeType of seq T for [bDistrLatticeType of seqlexi T]]. +Canonical seqlexi_orderType (T : orderType disp) := + [orderType of seq T for [orderType of seqlexi T]]. + +End DefaultSeqLexiOrder. +End DefaultSeqLexiOrder. + +(***************************************) +(* We declare a "copy" of the tuples, *) +(* which has canonical product order. *) +(***************************************) + +Module TupleProdOrder. +Import DefaultSeqProdOrder. + +Section TupleProdOrder. + +Definition type (disp : unit) n T := n.-tuple T. + +Context {disp disp' : unit}. +Local Notation "n .-tuple" := (type disp' n) : type_scope. + +Section Basics. +Variable (n : nat). + +Canonical eqType (T : eqType):= Eval hnf in [eqType of n.-tuple T]. +Canonical choiceType (T : choiceType):= Eval hnf in [choiceType of n.-tuple T]. +Canonical countType (T : countType):= Eval hnf in [countType of n.-tuple T]. +Canonical finType (T : finType):= Eval hnf in [finType of n.-tuple T]. +End Basics. + +Section POrder. +Implicit Types (T : porderType disp). + +Definition porderMixin n T := [porderMixin of n.-tuple T by <:]. +Canonical porderType n T := POrderType disp' (n.-tuple T) (porderMixin n T). + +Lemma leEtprod n T (t1 t2 : n.-tuple T) : + t1 <= t2 = [forall i, tnth t1 i <= tnth t2 i]. +Proof. +elim: n => [|n IHn] in t1 t2 *. + by rewrite tuple0 [t2]tuple0/= lexx; symmetry; apply/forallP => []. +case: (tupleP t1) (tupleP t2) => [x1 {t1}t1] [x2 {t2}t2]. +rewrite [_ <= _]le_cons [t1 <= t2 :> seq _]IHn. +apply/idP/forallP => [/andP[lex12 /forallP/= let12 i]|lext12]. + by case: (unliftP ord0 i) => [j ->|->]//; rewrite !tnthS. +rewrite (lext12 ord0)/=; apply/forallP=> i. +by have := lext12 (lift ord0 i); rewrite !tnthS. +Qed. + +Lemma ltEtprod n T (t1 t2 : n.-tuple T) : + t1 < t2 = [exists i, tnth t1 i != tnth t2 i] && + [forall i, tnth t1 i <= tnth t2 i]. +Proof. by rewrite lt_neqAle leEtprod eqEtuple negb_forall. Qed. + +End POrder. + +Section DistrLattice. +Variables (n : nat) (T : distrLatticeType disp). +Implicit Types (t : n.-tuple T). + +Definition meet t1 t2 : n.-tuple T := + [tuple of [seq x.1 `&` x.2 | x <- zip t1 t2]]. +Definition join t1 t2 : n.-tuple T := + [tuple of [seq x.1 `|` x.2 | x <- zip t1 t2]]. + +Fact tnth_meet t1 t2 i : tnth (meet t1 t2) i = tnth t1 i `&` tnth t2 i. +Proof. +rewrite tnth_map -(tnth_map fst) -(tnth_map snd) -/unzip1 -/unzip2. +by rewrite !(tnth_nth (tnth_default t1 i))/= unzip1_zip ?unzip2_zip ?size_tuple. +Qed. + +Fact tnth_join t1 t2 i : tnth (join t1 t2) i = tnth t1 i `|` tnth t2 i. +Proof. +rewrite tnth_map -(tnth_map fst) -(tnth_map snd) -/unzip1 -/unzip2. +by rewrite !(tnth_nth (tnth_default t1 i))/= unzip1_zip ?unzip2_zip ?size_tuple. +Qed. + +Fact meetC : commutative meet. +Proof. by move=> t1 t2; apply: eq_from_tnth => i; rewrite !tnth_meet meetC. Qed. + +Fact joinC : commutative join. +Proof. by move=> t1 t2; apply: eq_from_tnth => i; rewrite !tnth_join joinC. Qed. + +Fact meetA : associative meet. +Proof. +by move=> t1 t2 t3; apply: eq_from_tnth => i; rewrite !tnth_meet meetA. +Qed. + +Fact joinA : associative join. +Proof. +by move=> t1 t2 t3; apply: eq_from_tnth => i; rewrite !tnth_join joinA. +Qed. + +Fact joinKI t2 t1 : meet t1 (join t1 t2) = t1. +Proof. by apply: eq_from_tnth => i; rewrite tnth_meet tnth_join joinKI. Qed. + +Fact meetKU y x : join x (meet x y) = x. +Proof. by apply: eq_from_tnth => i; rewrite tnth_join tnth_meet meetKU. Qed. + +Fact leEmeet t1 t2 : (t1 <= t2) = (meet t1 t2 == t1). +Proof. +rewrite leEtprod eqEtuple; apply: eq_forallb => /= i. +by rewrite tnth_meet leEmeet. +Qed. + +Fact meetUl : left_distributive meet join. +Proof. +move=> t1 t2 t3; apply: eq_from_tnth => i. +by rewrite !(tnth_meet, tnth_join) meetUl. +Qed. + +Definition distrLatticeMixin := + DistrLattice.Mixin meetC joinC meetA joinA joinKI meetKU leEmeet meetUl. +Canonical distrLatticeType := DistrLatticeType (n.-tuple T) distrLatticeMixin. + +Lemma meetEtprod t1 t2 : + t1 `&` t2 = [tuple of [seq x.1 `&` x.2 | x <- zip t1 t2]]. +Proof. by []. Qed. + +Lemma joinEtprod t1 t2 : + t1 `|` t2 = [tuple of [seq x.1 `|` x.2 | x <- zip t1 t2]]. +Proof. by []. Qed. + +End DistrLattice. + +Section BDistrLattice. +Variables (n : nat) (T : bDistrLatticeType disp). +Implicit Types (t : n.-tuple T). + +Fact le0x t : [tuple of nseq n 0] <= t :> n.-tuple T. +Proof. by rewrite leEtprod; apply/forallP => i; rewrite tnth_nseq le0x. Qed. + +Canonical bDistrLatticeType := + BDistrLatticeType (n.-tuple T) (BDistrLattice.Mixin le0x). + +Lemma botEtprod : 0 = [tuple of nseq n 0] :> n.-tuple T. Proof. by []. Qed. + +End BDistrLattice. + +Section TBDistrLattice. +Variables (n : nat) (T : tbDistrLatticeType disp). +Implicit Types (t : n.-tuple T). + +Fact lex1 t : t <= [tuple of nseq n 1] :> n.-tuple T. +Proof. by rewrite leEtprod; apply/forallP => i; rewrite tnth_nseq lex1. Qed. + +Canonical tbDistrLatticeType := + TBDistrLatticeType (n.-tuple T) (TBDistrLattice.Mixin lex1). + +Lemma topEtprod : 1 = [tuple of nseq n 1] :> n.-tuple T. Proof. by []. Qed. + +End TBDistrLattice. + +Section CBDistrLattice. +Variables (n : nat) (T : cbDistrLatticeType disp). +Implicit Types (t : n.-tuple T). + +Definition sub t1 t2 : n.-tuple T := + [tuple of [seq x.1 `\` x.2 | x <- zip t1 t2]]. + +Fact tnth_sub t1 t2 i : tnth (sub t1 t2) i = tnth t1 i `\` tnth t2 i. +Proof. +rewrite tnth_map -(tnth_map fst) -(tnth_map snd) -/unzip1 -/unzip2. +by rewrite !(tnth_nth (tnth_default t1 i))/= unzip1_zip ?unzip2_zip ?size_tuple. +Qed. + +Lemma subKI t1 t2 : t2 `&` sub t1 t2 = 0. +Proof. +by apply: eq_from_tnth => i; rewrite tnth_meet tnth_sub subKI tnth_nseq. +Qed. + +Lemma joinIB t1 t2 : t1 `&` t2 `|` sub t1 t2 = t1. +Proof. +by apply: eq_from_tnth => i; rewrite tnth_join tnth_meet tnth_sub joinIB. +Qed. + +Definition cbDistrLatticeMixin := CBDistrLattice.Mixin subKI joinIB. +Canonical cbDistrLatticeType := + CBDistrLatticeType (n.-tuple T) cbDistrLatticeMixin. + +Lemma subEtprod t1 t2 : + t1 `\` t2 = [tuple of [seq x.1 `\` x.2 | x <- zip t1 t2]]. +Proof. by []. Qed. + +End CBDistrLattice. + +Section CTBDistrLattice. +Variables (n : nat) (T : ctbDistrLatticeType disp). +Implicit Types (t : n.-tuple T). + +Definition compl t : n.-tuple T := map_tuple compl t. + +Fact tnth_compl t i : tnth (compl t) i = ~` tnth t i. +Proof. by rewrite tnth_map. Qed. + +Lemma complE t : compl t = sub 1 t. +Proof. +by apply: eq_from_tnth => i; rewrite tnth_compl tnth_sub complE tnth_nseq. +Qed. + +Definition ctbDistrLatticeMixin := CTBDistrLattice.Mixin complE. +Canonical ctbDistrLatticeType := + CTBDistrLatticeType (n.-tuple T) ctbDistrLatticeMixin. + +Lemma complEtprod t : ~` t = [tuple of [seq ~` x | x <- t]]. +Proof. by []. Qed. + +End CTBDistrLattice. + +Canonical finPOrderType n (T : finPOrderType disp) := + [finPOrderType of n.-tuple T]. + +Canonical finDistrLatticeType n (T : finDistrLatticeType disp) := + [finDistrLatticeType of n.-tuple T]. + +Canonical finCDistrLatticeType n (T : finCDistrLatticeType disp) := + [finCDistrLatticeType of n.-tuple T]. + +End TupleProdOrder. + +Module Exports. + +Notation "n .-tupleprod[ disp ]" := (type disp n) + (at level 2, disp at next level, format "n .-tupleprod[ disp ]") : + type_scope. +Notation "n .-tupleprod" := (n.-tupleprod[prod_display]) + (at level 2, format "n .-tupleprod") : type_scope. + +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical distrLatticeType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical cbDistrLatticeType. +Canonical ctbDistrLatticeType. +Canonical finPOrderType. +Canonical finDistrLatticeType. +Canonical finCDistrLatticeType. + +Definition leEtprod := @leEtprod. +Definition ltEtprod := @ltEtprod. +Definition meetEtprod := @meetEtprod. +Definition joinEtprod := @joinEtprod. +Definition botEtprod := @botEtprod. +Definition topEtprod := @topEtprod. +Definition subEtprod := @subEtprod. +Definition complEtprod := @complEtprod. + +Definition tnth_meet := @tnth_meet. +Definition tnth_join := @tnth_join. +Definition tnth_sub := @tnth_sub. +Definition tnth_compl := @tnth_compl. + +End Exports. +End TupleProdOrder. +Import TupleProdOrder.Exports. + +Module DefaultTupleProdOrder. +Section DefaultTupleProdOrder. +Context {disp : unit}. + +Canonical tprod_porderType n (T : porderType disp) := + [porderType of n.-tuple T for [porderType of n.-tupleprod T]]. +Canonical tprod_distrLatticeType n (T : distrLatticeType disp) := + [distrLatticeType of n.-tuple T for [distrLatticeType of n.-tupleprod T]]. +Canonical tprod_bDistrLatticeType n (T : bDistrLatticeType disp) := + [bDistrLatticeType of n.-tuple T for [bDistrLatticeType of n.-tupleprod T]]. +Canonical tprod_tbDistrLatticeType n (T : tbDistrLatticeType disp) := + [tbDistrLatticeType of n.-tuple T for [tbDistrLatticeType of n.-tupleprod T]]. +Canonical tprod_cbDistrLatticeType n (T : cbDistrLatticeType disp) := + [cbDistrLatticeType of n.-tuple T for [cbDistrLatticeType of n.-tupleprod T]]. +Canonical tprod_ctbDistrLatticeType n (T : ctbDistrLatticeType disp) := + [ctbDistrLatticeType of n.-tuple T for + [ctbDistrLatticeType of n.-tupleprod T]]. +Canonical tprod_finPOrderType n (T : finPOrderType disp) := + [finPOrderType of n.-tuple T]. +Canonical tprod_finDistrLatticeType n (T : finDistrLatticeType disp) := + [finDistrLatticeType of n.-tuple T]. +Canonical tprod_finCDistrLatticeType n (T : finCDistrLatticeType disp) := + [finCDistrLatticeType of n.-tuple T]. + +End DefaultTupleProdOrder. +End DefaultTupleProdOrder. + +(*********************************************) +(* We declare a "copy" of the tuples, *) +(* which has canonical lexicographic order. *) +(*********************************************) + +Module TupleLexiOrder. +Section TupleLexiOrder. +Import DefaultSeqLexiOrder. + +Definition type (disp : unit) n T := n.-tuple T. + +Context {disp disp' : unit}. +Local Notation "n .-tuple" := (type disp' n) : type_scope. + +Section Basics. +Variable (n : nat). + +Canonical eqType (T : eqType):= Eval hnf in [eqType of n.-tuple T]. +Canonical choiceType (T : choiceType):= Eval hnf in [choiceType of n.-tuple T]. +Canonical countType (T : countType):= Eval hnf in [countType of n.-tuple T]. +Canonical finType (T : finType):= Eval hnf in [finType of n.-tuple T]. +End Basics. + +Section POrder. +Implicit Types (T : porderType disp). + +Definition porderMixin n T := [porderMixin of n.-tuple T by <:]. +Canonical porderType n T := POrderType disp' (n.-tuple T) (porderMixin n T). + + +Lemma lexi_tupleP n T (t1 t2 : n.-tuple T) : + reflect (exists k : 'I_n.+1, forall i : 'I_n, (i <= k)%N -> + tnth t1 i <= tnth t2 i ?= iff (i != k :> nat)) (t1 <= t2). +Proof. +elim: n => [|n IHn] in t1 t2 *. + by rewrite tuple0 [t2]tuple0/= lexx; constructor; exists ord0 => -[]. +case: (tupleP t1) (tupleP t2) => [x1 {t1}t1] [x2 {t2}t2]. +rewrite [_ <= _]lexi_cons; apply: (iffP idP) => [|[k leif_xt12]]. + case: comparableP => //= [ltx12 _|-> /IHn[k kP]]. + exists ord0 => i; rewrite leqn0 => /eqP/(@ord_inj n.+1 i ord0)->. + by apply/leifP; rewrite !tnth0. + exists (lift ord0 k) => i; case: (unliftP ord0 i) => [j ->|-> _]. + by rewrite !ltnS => /kP; rewrite !tnthS. + by apply/leifP; rewrite !tnth0 eqxx. +have /= := leif_xt12 ord0 isT; rewrite !tnth0 => leif_x12. +rewrite leif_x12/=; move: leif_x12 leif_xt12 => /leifP. +case: (unliftP ord0 k) => {k} [k-> /eqP<-{x2}|-> /lt_geF->//] leif_xt12. +rewrite lexx implyTb; apply/IHn; exists k => i le_ik. +by have := leif_xt12 (lift ord0 i) le_ik; rewrite !tnthS. +Qed. + +Lemma ltxi_tupleP n T (t1 t2 : n.-tuple T) : + reflect (exists k : 'I_n, forall i : 'I_n, (i <= k)%N -> + tnth t1 i <= tnth t2 i ?= iff (i != k :> nat)) (t1 < t2). +Proof. +elim: n => [|n IHn] in t1 t2 *. + by rewrite tuple0 [t2]tuple0/= ltxx; constructor => - [] []. +case: (tupleP t1) (tupleP t2) => [x1 {t1}t1] [x2 {t2}t2]. +rewrite [_ < _]ltxi_cons; apply: (iffP idP) => [|[k leif_xt12]]. + case: (comparableP x1 x2) => //= [ltx12 _|-> /IHn[k kP]]. + exists ord0 => i; rewrite leqn0 => /eqP/(@ord_inj n.+1 i ord0)->. + by apply/leifP; rewrite !tnth0. + exists (lift ord0 k) => i; case: (unliftP ord0 i) => {i} [i ->|-> _]. + by rewrite !ltnS => /kP; rewrite !tnthS. + by apply/leifP; rewrite !tnth0 eqxx. +have /= := leif_xt12 ord0 isT; rewrite !tnth0 => leif_x12. +rewrite leif_x12/=; move: leif_x12 leif_xt12 => /leifP. +case: (unliftP ord0 k) => {k} [k-> /eqP<-{x2}|-> /lt_geF->//] leif_xt12. +rewrite lexx implyTb; apply/IHn; exists k => i le_ik. +by have := leif_xt12 (lift ord0 i) le_ik; rewrite !tnthS. +Qed. + + +Lemma ltxi_tuplePlt n T (t1 t2 : n.-tuple T) : reflect + (exists2 k : 'I_n, forall i : 'I_n, (i < k)%N -> tnth t1 i = tnth t2 i + & tnth t1 k < tnth t2 k) + (t1 < t2). +Proof. +apply: (iffP (ltxi_tupleP _ _)) => [[k kP]|[k kP ltk12]]. + exists k => [i i_lt|]; last by rewrite (lt_leif (kP _ _)) ?eqxx ?leqnn. + by have /eqTleif->// := kP i (ltnW i_lt); rewrite ltn_eqF. +by exists k => i; case: ltngtP => //= [/kP-> _|/ord_inj-> _]; apply/leifP. +Qed. + +End POrder. + +Section Total. +Variables (n : nat) (T : orderType disp). +Implicit Types (t : n.-tuple T). + +Definition total : totalPOrderMixin [porderType of n.-tuple T] := + [totalOrderMixin of n.-tuple T by <:]. +Canonical distrLatticeType := DistrLatticeType (n.-tuple T) total. +Canonical orderType := OrderType (n.-tuple T) total. + +End Total. + +Section BDistrLattice. +Variables (n : nat) (T : finOrderType disp). +Implicit Types (t : n.-tuple T). + +Fact le0x t : [tuple of nseq n 0] <= t :> n.-tuple T. +Proof. by apply: sub_seqprod_lexi; apply: le0x (t : n.-tupleprod T). Qed. + +Canonical bDistrLatticeType := + BDistrLatticeType (n.-tuple T) (BDistrLattice.Mixin le0x). + +Lemma botEtlexi : 0 = [tuple of nseq n 0] :> n.-tuple T. Proof. by []. Qed. + +End BDistrLattice. + +Section TBDistrLattice. +Variables (n : nat) (T : finOrderType disp). +Implicit Types (t : n.-tuple T). + +Fact lex1 t : t <= [tuple of nseq n 1]. +Proof. by apply: sub_seqprod_lexi; apply: lex1 (t : n.-tupleprod T). Qed. + +Canonical tbDistrLatticeType := + TBDistrLatticeType (n.-tuple T) (TBDistrLattice.Mixin lex1). + +Lemma topEtlexi : 1 = [tuple of nseq n 1] :> n.-tuple T. Proof. by []. Qed. + +End TBDistrLattice. + +Canonical finPOrderType n (T : finPOrderType disp) := + [finPOrderType of n.-tuple T]. +Canonical finDistrLatticeType n (T : finOrderType disp) := + [finDistrLatticeType of n.-tuple T]. +Canonical finOrderType n (T : finOrderType disp) := + [finOrderType of n.-tuple T]. + +Lemma sub_tprod_lexi d n (T : POrder.Exports.porderType disp) : + subrel (<=%O : rel (n.-tupleprod[d] T)) (<=%O : rel (n.-tuple T)). +Proof. exact: sub_seqprod_lexi. Qed. + +End TupleLexiOrder. + +Module Exports. + +Notation "n .-tuplelexi[ disp ]" := (type disp n) + (at level 2, disp at next level, format "n .-tuplelexi[ disp ]") : + order_scope. +Notation "n .-tuplelexi" := (n.-tuplelexi[lexi_display]) + (at level 2, format "n .-tuplelexi") : order_scope. + +Canonical eqType. +Canonical choiceType. +Canonical countType. +Canonical finType. +Canonical porderType. +Canonical distrLatticeType. +Canonical orderType. +Canonical bDistrLatticeType. +Canonical tbDistrLatticeType. +Canonical finPOrderType. +Canonical finDistrLatticeType. +Canonical finOrderType. + +Definition lexi_tupleP := @lexi_tupleP. +Arguments lexi_tupleP {disp disp' n T t1 t2}. +Definition ltxi_tupleP := @ltxi_tupleP. +Arguments ltxi_tupleP {disp disp' n T t1 t2}. +Definition ltxi_tuplePlt := @ltxi_tuplePlt. +Arguments ltxi_tuplePlt {disp disp' n T t1 t2}. +Definition topEtlexi := @topEtlexi. +Definition botEtlexi := @botEtlexi. +Definition sub_tprod_lexi := @sub_tprod_lexi. + +End Exports. +End TupleLexiOrder. +Import TupleLexiOrder.Exports. + +Module DefaultTupleLexiOrder. +Section DefaultTupleLexiOrder. +Context {disp : unit}. + +Canonical tlexi_porderType n (T : porderType disp) := + [porderType of n.-tuple T for [porderType of n.-tuplelexi T]]. +Canonical tlexi_distrLatticeType n (T : orderType disp) := + [distrLatticeType of n.-tuple T for [distrLatticeType of n.-tuplelexi T]]. +Canonical tlexi_bDistrLatticeType n (T : finOrderType disp) := + [bDistrLatticeType of n.-tuple T for [bDistrLatticeType of n.-tuplelexi T]]. +Canonical tlexi_tbDistrLatticeType n (T : finOrderType disp) := + [tbDistrLatticeType of n.-tuple T for [tbDistrLatticeType of n.-tuplelexi T]]. +Canonical tlexi_orderType n (T : orderType disp) := + [orderType of n.-tuple T for [orderType of n.-tuplelexi T]]. +Canonical tlexi_finPOrderType n (T : finPOrderType disp) := + [finPOrderType of n.-tuple T]. +Canonical tlexi_finDistrLatticeType n (T : finOrderType disp) := + [finDistrLatticeType of n.-tuple T]. +Canonical tlexi_finOrderType n (T : finOrderType disp) := + [finOrderType of n.-tuple T]. + +End DefaultTupleLexiOrder. +End DefaultTupleLexiOrder. + +Module Syntax. +Export POSyntax. +Export DistrLatticeSyntax. +Export BDistrLatticeSyntax. +Export TBDistrLatticeSyntax. +Export CBDistrLatticeSyntax. +Export CTBDistrLatticeSyntax. +Export TotalSyntax. +Export ConverseSyntax. +Export DvdSyntax. +End Syntax. + +Module LTheory. +Export POCoercions. +Export ConversePOrder. +Export POrderTheory. + +Export ConverseDistrLattice. +Export DistrLatticeTheoryMeet. +Export DistrLatticeTheoryJoin. +Export BDistrLatticeTheory. +Export ConverseTBDistrLattice. +Export TBDistrLatticeTheory. +End LTheory. + +Module CTheory. +Export LTheory CBDistrLatticeTheory CTBDistrLatticeTheory. +End CTheory. + +Module TTheory. +Export LTheory TotalTheory. +End TTheory. + +Module Theory. +Export CTheory TotalTheory. +End Theory. + +End Order. + +Export Order.Syntax. + +Export Order.POrder.Exports. +Export Order.FinPOrder.Exports. +Export Order.DistrLattice.Exports. +Export Order.BDistrLattice.Exports. +Export Order.TBDistrLattice.Exports. +Export Order.FinDistrLattice.Exports. +Export Order.CBDistrLattice.Exports. +Export Order.CTBDistrLattice.Exports. +Export Order.FinCDistrLattice.Exports. +Export Order.Total.Exports. +Export Order.FinTotal.Exports. + +Export Order.TotalPOrderMixin.Exports. +Export Order.LtPOrderMixin.Exports. +Export Order.MeetJoinMixin.Exports. +Export Order.LeOrderMixin.Exports. +Export Order.LtOrderMixin.Exports. +Export Order.CanMixin.Exports. +Export Order.SubOrder.Exports. + +Export Order.NatOrder.Exports. +Export Order.NatDvd.Exports. +Export Order.BoolOrder.Exports. +Export Order.ProdOrder.Exports. +Export Order.SigmaOrder.Exports. +Export Order.ProdLexiOrder.Exports. +Export Order.SeqProdOrder.Exports. +Export Order.SeqLexiOrder.Exports. +Export Order.TupleProdOrder.Exports. +Export Order.TupleLexiOrder.Exports. + +Module DefaultProdOrder := Order.DefaultProdOrder. +Module DefaultSeqProdOrder := Order.DefaultSeqProdOrder. +Module DefaultTupleProdOrder := Order.DefaultTupleProdOrder. +Module DefaultProdLexiOrder := Order.DefaultProdLexiOrder. +Module DefaultSeqLexiOrder := Order.DefaultSeqLexiOrder. +Module DefaultTupleLexiOrder := Order.DefaultTupleLexiOrder. diff --git a/mathcomp/ssreflect/prime.v b/mathcomp/ssreflect/prime.v index 4e55abe..d8f5939 100644 --- a/mathcomp/ssreflect/prime.v +++ b/mathcomp/ssreflect/prime.v @@ -715,6 +715,9 @@ set k := logn p _; have: p ^ k %| m * n by rewrite pfactor_dvdn. by rewrite Gauss_dvdr ?coprime_expl // -pfactor_dvdn. Qed. +Lemma logn_coprime p m : coprime p m -> logn p m = 0. +Proof. by move=> coprime_pm; rewrite -[m]muln1 logn_Gauss// logn1. Qed. + Lemma lognM p m n : 0 < m -> 0 < n -> logn p (m * n) = logn p m + logn p n. Proof. case p_pr: (prime p); last by rewrite /logn p_pr. @@ -899,6 +902,13 @@ apply: eq_bigr => p _; rewrite ltnS lognE. by case: and3P => [[_ n_gt0 p_dv_n]|]; rewrite ?if_same // andbC dvdn_leq. Qed. +Lemma eq_partn_from_log m n (pi : nat_pred) : 0 < m -> 0 < n -> + {in pi, logn^~ m =1 logn^~ n} -> m`_pi = n`_pi. +Proof. +move=> m0 n0 eq_log; rewrite !(@widen_partn (maxn m n)) ?leq_maxl ?leq_maxr//. +by apply: eq_bigr => p /eq_log ->. +Qed. + Lemma partn0 pi : 0`_pi = 1. Proof. by apply: big1_seq => [] [|n]; rewrite andbC. Qed. @@ -983,6 +993,11 @@ move=> n_gt0; rewrite -{2}(partn_pi n_gt0) {2}/partn big_mkcond /=. by apply: eq_bigr => p _; rewrite -logn_gt0; case: (logn p _). Qed. +Lemma eqn_from_log m n : 0 < m -> 0 < n -> logn^~ m =1 logn^~ n -> m = n. +Proof. +by move=> ? ? /(@in1W _ predT)/eq_partn_from_log; rewrite !partnT// => ->. +Qed. + Lemma partnC pi n : n > 0 -> n`_pi * n`_pi^' = n. Proof. move=> n_gt0; rewrite -{3}(partnT n_gt0) /partn. @@ -1048,6 +1063,20 @@ apply/dvdn_biggcdP=> i Pi; rewrite -(partnC pi (F_gt0 i Pi)) dvdn_mul //. by rewrite partn_dvd ?F_gt0 // (@biggcdn_inf _ i). Qed. +Lemma logn_gcd p m n : 0 < m -> 0 < n -> + logn p (gcdn m n) = minn (logn p m) (logn p n). +Proof. +move=> m_gt0 n_gt0; case p_pr: (prime p); last by rewrite /logn p_pr. +by apply: (@expnI p); rewrite ?prime_gt1// expn_min -!p_part partn_gcd. +Qed. + +Lemma logn_lcm p m n : 0 < m -> 0 < n -> + logn p (lcmn m n) = maxn (logn p m) (logn p n). +Proof. +move=> m_gt0 n_gt0; rewrite /lcmn logn_div ?dvdn_mull ?dvdn_gcdr//. +by rewrite lognM// logn_gcd// -addn_min_max addnC addnK. +Qed. + Lemma sub_in_pnat pi rho n : {in \pi(n), {subset pi <= rho}} -> pi.-nat n -> rho.-nat n. Proof. diff --git a/mathcomp/ssreflect/ssrnat.v b/mathcomp/ssreflect/ssrnat.v index b518c96..bccb968 100644 --- a/mathcomp/ssreflect/ssrnat.v +++ b/mathcomp/ssreflect/ssrnat.v @@ -1450,6 +1450,15 @@ Proof. by case=> le_ab; rewrite eqn_leq le_ab. Qed. Lemma ltn_leqif a b C : a <= b ?= iff C -> (a < b) = ~~ C. Proof. by move=> le_ab; rewrite ltnNge (geq_leqif le_ab). Qed. +Lemma ltnNleqif x y C : x <= y ?= iff ~~ C -> (x < y) = C. +Proof. by move=> /ltn_leqif; rewrite negbK. Qed. + +Lemma eq_leqif x y C : x <= y ?= iff C -> (x == y) = C. +Proof. by move=> /leqifP; case: C ltngtP => [] []. Qed. + +Lemma eqTleqif x y C : x <= y ?= iff C -> C -> x = y. +Proof. by move=> /eq_leqif<-/eqP. Qed. + Lemma leqif_add m1 n1 C1 m2 n2 C2 : m1 <= n1 ?= iff C1 -> m2 <= n2 ?= iff C2 -> m1 + m2 <= n1 + n2 ?= iff C1 && C2. @@ -1538,21 +1547,21 @@ Let leq_total := leq_total. Lemma ltnW_homo : {homo f : m n / m < n} -> {homo f : m n / m <= n}. Proof. exact: homoW. Qed. -Lemma homo_inj_lt : injective f -> {homo f : m n / m <= n} -> +Lemma inj_homo_ltn : injective f -> {homo f : m n / m <= n} -> {homo f : m n / m < n}. Proof. exact: inj_homo. Qed. Lemma ltnW_nhomo : {homo f : m n /~ m < n} -> {homo f : m n /~ m <= n}. Proof. exact: homoW. Qed. -Lemma nhomo_inj_lt : injective f -> {homo f : m n /~ m <= n} -> +Lemma inj_nhomo_ltn : injective f -> {homo f : m n /~ m <= n} -> {homo f : m n /~ m < n}. Proof. exact: inj_homo. Qed. -Lemma incrn_inj : {mono f : m n / m <= n} -> injective f. +Lemma incn_inj : {mono f : m n / m <= n} -> injective f. Proof. exact: mono_inj. Qed. -Lemma decrn_inj : {mono f : m n /~ m <= n} -> injective f. +Lemma decn_inj : {mono f : m n /~ m <= n} -> injective f. Proof. exact: mono_inj. Qed. Lemma leqW_mono : {mono f : m n / m <= n} -> {mono f : m n / m < n}. @@ -1577,21 +1586,21 @@ Lemma ltnW_nhomo_in : {in D & D', {homo f : m n /~ m < n}} -> {in D & D', {homo f : m n /~ m <= n}}. Proof. exact: homoW_in. Qed. -Lemma homo_inj_lt_in : {in D & D', injective f} -> +Lemma inj_homo_ltn_in : {in D & D', injective f} -> {in D & D', {homo f : m n / m <= n}} -> {in D & D', {homo f : m n / m < n}}. Proof. exact: inj_homo_in. Qed. -Lemma nhomo_inj_lt_in : {in D & D', injective f} -> +Lemma inj_nhomo_ltn_in : {in D & D', injective f} -> {in D & D', {homo f : m n /~ m <= n}} -> {in D & D', {homo f : m n /~ m < n}}. Proof. exact: inj_homo_in. Qed. -Lemma incrn_inj_in : {in D &, {mono f : m n / m <= n}} -> +Lemma incn_inj_in : {in D &, {mono f : m n / m <= n}} -> {in D &, injective f}. Proof. exact: mono_inj_in. Qed. -Lemma decrn_inj_in : {in D &, {mono f : m n /~ m <= n}} -> +Lemma decn_inj_in : {in D &, {mono f : m n /~ m <= n}} -> {in D &, injective f}. Proof. exact: mono_inj_in. Qed. diff --git a/mathcomp/ssreflect/tuple.v b/mathcomp/ssreflect/tuple.v index dd73664..10d54f0 100644 --- a/mathcomp/ssreflect/tuple.v +++ b/mathcomp/ssreflect/tuple.v @@ -213,6 +213,9 @@ Definition thead (u : n.+1.-tuple T) := tnth u ord0. Lemma tnth0 x t : tnth [tuple of x :: t] ord0 = x. Proof. by []. Qed. +Lemma tnthS x t i : tnth [tuple of x :: t] (lift ord0 i) = tnth t i. +Proof. by rewrite (tnth_nth (tnth_default t i)). Qed. + Lemma theadE x t : thead [tuple of x :: t] = x. Proof. by []. Qed. @@ -231,6 +234,11 @@ Qed. Lemma tnth_map f t i : tnth [tuple of map f t] i = f (tnth t i) :> rT. Proof. by apply: nth_map; rewrite size_tuple. Qed. +Lemma tnth_nseq x i : tnth [tuple of nseq n x] i = x. +Proof. +by rewrite !(tnth_nth (tnth_default (nseq_tuple x) i)) nth_nseq ltn_ord. +Qed. + End SeqTuple. Lemma tnth_behead n T (t : n.+1.-tuple T) i : @@ -277,6 +285,10 @@ Canonical tuple_eqType := Eval hnf in EqType (n.-tuple T) tuple_eqMixin. Canonical tuple_predType := PredType (pred_of_seq : n.-tuple T -> pred T). +Lemma eqEtuple (t1 t2 : n.-tuple T) : + (t1 == t2) = [forall i, tnth t1 i == tnth t2 i]. +Proof. by apply/eqP/'forall_eqP => [->|/eq_from_tnth]. Qed. + Lemma memtE (t : n.-tuple T) : mem t = mem (tval t). Proof. by []. Qed. |
