diff options
| author | Enrico Tassi | 2015-03-09 11:07:53 +0100 |
|---|---|---|
| committer | Enrico Tassi | 2015-03-09 11:24:38 +0100 |
| commit | fc84c27eac260dffd8f2fb1cb56d599f1e3486d9 (patch) | |
| tree | c16205f1637c80833a4c4598993c29fa0fd8c373 /mathcomp/real_closed | |
Initial commit
Diffstat (limited to 'mathcomp/real_closed')
| -rw-r--r-- | mathcomp/real_closed/all.v | 9 | ||||
| -rw-r--r-- | mathcomp/real_closed/bigenough.v | 118 | ||||
| -rw-r--r-- | mathcomp/real_closed/cauchyreals.v | 1681 | ||||
| -rw-r--r-- | mathcomp/real_closed/complex.v | 1252 | ||||
| -rw-r--r-- | mathcomp/real_closed/ordered_qelim.v | 1180 | ||||
| -rw-r--r-- | mathcomp/real_closed/polyorder.v | 273 | ||||
| -rw-r--r-- | mathcomp/real_closed/polyrcf.v | 1857 | ||||
| -rw-r--r-- | mathcomp/real_closed/qe_rcf.v | 1008 | ||||
| -rw-r--r-- | mathcomp/real_closed/qe_rcf_th.v | 1293 | ||||
| -rw-r--r-- | mathcomp/real_closed/realalg.v | 1530 |
10 files changed, 10201 insertions, 0 deletions
diff --git a/mathcomp/real_closed/all.v b/mathcomp/real_closed/all.v new file mode 100644 index 0000000..c9de2b6 --- /dev/null +++ b/mathcomp/real_closed/all.v @@ -0,0 +1,9 @@ +Require Export bigenough. +Require Export cauchyreals. +Require Export complex. +Require Export ordered_qelim. +Require Export polyorder. +Require Export polyrcf. +Require Export qe_rcf_th. +Require Export qe_rcf. +Require Export realalg. diff --git a/mathcomp/real_closed/bigenough.v b/mathcomp/real_closed/bigenough.v new file mode 100644 index 0000000..d9c89ca --- /dev/null +++ b/mathcomp/real_closed/bigenough.v @@ -0,0 +1,118 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. + +(****************************************************************************) +(* This is a small library to do epsilon - N reasonning. *) +(* In order to use it, one only has to know the following tactics: *) +(* *) +(* pose_big_enough i == pose a big enough natural number i *) +(* pose_big_modulus m F == pose a function m : F -> nat which should *) +(* provide a big enough return value *) +(* exists_big_modulus m F := pose_big_modulus m F; exists m *) +(* big_enough == replaces a big enough constraint x <= i *) +(* by true and implicity remembers that i should *) +(* be bigger than x. *) +(* close == all "pose" tactics create a dummy subgoal to *) +(* force the user to explictely indicate that all *) +(* constraints have been found *) +(****************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Module BigEnough. + +Record big_rel_class_of T (leq : rel T) := + BigRelClass { + leq_big_internal_op : rel T; + bigger_than_op : seq T -> T; + _ : leq_big_internal_op = leq; + _ : forall i s, leq_big_internal_op i (bigger_than_op (i :: s)); + _ : forall i j s, leq_big_internal_op i (bigger_than_op s) -> + leq_big_internal_op i (bigger_than_op (j :: s)) +}. + +Record big_rel_of T := BigRelOf { + leq_big :> rel T; + big_rel_class : big_rel_class_of leq_big +}. + +Definition bigger_than_of T (b : big_rel_of T) + (phb : phantom (rel T) b) := + bigger_than_op (big_rel_class b). +Notation bigger_than leq := (@bigger_than_of _ _ (Phantom (rel _) leq)). + +Definition leq_big_internal_of T (b : big_rel_of T) + (phb : phantom (rel T) b) := + leq_big_internal_op (big_rel_class b). +Notation leq_big_internal leq := (@leq_big_internal_of _ _ (Phantom (rel _) leq)). + +Lemma next_bigger_than T (b : big_rel_of T) i j s : + leq_big_internal b i (bigger_than b s) -> + leq_big_internal b i (bigger_than b (j :: s)). +Proof. by case: b i j s => [? []]. Qed. + +Lemma instantiate_bigger_than T (b : big_rel_of T) i s : + leq_big_internal b i (bigger_than b (i :: s)). +Proof. by case: b i s => [? []]. Qed. + +Lemma leq_big_internalE T (b : big_rel_of T) : leq_big_internal b = leq_big b. +Proof. by case: b => [? []]. Qed. + +(* Lemma big_enough T (b : big_rel_of T) i s : *) +(* leq_big_internal b i (bigger_than b s) -> *) +(* leq_big b i (bigger_than b s). *) +(* Proof. by rewrite leq_big_internalE. Qed. *) + +Lemma context_big_enough P T (b : big_rel_of T) i s : + leq_big_internal b i (bigger_than b s) -> + P true -> + P (leq_big b i (bigger_than b s)). +Proof. by rewrite leq_big_internalE => ->. Qed. + +Definition big_rel_leq_class : big_rel_class_of leq. +Proof. +exists leq (foldr maxn 0%N) => [|i s|i j s /leq_trans->] //; +by rewrite (leq_maxl, leq_maxr). +Qed. +Canonical big_enough_nat := BigRelOf big_rel_leq_class. + +Definition closed T (i : T) := {j : T | j = i}. +Ltac close := match goal with + | |- context [closed ?i] => + instantiate (1 := [::]) in (Value of i); exists i + end. + +Ltac pose_big_enough i := + evar (i : nat); suff : closed i; first do + [move=> _; instantiate (1 := bigger_than leq _) in (Value of i)]. + +Ltac pose_big_modulus m F := + evar (m : F -> nat); suff : closed m; first do + [move=> _; instantiate (1 := (fun e => bigger_than leq _)) in (Value of m)]. + +Ltac exists_big_modulus m F := pose_big_modulus m F; first exists m. + +Ltac olddone := + trivial; hnf; intros; solve + [ do ![solve [trivial | apply: sym_equal; trivial] + | discriminate | contradiction | split] + | case not_locked_false_eq_true; assumption + | match goal with H : ~ _ |- _ => solve [case H; trivial] end]. + +Ltac big_enough := + do ?[ apply context_big_enough; + first do [do ?[ now apply instantiate_bigger_than + | apply next_bigger_than]]]. + +Ltac big_enough_trans := + match goal with + | [leq_nm : is_true (?n <= ?m)%N |- is_true (?x <= ?m)] => + apply: leq_trans leq_nm; big_enough; olddone + | _ => big_enough; olddone + end. + +Ltac done := do [olddone|big_enough_trans]. + +End BigEnough. diff --git a/mathcomp/real_closed/cauchyreals.v b/mathcomp/real_closed/cauchyreals.v new file mode 100644 index 0000000..83504be --- /dev/null +++ b/mathcomp/real_closed/cauchyreals.v @@ -0,0 +1,1681 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import bigop ssralg ssrnum ssrint rat poly polydiv polyorder. +Require Import perm matrix mxpoly polyXY binomial bigenough. + +(***************************************************************************) +(* This is a standalone construction of Cauchy reals over an arbitrary *) +(* discrete archimedian field R. *) +(* creals R == setoid of Cauchy sequences, it is not discrete and *) +(* cannot be equipped with any ssreflect algebraic structure *) +(***************************************************************************) + +Import GRing.Theory Num.Theory Num.Def BigEnough. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Delimit Scope creal_scope with CR. + +Section poly_extra. + +Local Open Scope ring_scope. + +Lemma monic_monic_from_neq0 (F : fieldType) (p : {poly F}) : + (p != 0)%B -> (lead_coef p) ^-1 *: p \is monic. +Proof. by move=> ?; rewrite monicE lead_coefZ mulVf ?lead_coef_eq0. Qed. + +(* GG -- lemmas with ssrnum dependencies cannot go in poly! *) +Lemma size_derivn (R : realDomainType) (p : {poly R}) n : + size p^`(n) = (size p - n)%N. +Proof. +elim: n=> [|n ihn]; first by rewrite derivn0 subn0. +by rewrite derivnS size_deriv ihn -subnS. +Qed. + +Lemma size_nderivn (R : realDomainType) (p : {poly R}) n : + size p^`N(n) = (size p - n)%N. +Proof. +rewrite -size_derivn nderivn_def -mulr_natl. +by rewrite -polyC1 -!polyC_muln size_Cmul // pnatr_eq0 -lt0n fact_gt0. +Qed. + +End poly_extra. + +Local Notation eval := horner_eval. + +Section ordered_extra. + +Definition gtr0E := (invr_gt0, exprn_gt0, ltr0n, @ltr01). +Definition ger0E := (invr_ge0, exprn_ge0, ler0n, @ler01). + +End ordered_extra. + +Section polyorder_extra. + +Variable F : realDomainType. + +Local Open Scope ring_scope. + +Definition poly_bound (p : {poly F}) (a r : F) : F + := 1 + \sum_(i < size p) `|p`_i| * (`|a| + `|r|) ^+ i. + +Lemma poly_boundP p a r x : `|x - a| <= r -> + `|p.[x]| <= poly_bound p a r. +Proof. +have [r_ge0|r_lt0] := lerP 0 r; last first. + by move=> hr; have := ler_lt_trans hr r_lt0; rewrite normr_lt0. +rewrite ler_distl=> /andP[lx ux]. +rewrite ler_paddl //. +elim/poly_ind: p=> [|p c ihp]. + by rewrite horner0 normr0 size_poly0 big_ord0. +rewrite hornerMXaddC size_MXaddC. +have [->|p_neq0 /=] := altP eqP. + rewrite horner0 !mul0r !add0r size_poly0. + have [->|c_neq0] /= := altP eqP; first by rewrite normr0 big_ord0. + rewrite big_ord_recl big_ord0 addr0 coefC /=. + by rewrite ler_pmulr ?normr_gt0 // ler_addl ler_maxr !normr_ge0. +rewrite big_ord_recl coefD coefMX coefC eqxx add0r. +rewrite (ler_trans (ler_norm_add _ _)) // addrC ler_add //. + by rewrite expr0 mulr1. +rewrite normrM. +move: ihp=> /(ler_wpmul2r (normr_ge0 x)) /ler_trans-> //. +rewrite mulr_suml ler_sum // => i _. +rewrite coefD coefC coefMX /= addr0 exprSr mulrA. +rewrite ler_wpmul2l //. + by rewrite ?mulr_ge0 ?exprn_ge0 ?ler_maxr ?addr_ge0 ?normr_ge0 // ltrW. +rewrite (ger0_norm r_ge0) ler_norml opprD. +rewrite (ler_trans _ lx) ?(ler_trans ux) // ler_add2r. + by rewrite ler_normr lerr. +by rewrite ler_oppl ler_normr lerr orbT. +Qed. + +Lemma poly_bound_gt0 p a r : 0 < poly_bound p a r. +Proof. +rewrite ltr_paddr // sumr_ge0 // => i _. +by rewrite mulr_ge0 ?exprn_ge0 ?addr_ge0 ?ler_maxr ?normr_ge0 // ltrW. +Qed. + +Lemma poly_bound_ge0 p a r : 0 <= poly_bound p a r. +Proof. by rewrite ltrW // poly_bound_gt0. Qed. + +Definition poly_accr_bound (p : {poly F}) (a r : F) : F + := (maxr 1 (2%:R * r)) ^+ (size p).-1 + * (1 + \sum_(i < (size p).-1) poly_bound p^`N(i.+1) a r). + +Lemma poly_accr_bound1P p a r x y : + `|x - a| <= r -> `|y - a| <= r -> + `|p.[y] - p.[x]| <= `|y - x| * poly_accr_bound p a r. +Proof. +have [|r_lt0] := lerP 0 r; last first. + by move=> hr; have := ler_lt_trans hr r_lt0; rewrite normr_lt0. +rewrite le0r=> /orP[/eqP->|r_gt0 hx hy]. + by rewrite !normr_le0 !subr_eq0=> /eqP-> /eqP->; rewrite !subrr normr0 mul0r. +rewrite mulrA mulrDr mulr1 ler_paddl ?mulr_ge0 ?normr_ge0 //=. + by rewrite exprn_ge0 ?ler_maxr ?mulr_ge0 ?ger0E ?ltrW. +rewrite -{1}(addNKr x y) [- _ + _]addrC /= -mulrA. +rewrite nderiv_taylor; last exact: mulrC. +have [->|p_neq0] := eqVneq p 0. + rewrite size_poly0 big_ord0 horner0 subr0 normr0 mulr_ge0 ?normr_ge0 //. + by rewrite big_ord0 mulr0 lerr. +rewrite -[size _]prednK ?lt0n ?size_poly_eq0 //. +rewrite big_ord_recl expr0 mulr1 nderivn0 addrC addKr !mulr_sumr. +have := ler_trans (ler_norm_sum _ _ _); apply. +rewrite ler_sum // => i _. +rewrite exprSr mulrA !normrM mulrC ler_wpmul2l ?normr_ge0 //. +suff /ler_wpmul2l /ler_trans : + `|(y - x) ^+ i| <= maxr 1 (2%:R * r) ^+ (size p).-1. + apply; rewrite ?normr_ge0 // mulrC ler_wpmul2l ?poly_boundP //. + by rewrite ?exprn_ge0 // ler_maxr ler01 mulr_ge0 ?ler0n ?ltrW. +case: maxrP=> hr. + rewrite expr1n normrX exprn_ile1 ?normr_ge0 //. + rewrite (ler_trans (ler_dist_add a _ _)) // addrC distrC. + by rewrite (ler_trans _ hr) // mulrDl ler_add ?mul1r. +rewrite (@ler_trans _ ((2%:R * r) ^+ i)) //. + rewrite normrX @ler_expn2r -?topredE /= ?normr_ge0 ?mulr_ge0 ?ler0n //. + by rewrite ltrW. + rewrite (ler_trans (ler_dist_add a _ _)) // addrC distrC. + by rewrite mulrDl ler_add ?mul1r. +by rewrite ler_eexpn2l // ltnW. +Qed. + +Lemma poly_accr_bound_gt0 p a r : 0 < poly_accr_bound p a r. +Proof. +rewrite /poly_accr_bound pmulr_rgt0 //. + rewrite ltr_paddr ?ltr01 //. + by rewrite sumr_ge0 // => i; rewrite poly_bound_ge0. +by rewrite exprn_gt0 // ltr_maxr ltr01 pmulr_rgt0 ?ltr0n. +Qed. + +Lemma poly_accr_bound_ge0 p a r : 0 <= poly_accr_bound p a r. +Proof. by rewrite ltrW // poly_accr_bound_gt0. Qed. + +(* Todo : move to polyorder => need char 0 *) +Lemma gdcop_eq0 (p q : {poly F}) : + (gdcop p q == 0)%B = (q == 0)%B && (p != 0)%B. +Proof. +have [[->|q_neq0] [->|p_neq0] /=] := (altP (q =P 0), altP (p =P 0)). ++ by rewrite gdcop0 eqxx oner_eq0. ++ by rewrite gdcop0 (negPf p_neq0) eqxx. ++ apply/negP=> /eqP hg; have := coprimep_gdco 0 q_neq0. + by rewrite hg coprimep0 eqp01. +by apply/negP=> /eqP hg; have := dvdp_gdco p q; rewrite hg dvd0p; apply/negP. +Qed. + +End polyorder_extra. + +Section polyXY_order_extra. + +Variable F : realFieldType. +Local Open Scope ring_scope. + +Local Notation "p ^ f" := (map_poly f p) : ring_scope. +Local Notation "'Y" := 'X%:P. + +Definition norm_poly2 (p : {poly {poly F}}) := p ^ (map_poly (fun x => `|x|)). + +Lemma coef_norm_poly2 p i j : (norm_poly2 p)`_i`_j = `|p`_i`_j|. +Proof. +rewrite !coef_map_id0 ?normr0 //; last first. +by rewrite /map_poly poly_def size_poly0 big_ord0. +Qed. + +Lemma size_norm_poly2 p : size (norm_poly2 p) = size p. +Proof. +rewrite /norm_poly2; have [->|p0] := eqVneq p 0. + by rewrite /map_poly poly_def !(size_poly0, big_ord0). +rewrite /map_poly size_poly_eq // -size_poly_eq0 size_poly_eq //. + by rewrite -lead_coefE size_poly_eq0 lead_coef_eq0. +by rewrite -!lead_coefE normr_eq0 !lead_coef_eq0. +Qed. + +End polyXY_order_extra. + +Section polyorder_field_extra. + +Variable F : realFieldType. + +Local Open Scope ring_scope. + +Definition poly_accr_bound2 (p : {poly F}) (a r : F) : F + := (maxr 1 (2%:R * r)) ^+ (size p).-2 + * (1 + \sum_(i < (size p).-2) poly_bound p^`N(i.+2) a r). + +Lemma poly_accr_bound2_gt0 p a r : 0 < poly_accr_bound2 p a r. +Proof. +rewrite /poly_accr_bound pmulr_rgt0 //. + rewrite ltr_paddr ?ltr01 //. + by rewrite sumr_ge0 // => i; rewrite poly_bound_ge0. +by rewrite exprn_gt0 // ltr_maxr ltr01 pmulr_rgt0 ?ltr0n. +Qed. + +Lemma poly_accr_bound2_ge0 p a r : 0 <= poly_accr_bound2 p a r. +Proof. by rewrite ltrW // poly_accr_bound2_gt0. Qed. + +Lemma poly_accr_bound2P p (a r x y : F) : (x != y)%B -> + `|x - a| <= r -> `|y - a| <= r -> + `|(p.[y] - p.[x]) / (y - x) - p^`().[x]| + <= `|y - x| * poly_accr_bound2 p a r. +Proof. +have [|r_lt0] := lerP 0 r; last first. + by move=> _ hr; have := ler_lt_trans hr r_lt0; rewrite normr_lt0. +rewrite le0r=> /orP[/eqP->|r_gt0]. + rewrite !normr_le0 !subr_eq0. + by move=> nxy /eqP xa /eqP xb; rewrite xa xb eqxx in nxy. +move=> neq_xy hx hy. +rewrite mulrA mulrDr mulr1 ler_paddl ?mulr_ge0 ?normr_ge0 //=. + by rewrite exprn_ge0 ?ler_maxr ?mulr_ge0 ?ger0E ?ltrW. +rewrite -{1}(addNKr x y) [- _ + _]addrC /= -mulrA. +rewrite nderiv_taylor; last exact: mulrC. +have [->|p_neq0] := eqVneq p 0. + by rewrite derivC !horner0 size_poly0 !(big_ord0, subrr, mul0r) normr0 !mulr0. +rewrite -[size _]prednK ?lt0n ?size_poly_eq0 //. +rewrite big_ord_recl expr0 mulr1 nderivn0 /= -size_deriv. +have [->|p'_neq0] := eqVneq p^`() 0. + by rewrite horner0 size_poly0 !big_ord0 addr0 !(subrr, mul0r) normr0 !mulr0. +rewrite -[size _]prednK ?lt0n ?size_poly_eq0 // big_ord_recl expr1. +rewrite addrAC subrr add0r mulrDl mulfK; last by rewrite subr_eq0 eq_sym. +rewrite nderivn1 addrAC subrr add0r mulr_sumr normrM normfV. +rewrite ler_pdivr_mulr ?normr_gt0; last by rewrite subr_eq0 eq_sym. +rewrite mulrAC -expr2 mulrC mulr_suml. +have := ler_trans (ler_norm_sum _ _ _); apply. +rewrite ler_sum // => i _ /=; rewrite /bump /= !add1n. +rewrite normrM normrX 3!exprSr expr1 !mulrA !ler_wpmul2r ?normr_ge0 //. +suff /ler_wpmul2l /ler_trans : + `|(y - x)| ^+ i <= maxr 1 (2%:R * r) ^+ (size p^`()).-1. + apply; rewrite ?normr_ge0 // mulrC ler_wpmul2l ?poly_boundP //. + by rewrite ?exprn_ge0 // ler_maxr ler01 mulr_ge0 ?ler0n ?ltrW. +case: maxrP=> hr. + rewrite expr1n exprn_ile1 ?normr_ge0 //. + rewrite (ler_trans (ler_dist_add a _ _)) // addrC distrC. + by rewrite (ler_trans _ hr) // mulrDl ler_add ?mul1r. +rewrite (@ler_trans _ ((2%:R * r) ^+ i)) //. + rewrite @ler_expn2r -?topredE /= ?normr_ge0 ?mulr_ge0 ?ler0n //. + by rewrite ltrW. + rewrite (ler_trans (ler_dist_add a _ _)) // addrC distrC. + by rewrite mulrDl ler_add ?mul1r. +by rewrite ler_eexpn2l // ltnW. +Qed. + +End polyorder_field_extra. + +Section monotony. + +Variable F : realFieldType. + +Local Open Scope ring_scope. + +Definition accr_pos p (a r : F) := + ({ k | 0 < k & forall x y, (x != y)%B -> + `|x - a| <= r -> `|y - a| <= r -> (p.[x] - p.[y]) / (x - y) > k } + * forall x, `|x - a| <= r -> p^`().[x] > 0)%type. + +Definition accr_neg p (a r : F) := + ({ k | 0 < k & forall x y, (x != y)%B -> + `|x - a| <= r -> `|y - a| <= r -> (p.[x] - p.[y]) / (x - y) < - k} + * forall x, `|x - a| <= r -> p^`().[x] < 0)%type. + +Definition strong_mono p (a r : F) := (accr_pos p a r + accr_neg p a r)%type. + +Lemma accr_pos_incr p a r : accr_pos p a r -> + forall x y, `|x - a| <= r -> `|y - a| <= r -> (p.[x] <= p.[y]) = (x <= y). +Proof. +move=> [[k k_gt0 hk] _] x y hx hy. +have [->|neq_xy] := eqVneq x y; first by rewrite !lerr. +have hkxy := hk _ _ neq_xy hx hy. +have := ltr_trans k_gt0 hkxy. +have [lpxpy|lpypx|->] := ltrgtP p.[x] p.[y]. ++ by rewrite nmulr_rgt0 ?subr_lt0 // ?invr_lt0 subr_lt0=> /ltrW->. ++ by rewrite pmulr_rgt0 ?subr_gt0 // ?invr_gt0 subr_gt0 lerNgt=> ->. +by rewrite subrr mul0r ltrr. +Qed. + +Lemma accr_neg_decr p a r : accr_neg p a r -> + forall x y, `|x - a| <= r -> `|y - a| <= r -> (p.[x] <= p.[y]) = (y <= x). +Proof. +move=> [] [k]; rewrite -oppr_lt0=> Nk_lt0 hk _ x y hx hy. +have [->|neq_xy] := eqVneq x y; first by rewrite !lerr. +have hkxy := hk _ _ neq_xy hx hy. +have := ltr_trans hkxy Nk_lt0. +have [lpxpy|lpypx|->] := ltrgtP p.[x] p.[y]. ++ by rewrite nmulr_rlt0 ?subr_lt0 // ?invr_gt0 subr_gt0=> /ltrW->. ++ by rewrite pmulr_rlt0 ?subr_gt0 // ?invr_lt0 subr_lt0 lerNgt=> ->. +by rewrite subrr mul0r ltrr. +Qed. + +Lemma accr_negN p a r : accr_pos p a r -> accr_neg (- p) a r. +Proof. +case=> [[k k_gt0 hk] h]. +split; [ exists k=> // x y nxy hx hy; + by rewrite !hornerN -opprD mulNr ltr_opp2; apply: hk + | by move=> x hx; rewrite derivN hornerN oppr_lt0; apply: h ]. +Qed. + +Lemma accr_posN p a r : accr_neg p a r -> accr_pos (- p) a r. +Proof. +case=> [[k k_gt0 hk] h]. +split; [ exists k=> // x y nxy hx hy; + by rewrite !hornerN -opprD mulNr ltr_oppr; apply: hk + | by move=> x hx; rewrite derivN hornerN oppr_gt0; apply: h ]. +Qed. + +Lemma strong_monoN p a r : strong_mono p a r -> strong_mono (- p) a r. +Proof. by case=> [] hp; [right; apply: accr_negN|left; apply: accr_posN]. Qed. + +Lemma strong_mono_bound p a r : strong_mono p a r + -> {k | 0 < k & forall x y, `|x - a| <= r -> `|y - a| <= r -> + `| x - y | <= k * `| p.[x] - p.[y] | }. +Proof. +case=> [] [[k k_gt0 hk] _]; exists k^-1; rewrite ?invr_gt0=> // x y hx hy; +have [->|neq_xy] := eqVneq x y; do ?[by rewrite !subrr normr0 mulr0]; +move: (hk _ _ neq_xy hx hy); rewrite 1?ltr_oppr ler_pdivl_mull //; +rewrite -ler_pdivl_mulr ?normr_gt0 ?subr_eq0 // => /ltrW /ler_trans-> //; +by rewrite -normfV -normrM ler_normr lerr ?orbT. +Qed. + +Definition merge_intervals (ar1 ar2 : F * F) := + let l := minr (ar1.1 - ar1.2) (ar2.1 - ar2.2) in + let u := maxr (ar1.1 + ar1.2) (ar2.1 + ar2.2) in + ((l + u) / 2%:R, (u - l) / 2%:R). +Local Notation center ar1 ar2 := ((merge_intervals ar1 ar2).1). +Local Notation radius ar1 ar2 := ((merge_intervals ar1 ar2).2). + +Lemma split_interval (a1 a2 r1 r2 x : F) : + 0 < r1 -> 0 < r2 -> `|a1 - a2| <= r1 + r2 -> + (`|x - center (a1, r1) (a2, r2)| <= radius (a1, r1) (a2, r2)) + = (`|x - a1| <= r1) || (`|x - a2| <= r2). +Proof. +move=> r1_gt0 r2_gt0 le_ar. +rewrite /merge_intervals /=. +set l := minr _ _; set u := maxr _ _. +rewrite ler_pdivl_mulr ?gtr0E // -{2}[2%:R]ger0_norm ?ger0E //. +rewrite -normrM mulrBl mulfVK ?pnatr_eq0 // ler_distl. +rewrite opprB addrCA addrK (addrC (l + u)) addrA addrNK. +rewrite -!mulr2n !mulr_natr !ler_muln2r !orFb. +rewrite ler_minl ler_maxr !ler_distl. +have [] := lerP=> /= a1N; have [] := lerP=> //= a1P; +have [] := lerP=> //= a2P; rewrite ?(andbF, andbT) //; symmetry. + rewrite ltrW // (ler_lt_trans _ a1P) //. + rewrite (monoLR (addrK _) (ler_add2r _)) -addrA. + rewrite (monoRL (addNKr _) (ler_add2l _)) addrC. + by rewrite (ler_trans _ le_ar) // ler_normr opprB lerr orbT. +rewrite ltrW // (ltr_le_trans a1N) //. +rewrite (monoLR (addrK _) (ler_add2r _)) -addrA. +rewrite (monoRL (addNKr _) (ler_add2l _)) addrC ?[r2 + _]addrC. +by rewrite (ler_trans _ le_ar) // ler_normr lerr. +Qed. + +Lemma merge_mono p a1 a2 r1 r2 : + 0 < r1 -> 0 < r2 -> + `|a1 - a2| <= (r1 + r2) -> + strong_mono p a1 r1 -> strong_mono p a2 r2 -> + strong_mono p (center (a1, r1) (a2, r2)) (radius (a1, r1) (a2, r2)). +Proof. +move=> r1_gt0 r2_gt0 har sm1; wlog : p sm1 / accr_pos p a1 r1. + move=> hwlog; case: (sm1); first exact: hwlog. + move=> accr_p smp; rewrite -[p]opprK; apply: strong_monoN. + apply: hwlog=> //; do ?exact: strong_monoN. + exact: accr_posN. +case=> [[k1 k1_gt0 hk1]] h1. +move=> [] accr2_p; last first. + set m := (r2 * a1 + r1 * a2) / (r1 + r2). + have pm_gt0 := h1 m. + case: accr2_p=> [_] /(_ m) pm_lt0. + suff: 0 < 0 :> F by rewrite ltrr. + have r_gt0 : 0 < r1 + r2 by rewrite ?addr_gt0. + apply: (ltr_trans (pm_gt0 _) (pm_lt0 _)). + rewrite -(@ler_pmul2l _ (r1 + r2)) //. + rewrite -{1}[r1 + r2]ger0_norm ?(ltrW r_gt0) //. + rewrite -normrM mulrBr /m mulrC mulrVK ?unitfE ?gtr_eqF //. + rewrite mulrDl opprD addrA addrC addrA addKr. + rewrite distrC -mulrBr normrM ger0_norm ?(ltrW r1_gt0) //. + by rewrite mulrC ler_wpmul2r // ltrW. + rewrite -(@ler_pmul2l _ (r1 + r2)) //. + rewrite -{1}[r1 + r2]ger0_norm ?(ltrW r_gt0) //. + rewrite -normrM mulrBr /m mulrC mulrVK ?unitfE ?gtr_eqF //. + rewrite mulrDl opprD addrA addrK. + rewrite -mulrBr normrM ger0_norm ?(ltrW r2_gt0) //. + by rewrite mulrC ler_wpmul2r // ltrW. +case: accr2_p=> [[k2 k2_gt0 hk2]] h2. +left; split; last by move=> x; rewrite split_interval // => /orP [/h1|/h2]. +exists (minr k1 k2); first by rewrite ltr_minr k1_gt0. +move=> x y neq_xy; rewrite !split_interval //. +wlog lt_xy: x y neq_xy / y < x. + move=> hwlog; have [] := ltrP y x; first exact: hwlog. + rewrite ler_eqVlt (negPf neq_xy) /= => /hwlog hwlog' hx hy. + rewrite -mulrNN -!invrN !opprB. + by apply: hwlog'; rewrite // eq_sym. +move=> {h1} {h2} {sm1}. +wlog le_xr1 : a1 a2 r1 r2 k1 k2 + r1_gt0 r2_gt0 k1_gt0 k2_gt0 har hk1 hk2 / `|x - a1| <= r1. + move=> hwlog h; move: (h)=> /orP [/hwlog|]; first exact. + move=> /(hwlog a2 a1 r2 r1 k2 k1) hwlog' ley; rewrite minrC. + by apply: hwlog'; rewrite 1?orbC // distrC [r2 + _]addrC. +move=> _. +have [le_yr1|gt_yr1] := (lerP _ r1)=> /= [_|le_yr2]. + by rewrite ltr_minl hk1. +rewrite ltr_pdivl_mulr ?subr_gt0 //. +pose z := a1 - r1. +have hz1 : `|z - a1| <= r1 by rewrite addrC addKr normrN gtr0_norm. +have gt_yr1' : y + r1 < a1. + rewrite addrC; move: gt_yr1. + rewrite (monoLR (addrNK _) (ltr_add2r _)). + rewrite /z ltr_normr opprB=> /orP[|-> //]. + rewrite (monoRL (addrK a1) (ltr_add2r _))=> /ltr_trans /(_ lt_xy). + by rewrite ltrNge addrC; move: le_xr1; rewrite ler_distl=> /andP [_ ->]. +have lt_yz : y < z by rewrite (monoRL (addrK _) (ltr_add2r _)). +have hz2 : `|z - a2| <= r2. + move: (har); rewrite ler_norml=> /andP [la ua]. + rewrite addrAC ler_distl ua andbT. + rewrite -[a1](addrNK y) -[_ - _ + _ - _]addrA. + rewrite ler_add //. + by rewrite (monoRL (addrK _) (ler_add2r _)) addrC ltrW. + by move: le_yr2; rewrite ler_norml=> /andP[]. +have [<-|neq_zx] := eqVneq z x. + by rewrite -ltr_pdivl_mulr ?subr_gt0 // ltr_minl hk2 ?orbT // gtr_eqF. +have lt_zx : z < x. + rewrite ltr_neqAle neq_zx /=. + move: le_xr1; rewrite distrC ler_norml=> /andP[_]. + by rewrite !(monoLR (addrK _) (ler_add2r _)) addrC. +rewrite -{1}[x](addrNK z) -{1}[p.[x]](addrNK p.[z]). +rewrite !addrA -![_ - _ + _ - _]addrA mulrDr ltr_add //. + rewrite -ltr_pdivl_mulr ?subr_gt0 //. + by rewrite ltr_minl hk1 ?gtr_eqF. +rewrite -ltr_pdivl_mulr ?subr_gt0 //. +by rewrite ltr_minl hk2 ?orbT ?gtr_eqF. +Qed. + +End monotony. + +Section CauchyReals. + +Local Open Scope nat_scope. +Local Open Scope creal_scope. +Local Open Scope ring_scope. + +Definition asympt1 (R : numDomainType) (P : R -> nat -> Prop) + := {m : R -> nat | forall eps i, 0 < eps -> (m eps <= i)%N -> P eps i}. + +Definition asympt2 (R : numDomainType) (P : R -> nat -> nat -> Prop) + := {m : R -> nat | forall eps i j, 0 < eps -> (m eps <= i)%N -> (m eps <= j)%N -> P eps i j}. + +Notation "{ 'asympt' e : i / P }" := (asympt1 (fun e i => P)) + (at level 0, e ident, i ident, format "{ 'asympt' e : i / P }") : type_scope. + +Notation "{ 'asympt' e : i j / P }" := (asympt2 (fun e i j => P)) + (at level 0, e ident, i ident, j ident, format "{ 'asympt' e : i j / P }") : type_scope. + +Lemma asympt1modP (R : numDomainType) P (a : asympt1 P) e i : + 0 < e :> R -> (projT1 a e <= i)%N -> P e i. +Proof. by case: a e i. Qed. + +Lemma asympt2modP (R : numDomainType) P (a : asympt2 P) e i j : + 0 < e :> R -> (projT1 a e <= i)%N -> (projT1 a e <= j)%N -> P e i j. +Proof. by case: a e i j. Qed. + +Variable F : realFieldType. + +(* Lemma asympt_mulLR (k : F) (hk : 0 < k) (P : F -> nat -> Prop) : *) +(* {asympt e : i / P e i} -> {asympt e : i / P (e * k) i}. *) +(* Proof. *) +(* case=> m hm; exists (fun e => m (e * k))=> e i he hi. *) +(* by apply: hm=> //; rewrite -ltr_pdivr_mulr // mul0r. *) +(* Qed. *) + +(* Lemma asympt_mulRL (k : F) (hk : 0 < k) (P : F -> nat -> Prop) : *) +(* {asympt e : i / P (e * k) i} -> {asympt e : i / P e i}. *) +(* Proof. *) +(* case=> m hm; exists (fun e => m (e / k))=> e i he hi. *) +(* rewrite -[e](@mulfVK _ k) ?gtr_eqF //. *) +(* by apply: hm=> //; rewrite -ltr_pdivr_mulr ?invr_gt0 // mul0r. *) +(* Qed. *) + +Lemma asymptP (P1 : F -> nat -> Prop) (P2 : F -> nat -> Prop) : + (forall e i, 0 < e -> P1 e i -> P2 e i) -> + {asympt e : i / P1 e i} -> {asympt e : i / P2 e i}. +Proof. +by move=> hP; case=> m hm; exists m=> e i he me; apply: hP=> //; apply: hm. +Qed. + +(* Lemma asympt2_mulLR (k : F) (hk : 0 < k) (P : F -> nat -> nat -> Prop) : *) +(* {asympt e : i j / P e i j} -> {asympt e : i j / P (e * k) i j}. *) +(* Proof. *) +(* case=> m hm; exists (fun e => m (e * k))=> e i j he hi hj. *) +(* by apply: hm=> //; rewrite -ltr_pdivr_mulr // mul0r. *) +(* Qed. *) + +(* Lemma asympt2_mulRL (k : F) (hk : 0 < k) (P : F -> nat -> nat -> Prop) : *) +(* {asympt e : i j / P (e * k) i j} -> {asympt e : i j / P e i j}. *) +(* Proof. *) +(* case=> m hm; exists (fun e => m (e / k))=> e i j he hi hj. *) +(* rewrite -[e](@mulfVK _ k) ?gtr_eqF //. *) +(* by apply: hm=> //; rewrite -ltr_pdivr_mulr ?invr_gt0 // mul0r. *) +(* Qed. *) + +(* Lemma asympt2P (P1 : F -> nat -> nat -> Prop) (P2 : F -> nat -> nat -> Prop) : *) +(* (forall e i j, 0 < e -> P1 e i j -> P2 e i j) -> *) +(* {asympt e : i j / P1 e i j} -> {asympt e : i j / P2 e i j}. *) +(* Proof. *) +(* move=> hP; case=> m hm; exists m=> e i j he mei mej. *) +(* by apply: hP=> //; apply: hm. *) +(* Qed. *) + +Lemma splitf (n : nat) (e : F) : e = iterop n +%R (e / n%:R) e. +Proof. +case: n=> // n; set e' := (e / _). +have -> : e = e' * n.+1%:R by rewrite mulfVK ?pnatr_eq0. +move: e'=> {e} e; rewrite iteropS. +by elim: n=> /= [|n <-]; rewrite !mulr_natr ?mulr1n. +Qed. + +Lemma splitD (x y e : F) : x < e / 2%:R -> y < e / 2%:R -> x + y < e. +Proof. by move=> hx hy; rewrite [e](splitf 2) ltr_add. Qed. + +Lemma divrn_gt0 (e : F) (n : nat) : 0 < e -> (0 < n)%N -> 0 < e / n%:R. +Proof. by move=> e_gt0 n_gt0; rewrite pmulr_rgt0 ?gtr0E. Qed. + +Lemma split_norm_add (x y e : F) : + `|x| < e / 2%:R -> `|y| < e / 2%:R -> `|x + y| < e. +Proof. by move=> hx hy; rewrite (ler_lt_trans (ler_norm_add _ _)) // splitD. Qed. + +Lemma split_norm_sub (x y e : F) : + `|x| < e / 2%:R -> `|y| < e / 2%:R -> `|x - y| < e. +Proof. by move=> hx hy; rewrite (ler_lt_trans (ler_norm_sub _ _)) // splitD. Qed. + +Lemma split_dist_add (z x y e : F) : + `|x - z| < e / 2%:R -> `|z - y| < e / 2%:R -> `|x - y| < e. +Proof. +by move=> *; rewrite (ler_lt_trans (ler_dist_add z _ _)) ?splitD // 1?distrC. +Qed. + +Definition creal_axiom (x : nat -> F) := {asympt e : i j / `|x i - x j| < e}. + +CoInductive creal := CReal {cauchyseq :> nat -> F; _ : creal_axiom cauchyseq}. +Bind Scope creal_scope with creal. + +Lemma crealP (x : creal) : {asympt e : i j / `|x i - x j| < e}. +Proof. by case: x. Qed. + +Definition cauchymod := + nosimpl (fun (x : creal) => let: CReal _ m := x in projT1 m). + +Lemma cauchymodP (x : creal) eps i j : 0 < eps -> + (cauchymod x eps <= i)%N -> (cauchymod x eps <= j)%N -> `|x i - x j| < eps. +Proof. by case: x=> [x [m mP] //] /mP; apply. Qed. + +Definition neq_creal (x y : creal) : Prop := + exists eps, (0 < eps) && + (eps * 3%:R <= `|x (cauchymod x eps) - y (cauchymod y eps)|). +Notation "!=%CR" := neq_creal : creal_scope. +Notation "x != y" := (neq_creal x y) : creal_scope. + +Definition eq_creal x y := (~ (x != y)%CR). +Notation "x == y" := (eq_creal x y) : creal_scope. + +Lemma ltr_distl_creal (e : F) (i : nat) (x : creal) (j : nat) (a b : F) : + 0 < e -> (cauchymod x e <= i)%N -> (cauchymod x e <= j)%N -> + `| x i - a | <= b - e -> `| x j - a | < b. +Proof. +move=> e_gt0 hi hj hb. +rewrite (ler_lt_trans (ler_dist_add (x i) _ _)) ?ltr_le_add //. +by rewrite -[b](addrNK e) addrC ler_lt_add ?cauchymodP. +Qed. + +Lemma ltr_distr_creal (e : F) (i : nat) (x : creal) (j : nat) (a b : F) : + 0 < e -> (cauchymod x e <= i)%N -> (cauchymod x e <= j)%N -> + a + e <= `| x i - b | -> a < `| x j - b |. +Proof. +move=> e_gt0 hi hj hb; apply: contraLR hb; rewrite -ltrNge -lerNgt. +by move=> ha; rewrite (@ltr_distl_creal e j) // addrK. +Qed. + +(* Lemma asympt_neq (x y : creal) : x != y -> *) +(* {e | 0 < e & forall i, (cauchymod x e <= i)%N -> *) +(* (cauchymod y e <= i)%N -> `|x i - y i| >= e}. *) +(* Proof. *) +(* case/sigW=> e /andP[e_gt0 hxy]. *) +(* exists e=> // i hi hj; move: hxy; rewrite !lerNgt; apply: contra=> hxy. *) +(* rewrite !mulrDr !mulr1 distrC (@ltr_distl_creal i) //. *) +(* by rewrite distrC ltrW // (@ltr_distl_creal i) // ltrW. *) +(* Qed. *) + +Definition lbound (x y : creal) (neq_xy : x != y) : F := + projT1 (sigW neq_xy). + +Lemma lboundP (x y : creal) (neq_xy : x != y) i : + (cauchymod x (lbound neq_xy) <= i)%N -> + (cauchymod y (lbound neq_xy) <= i)%N -> lbound neq_xy <= `|x i - y i|. +Proof. +rewrite /lbound; case: (sigW _)=> /= d /andP[d_gt0 hd] hi hj. +apply: contraLR hd; rewrite -!ltrNge=> hd. +rewrite (@ltr_distl_creal d i) // distrC ltrW // (@ltr_distl_creal d i) //. +by rewrite distrC ltrW // !mulrDr mulr1 !addrA !addrK. +Qed. + +Notation lbound_of p := (@lboundP _ _ p _ _ _). + +Lemma lbound_gt0 (x y : creal) (neq_xy : x != y) : lbound neq_xy > 0. +Proof. by rewrite /lbound; case: (sigW _)=> /= d /andP[]. Qed. + +Definition lbound_ge0 x y neq_xy := (ltrW (@lbound_gt0 x y neq_xy)). + +Lemma cst_crealP (x : F) : creal_axiom (fun _ => x). +Proof. by exists (fun _ => 0%N)=> *; rewrite subrr normr0. Qed. +Definition cst_creal (x : F) := CReal (cst_crealP x). +Notation "x %:CR" := (cst_creal x) + (at level 2, left associativity, format "x %:CR") : creal_scope. +Notation "0" := (0 %:CR) : creal_scope. + +Lemma lbound0P (x : creal) (x_neq0 : x != 0) i : + (cauchymod x (lbound x_neq0) <= i)%N -> + (cauchymod 0%CR (lbound x_neq0) <= i)%N -> lbound x_neq0 <= `|x i|. +Proof. by move=> cx c0; rewrite -[X in `|X|]subr0 -[0]/(0%CR i) lboundP. Qed. + +Notation lbound0_of p := (@lbound0P _ p _ _ _). + +Lemma neq_crealP e i j (e_gt0 : 0 < e) (x y : creal) : + (cauchymod x (e / 5%:R) <= i)%N -> (cauchymod y (e / 5%:R) <= j)%N -> + e <= `|x i - y j| -> x != y. +Proof. +move=> hi hj he; exists (e / 5%:R); rewrite pmulr_rgt0 ?gtr0E //=. +rewrite distrC ltrW // (@ltr_distr_creal (e / 5%:R) j) ?pmulr_rgt0 ?gtr0E //. +rewrite distrC ltrW // (@ltr_distr_creal (e / 5%:R) i) ?pmulr_rgt0 ?gtr0E //. +by rewrite mulr_natr -!mulrSr -mulrnAr -mulr_natr mulVf ?pnatr_eq0 ?mulr1. +Qed. + +Lemma eq_crealP (x y : creal) : {asympt e : i / `|x i - y i| < e} -> + (x == y)%CR. +Proof. +case=> m hm neq_xy; pose d := lbound neq_xy. +pose_big_enough i. + have := (hm d i); rewrite lbound_gt0; big_enough => /(_ isT isT). + by apply/negP; rewrite -lerNgt lboundP. +by close. +Qed. + +Lemma eq0_crealP (x : creal) : {asympt e : i / `|x i| < e} -> x == 0. +Proof. +by move=> hx; apply: eq_crealP; apply: asymptP hx=> e i; rewrite subr0. +Qed. + +Lemma asympt_eq (x y : creal) (eq_xy : x == y) : + {asympt e : i / `|x i - y i| < e}. +Proof. +exists_big_modulus m F. + move=> e i e0 hi; rewrite ltrNge; apply/negP=> he; apply: eq_xy. + by apply: (@neq_crealP e i i). +by close. +Qed. + +Lemma asympt_eq0 (x : creal) : x == 0 -> {asympt e : i / `|x i| < e}. +Proof. by move/asympt_eq; apply: asymptP=> e i; rewrite subr0. Qed. + +Definition eq_mod (x y : creal) (eq_xy : x == y) := projT1 (asympt_eq eq_xy). +Lemma eq_modP (x y : creal) (eq_xy : x == y) eps i : 0 < eps -> + (eq_mod eq_xy eps <= i)%N -> `|x i - y i| < eps. +Proof. +by move=> eps_gt0; rewrite /eq_mod; case: (asympt_eq _)=> /= m hm /hm; apply. +Qed. +Lemma eq0_modP (x : creal) (x_eq0 : x == 0) eps i : 0 < eps -> + (eq_mod x_eq0 eps <= i)%N -> `|x i| < eps. +Proof. +by move=> eps_gt0 hi; rewrite -[X in `|X|]subr0 -[0]/(0%CR i) eq_modP. +Qed. + +Lemma eq_creal_refl x : x == x. +Proof. +apply: eq_crealP; exists (fun _ => 0%N). +by move=> e i e_gt0 _; rewrite subrr normr0. +Qed. +Hint Resolve eq_creal_refl. + +Lemma neq_creal_sym x y : x != y -> y != x. +Proof. +move=> neq_xy; pose_big_enough i. + apply: (@neq_crealP (lbound neq_xy) i i); + by rewrite ?lbound_gt0 1?distrC ?(lbound_of neq_xy). +by close. +Qed. + +Lemma eq_creal_sym x y : x == y -> y == x. +Proof. by move=> eq_xy /neq_creal_sym. Qed. + +Lemma eq_creal_trans x y z : x == y -> y == z -> x == z. +Proof. +move=> eq_xy eq_yz; apply: eq_crealP; exists_big_modulus m F. + by move=> e i *; rewrite (@split_dist_add (y i)) ?eq_modP ?divrn_gt0. +by close. +Qed. + +Lemma creal_neq_always (x y : creal) i (neq_xy : x != y) : + (cauchymod x (lbound neq_xy) <= i)%N -> + (cauchymod y (lbound neq_xy) <= i)%N -> (x i != y i)%B. +Proof. +move=> hx hy; rewrite -subr_eq0 -normr_gt0. +by rewrite (ltr_le_trans _ (lbound_of neq_xy)) ?lbound_gt0. +Qed. + +Definition creal_neq0_always (x : creal) := @creal_neq_always x 0. + +Definition lt_creal (x y : creal) : Prop := + exists eps, (0 < eps) && + (x (cauchymod x eps) + eps * 3%:R <= y (cauchymod y eps)). +Notation "<%CR" := lt_creal : creal_scope. +Notation "x < y" := (lt_creal x y) : creal_scope. + +Definition le_creal (x y : creal) : Prop := ~ (y < x)%CR. +Notation "<=%CR" := le_creal : creal_scope. +Notation "x <= y" := (le_creal x y) : creal_scope. + +Lemma ltr_creal (e : F) (i : nat) (x : creal) (j : nat) (a : F) : + 0 < e -> (cauchymod x e <= i)%N -> (cauchymod x e <= j)%N -> + x i <= a - e -> x j < a. +Proof. +move=> e_gt0 hi hj ha; have := cauchymodP e_gt0 hj hi. +rewrite ltr_distl=> /andP[_ /ltr_le_trans-> //]. +by rewrite -(ler_add2r (- e)) addrK. +Qed. + +Lemma gtr_creal (e : F) (i : nat) (x : creal) (j : nat) (a : F) : + 0 < e -> (cauchymod x e <= i)%N -> (cauchymod x e <= j)%N -> + a + e <= x i-> a < x j. +Proof. +move=> e_gt0 hi hj ha; have := cauchymodP e_gt0 hj hi. +rewrite ltr_distl=> /andP[/(ler_lt_trans _)-> //]. +by rewrite -(ler_add2r e) addrNK. +Qed. + +Definition diff (x y : creal) (lt_xy : (x < y)%CR) : F := projT1 (sigW lt_xy). + +Lemma diff_gt0 (x y : creal) (lt_xy : (x < y)%CR) : diff lt_xy > 0. +Proof. by rewrite /diff; case: (sigW _)=> /= d /andP[]. Qed. + +Definition diff_ge0 x y lt_xy := (ltrW (@diff_gt0 x y lt_xy)). + +Lemma diffP (x y : creal) (lt_xy : (x < y)%CR) i : + (cauchymod x (diff lt_xy) <= i)%N -> + (cauchymod y (diff lt_xy) <= i)%N -> x i + diff lt_xy <= y i. +Proof. +rewrite /diff; case: (sigW _)=> /= e /andP[e_gt0 he] hi hj. +rewrite ltrW // (@gtr_creal e (cauchymod y e)) // (ler_trans _ he) //. +rewrite !mulrDr mulr1 !addrA !ler_add2r ltrW //. +by rewrite (@ltr_creal e (cauchymod x e)) // addrK. +Qed. + +Notation diff_of p := (@diffP _ _ p _ _ _). + +Lemma diff0P (x : creal) (x_gt0 : (0 < x)%CR) i : + (cauchymod x (diff x_gt0) <= i)%N -> + (cauchymod 0%CR (diff x_gt0) <= i)%N -> diff x_gt0 <= x i. +Proof. by move=> cx c0; rewrite -[diff _]add0r -[0]/(0%CR i) diffP. Qed. + +Notation diff0_of p := (@diff0P _ p _ _ _). + +Lemma lt_crealP e i j (e_gt0 : 0 < e) (x y : creal) : + (cauchymod x (e / 5%:R) <= i)%N -> (cauchymod y (e / 5%:R) <= j)%N -> + x i + e <= y j -> (x < y)%CR. +Proof. +move=> hi hj he; exists (e / 5%:R); rewrite pmulr_rgt0 ?gtr0E //=. +rewrite ltrW // (@gtr_creal (e / 5%:R) j) ?pmulr_rgt0 ?gtr0E //. +rewrite (ler_trans _ he) // -addrA (monoLR (addrNK _) (ler_add2r _)). +rewrite ltrW // (@ltr_creal (e / 5%:R) i) ?pmulr_rgt0 ?gtr0E //. +rewrite -!addrA ler_addl !addrA -mulrA -{1}[e]mulr1 -!(mulrBr, mulrDr). +rewrite pmulr_rge0 // {1}[1](splitf 5) /= !mul1r !mulrDr mulr1. +by rewrite !opprD !addrA !addrK addrN. +Qed. + +Lemma le_crealP i (x y : creal) : + (forall j, (i <= j)%N -> x j <= y j) -> (x <= y)%CR. +Proof. +move=> hi lt_yx; pose_big_enough j. + have := hi j; big_enough => /(_ isT); apply/negP; rewrite -ltrNge. + by rewrite (ltr_le_trans _ (diff_of lt_yx)) ?ltr_spaddr ?diff_gt0. +by close. +Qed. + +Lemma le_creal_refl (x : creal) : (x <= x)%CR. +Proof. by apply: (@le_crealP 0%N). Qed. +Hint Resolve le_creal_refl. + +Lemma lt_neq_creal (x y : creal) : (x < y)%CR -> x != y. +Proof. +move=> ltxy; pose_big_enough i. + apply: (@neq_crealP (diff ltxy) i i) => //; first by rewrite diff_gt0. + by rewrite distrC lerNgt ltr_distl negb_and -!lerNgt diffP ?orbT. +by close. +Qed. + +Lemma creal_lt_always (x y : creal) i (lt_xy : (x < y)%CR) : + (cauchymod x (diff lt_xy) <= i)%N -> + (cauchymod y (diff lt_xy) <= i)%N -> x i < y i. +Proof. +by move=> hx hy; rewrite (ltr_le_trans _ (diff_of lt_xy)) ?ltr_addl ?diff_gt0. +Qed. + +Definition creal_gt0_always := @creal_lt_always 0. + +Lemma eq_le_creal (x y : creal) : x == y -> (x <= y)%CR. +Proof. by move=> /eq_creal_sym ? /lt_neq_creal. Qed. + +Lemma asympt_le (x y : creal) (le_xy : (x <= y)%CR) : + {asympt e : i / x i < y i + e}. +Proof. +exists_big_modulus m F. + move=> e i e0 hm; rewrite ltrNge; apply/negP=> he; apply: le_xy. + by apply: (@lt_crealP e i i). +by close. +Qed. + +Lemma asympt_ge0 (x : creal) : (0 <= x)%CR -> {asympt e : i / - e < x i}. +Proof. by move/asympt_le; apply: asymptP=> *; rewrite -subr_gt0 opprK. Qed. + +Definition le_mod (x y : creal) (le_xy : (x <= y)%CR) := projT1 (asympt_le le_xy). + +Lemma le_modP (x y : creal) (le_xy : (x <= y)%CR) eps i : 0 < eps -> + (le_mod le_xy eps <= i)%N -> x i < y i + eps. +Proof. +by move=> eps_gt0; rewrite /le_mod; case: (asympt_le _)=> /= m hm /hm; apply. +Qed. + +Lemma ge0_modP (x : creal) (x_ge0 : (0 <= x)%CR) eps i : 0 < eps -> + (le_mod x_ge0 eps <= i)%N -> - eps < x i. +Proof. +by move=> eps_gt0 hi; rewrite -(ltr_add2r eps) addNr -[0]/(0%CR i) le_modP. +Qed. + +Lemma opp_crealP (x : creal) : creal_axiom (fun i => - x i). +Proof. by case: x=> [x [m mP]]; exists m=> *; rewrite /= -opprD normrN mP. Qed. +Definition opp_creal (x : creal) := CReal (opp_crealP x). +Notation "-%CR" := opp_creal : creal_scope. +Notation "- x" := (opp_creal x) : creal_scope. + +Lemma add_crealP (x y : creal) : creal_axiom (fun i => x i + y i). +Proof. +exists_big_modulus m F. + move=> e i j he hi hj; rewrite opprD addrAC addrA -addrA [- _ + _]addrC. + by rewrite split_norm_add ?cauchymodP ?divrn_gt0. +by close. +Qed. +Definition add_creal (x y : creal) := CReal (add_crealP x y). +Notation "+%CR" := add_creal : creal_scope. +Notation "x + y" := (add_creal x y) : creal_scope. +Notation "x - y" := (x + - y)%CR : creal_scope. + + +Lemma ubound_subproof (x : creal) : {b : F | b > 0 & forall i, `|x i| <= b}. +Proof. +pose_big_enough i; first set b := 1 + `|x i|. + exists (foldl maxr b [seq `|x n| | n <- iota 0 i]) => [|n]. + have : 0 < b by rewrite ltr_spaddl. + by elim: iota b => //= a l IHl b b_gt0; rewrite IHl ?ltr_maxr ?b_gt0. + have [|le_in] := (ltnP n i). + elim: i b => [|i IHi] b //. + rewrite ltnS -addn1 iota_add add0n map_cat foldl_cat /= ler_maxr leq_eqVlt. + by case/orP=> [/eqP->|/IHi->] //; rewrite lerr orbT. + set xn := `|x n|; suff : xn <= b. + by elim: iota xn b => //= a l IHl xn b Hxb; rewrite IHl ?ler_maxr ?Hxb. + rewrite -ler_subl_addr (ler_trans (ler_norm _)) //. + by rewrite (ler_trans (ler_dist_dist _ _)) ?ltrW ?cauchymodP. +by close. +Qed. + +Definition ubound (x : creal) := + nosimpl (let: exist2 b _ _ := ubound_subproof x in b). + +Lemma uboundP (x : creal) i : `|x i| <= ubound x. +Proof. by rewrite /ubound; case: ubound_subproof. Qed. + +Lemma ubound_gt0 x : 0 < ubound x. +Proof. by rewrite /ubound; case: ubound_subproof. Qed. + +Definition ubound_ge0 x := (ltrW (ubound_gt0 x)). + +Lemma mul_crealP (x y : creal) : creal_axiom (fun i => x i * y i). +Proof. +exists_big_modulus m F. + move=> e i j e_gt0 hi hj. + rewrite -[_ * _]subr0 -(subrr (x j * y i)) opprD opprK addrA. + rewrite -mulrBl -addrA -mulrBr split_norm_add // !normrM. + have /ler_wpmul2l /ler_lt_trans-> // := uboundP y i. + rewrite -ltr_pdivl_mulr ?ubound_gt0 ?cauchymodP //. + by rewrite !pmulr_rgt0 ?invr_gt0 ?ubound_gt0 ?ltr0n. + rewrite mulrC; have /ler_wpmul2l /ler_lt_trans-> // := uboundP x j. + rewrite -ltr_pdivl_mulr ?ubound_gt0 ?cauchymodP //. + by rewrite !pmulr_rgt0 ?gtr0E ?ubound_gt0. +by close. +Qed. +Definition mul_creal (x y : creal) := CReal (mul_crealP x y). +Notation "*%CR" := mul_creal : creal_scope. +Notation "x * y" := (mul_creal x y) : creal_scope. + +Lemma inv_crealP (x : creal) (x_neq0 : x != 0) : creal_axiom (fun i => (x i)^-1). +Proof. +pose d := lbound x_neq0. +exists_big_modulus m F. + (* exists (fun e => [CC x # e * d ^+ 2; ! x_neq0]). *) + move=> e i j e_gt0 hi hj. + have /andP[xi_neq0 xj_neq0] : (x i != 0) && (x j != 0). + by rewrite -!normr_gt0 !(ltr_le_trans _ (lbound0_of x_neq0)) ?lbound_gt0. + rewrite -(@ltr_pmul2r _ `|x i * x j|); last by rewrite normr_gt0 mulf_neq0. + rewrite -normrM !mulrBl mulrA mulVf // mulrCA mulVf // mul1r mulr1. + apply: (@ltr_le_trans _ (e * d ^+ 2)). + by apply: cauchymodP; rewrite // !pmulr_rgt0 ?lbound_gt0. + rewrite ler_wpmul2l ?(ltrW e_gt0) // normrM. + have /(_ j) hx := lbound0_of x_neq0; rewrite /=. + have -> // := (ler_trans (@ler_wpmul2l _ d _ _ _ (hx _ _))). + by rewrite ltrW // lbound_gt0. + by rewrite ler_wpmul2r ?normr_ge0 // lbound0P. +by close. +Qed. +Definition inv_creal (x : creal) (x_neq0 : x != 0) := CReal (inv_crealP x_neq0). +Notation "x_neq0 ^-1" := (inv_creal x_neq0) : creal_scope. +Notation "x / y_neq0" := (x * (y_neq0 ^-1))%CR : creal_scope. + +Lemma norm_crealP (x : creal) : creal_axiom (fun i => `|x i|). +Proof. +exists (cauchymod x). +by move=> *; rewrite (ler_lt_trans (ler_dist_dist _ _)) ?cauchymodP. +Qed. +Definition norm_creal x := CReal (norm_crealP x). +Local Notation "`| x |" := (norm_creal x) : creal_scope. + +Lemma horner_crealP (p : {poly F}) (x : creal) : + creal_axiom (fun i => p.[x i]). +Proof. +exists_big_modulus m F=> [e i j e_gt0 hi hj|]. + rewrite (ler_lt_trans (@poly_accr_bound1P _ p (x (cauchymod x 1)) 1 _ _ _ _)); + do ?[by rewrite ?e_gt0 | by rewrite ltrW // cauchymodP]. + rewrite -ltr_pdivl_mulr ?poly_accr_bound_gt0 ?cauchymodP //. + by rewrite pmulr_rgt0 ?invr_gt0 ?poly_accr_bound_gt0. +by close. +Qed. +Definition horner_creal (p : {poly F}) (x : creal) := CReal (horner_crealP p x). +Notation "p .[ x ]" := (horner_creal p x) : creal_scope. + +Lemma neq_creal_horner p (x y : creal) : p.[x] != p.[y] -> x != y. +Proof. +move=> neq_px_py. +pose d := lbound neq_px_py. +pose_big_enough i. + pose k := 2%:R + poly_accr_bound p (y i) d. + have /andP[d_gt0 k_gt0] : (0 < d) && (0 < k). + rewrite ?(ltr_spaddl, poly_accr_bound_ge0); + by rewrite ?ltr0n ?ltrW ?ltr01 ?lbound_gt0. + pose_big_enough j. + apply: (@neq_crealP (d / k) j j) => //. + by rewrite ?(pmulr_lgt0, invr_gt0, ltr0n). + rewrite ler_pdivr_mulr //. + have /(_ j) // := (lbound_of neq_px_py). + big_enough=> /(_ isT isT). + apply: contraLR; rewrite -!ltrNge=> hxy. + rewrite (ler_lt_trans (@poly_accr_bound1P _ _ (y i) d _ _ _ _)) //. + + by rewrite ltrW // cauchymodP. + + rewrite ltrW // (@split_dist_add (y j)) //; last first. + by rewrite cauchymodP ?divrn_gt0. + rewrite ltr_pdivl_mulr ?ltr0n // (ler_lt_trans _ hxy) //. + by rewrite ler_wpmul2l ?normr_ge0 // ler_paddr // poly_accr_bound_ge0. + rewrite (ler_lt_trans _ hxy) // ler_wpmul2l ?normr_ge0 //. + by rewrite ler_paddl // ?ler0n. + by close. +by close. +Qed. + +Lemma eq_creal_horner p (x y : creal) : x == y -> p.[x] == p.[y]. +Proof. by move=> hxy /neq_creal_horner. Qed. + +Import Setoid Relation_Definitions. + +Add Relation creal eq_creal + reflexivity proved by eq_creal_refl + symmetry proved by eq_creal_sym + transitivity proved by eq_creal_trans +as eq_creal_rel. +Global Existing Instance eq_creal_rel. + +Add Morphism add_creal with + signature eq_creal ==> eq_creal ==> eq_creal as add_creal_morph. +Proof. +move=> x y eq_xy z t eq_zt; apply: eq_crealP. +exists_big_modulus m F. + move=> e i e_gt0 hi; rewrite opprD addrA [X in X + _]addrAC -addrA. + by rewrite split_norm_add ?eq_modP ?divrn_gt0. +by close. +Qed. +Global Existing Instance add_creal_morph_Proper. + + +Add Morphism opp_creal with + signature eq_creal ==> eq_creal as opp_creal_morph. +Proof. +move=> x y /asympt_eq [m hm]; apply: eq_crealP; exists m. +by move=> e i e_gt0 hi /=; rewrite -opprD normrN hm. +Qed. +Global Existing Instance opp_creal_morph_Proper. + +Add Morphism mul_creal with + signature eq_creal ==> eq_creal ==> eq_creal as mul_creal_morph. +Proof. +move=> x y eq_xy z t eq_zt; apply: eq_crealP. +exists_big_modulus m F. + move=> e i e_gt0 hi. + rewrite (@split_dist_add (y i * z i)) // -(mulrBl, mulrBr) normrM. + have /ler_wpmul2l /ler_lt_trans-> // := uboundP z i. + rewrite -ltr_pdivl_mulr ?ubound_gt0 ?eq_modP //. + by rewrite !pmulr_rgt0 ?invr_gt0 ?ubound_gt0 ?ltr0n. + rewrite mulrC; have /ler_wpmul2l /ler_lt_trans-> // := uboundP y i. + rewrite -ltr_pdivl_mulr ?ubound_gt0 ?eq_modP //. + by rewrite !pmulr_rgt0 ?invr_gt0 ?ubound_gt0 ?ltr0n. +by close. +Qed. +Global Existing Instance mul_creal_morph_Proper. + +Lemma eq_creal_inv (x y : creal) (x_neq0 : x != 0) (y_neq0 : y != 0) : + (x == y) -> (x_neq0^-1 == y_neq0^-1). +Proof. +move=> eq_xy; apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi /=. + rewrite -(@ltr_pmul2r _ (lbound x_neq0 * lbound y_neq0)); + do ?by rewrite ?pmulr_rgt0 ?lbound_gt0. + rewrite (@ler_lt_trans _ (`|(x i)^-1 - (y i)^-1| * (`|x i| * `|y i|))) //. + rewrite ler_wpmul2l ?normr_ge0 //. + rewrite (@ler_trans _ (`|x i| * lbound y_neq0)) //. + by rewrite ler_wpmul2r ?lbound_ge0 ?lbound0P. + by rewrite ler_wpmul2l ?normr_ge0 ?lbound0P. + rewrite -!normrM mulrBl mulKf ?creal_neq0_always //. + rewrite mulrCA mulVf ?mulr1 ?creal_neq0_always //. + by rewrite distrC eq_modP ?pmulr_rgt0 ?lbound_gt0. +by close. +Qed. + +Add Morphism horner_creal with + signature (@eq _) ==> eq_creal ==> eq_creal as horner_creal_morph. +Proof. exact: eq_creal_horner. Qed. +Global Existing Instance horner_creal_morph_Proper. + +Add Morphism lt_creal with + signature eq_creal ==> eq_creal ==> iff as lt_creal_morph. +Proof. +move=> x y eq_xy z t eq_zt. +wlog lxz : x y z t eq_xy eq_zt / (x < z)%CR. + move=> hwlog; split=> h1; move: (h1) => /hwlog; apply=> //; + by apply: eq_creal_sym. +split=> // _. +pose e' := diff lxz / 4%:R. +have e'_gt0 : e' > 0 by rewrite pmulr_rgt0 ?gtr0E ?diff_gt0. +have le_zt : (z <= t)%CR by apply: eq_le_creal. +have le_xy : (y <= x)%CR by apply: eq_le_creal; apply: eq_creal_sym. +pose_big_enough i. + apply: (@lt_crealP e' i i)=> //. + rewrite ltrW // -(ltr_add2r e'). + rewrite (ler_lt_trans _ (@le_modP _ _ le_zt _ _ _ _)) //. + rewrite -addrA (monoLR (@addrNK _ _) (@ler_add2r _ _)) ltrW //. + rewrite (ltr_le_trans (@le_modP _ _ le_xy e' _ _ _)) //. + rewrite -(monoLR (@addrNK _ _) (@ler_add2r _ _)) ltrW //. + rewrite (ltr_le_trans _ (diff_of lxz)) //. + rewrite -addrA ler_lt_add // /e' -!mulrDr gtr_pmulr ?diff_gt0 //. + by rewrite [X in _ < X](splitf 4) /= mul1r !ltr_addr ?gtr0E. +by close. +Qed. +Global Existing Instance lt_creal_morph_Proper. + +Add Morphism le_creal with + signature eq_creal ==> eq_creal ==> iff as le_creal_morph. +Proof. by move=> x y exy z t ezt; rewrite /le_creal exy ezt. Qed. +Global Existing Instance le_creal_morph_Proper. + +Add Morphism norm_creal + with signature eq_creal ==> eq_creal as norm_creal_morph. +Proof. +move=> x y hxy; apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi. + by rewrite (ler_lt_trans (ler_dist_dist _ _)) ?eq_modP. +by close. +Qed. +Global Existing Instance norm_creal_morph_Proper. + +Lemma neq_creal_ltVgt (x y : creal) : x != y -> {(x < y)%CR} + {(y < x)%CR}. +Proof. +move=> neq_xy; pose_big_enough i. + have := (@lboundP _ _ neq_xy i); big_enough => /(_ isT isT). + have [le_xy|/ltrW le_yx'] := lerP (x i) (y i). + rewrite -(ler_add2r (x i)) ?addrNK addrC. + move=> /lt_crealP; rewrite ?lbound_gt0; big_enough. + by do 3!move/(_ isT); left. + rewrite -(ler_add2r (y i)) ?addrNK addrC. + move=> /lt_crealP; rewrite ?lbound_gt0; big_enough. + by do 3!move/(_ isT); right. +by close. +Qed. + +Lemma lt_creal_neq (x y : creal) : (x < y -> x != y)%CR. +Proof. +move=> lxy; pose_big_enough i. + apply: (@neq_crealP (diff lxy) i i); rewrite ?diff_gt0 //. + rewrite distrC ler_normr (monoRL (addrK _) (ler_add2r _)) addrC. + by rewrite (diff_of lxy). +by close. +Qed. + +Lemma gt_creal_neq (x y : creal) : (y < x -> x != y)%CR. +Proof. by move/lt_creal_neq /neq_creal_sym. Qed. + +Lemma lt_creal_trans (x y z : creal) : (x < y -> y < z -> x < z)%CR. +Proof. +move=> lt_xy lt_yz; pose_big_enough i. + apply: (@lt_crealP (diff lt_xy + diff lt_yz) i i) => //. + by rewrite addr_gt0 ?diff_gt0. + rewrite (ler_trans _ (diff_of lt_yz)) //. + by rewrite addrA ler_add2r (diff_of lt_xy). +by close. +Qed. + +Lemma lt_crealW (x y : creal) : (x < y)%CR -> (x <= y)%CR. +Proof. by move=> /lt_creal_trans /(_ _) /le_creal_refl. Qed. + +Add Morphism neq_creal with + signature eq_creal ==> eq_creal ==> iff as neq_creal_morph. +Proof. +move=> x y eq_xy z t eq_zt; split=> /neq_creal_ltVgt []. ++ by rewrite eq_xy eq_zt=> /lt_creal_neq. ++ by rewrite eq_xy eq_zt=> /gt_creal_neq. ++ by rewrite -eq_xy -eq_zt=> /lt_creal_neq. +by rewrite -eq_xy -eq_zt=> /gt_creal_neq. +Qed. +Global Existing Instance neq_creal_morph_Proper. + +Local Notation m0 := (fun (_ : F) => 0%N). + +Lemma add_0creal x : 0 + x == x. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite add0r subrr normr0. Qed. + +Lemma add_creal0 x : x + 0 == x. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite addr0 subrr normr0. Qed. + +Lemma mul_creal0 x : x * 0 == 0. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite mulr0 subrr normr0. Qed. + +Lemma mul_0creal x : 0 * x == 0. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite mul0r subrr normr0. Qed. + +Lemma mul_creal1 x : x * 1%:CR == x. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite mulr1 subrr normr0. Qed. + +Lemma mul_1creal x : 1%:CR * x == x. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite mul1r subrr normr0. Qed. + +Lemma opp_creal0 : - 0 == 0. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite oppr0 addr0 normr0. Qed. + +Lemma horner_crealX (x : creal) : 'X.[x] == x. +Proof. by apply: eq_crealP; exists m0=> *; rewrite /= hornerX subrr normr0. Qed. + +Lemma horner_crealM (p q : {poly F}) (x : creal) : + ((p * q).[x] == p.[x] * q.[x])%CR. +Proof. +by apply: eq_crealP; exists m0=> * /=; rewrite hornerM subrr normr0. +Qed. + +Lemma neq_creal_cst x y : reflect (cst_creal x != cst_creal y) (x != y). +Proof. +apply: (iffP idP)=> neq_xy; pose_big_enough i. ++ by apply (@neq_crealP `|x - y| i i); rewrite ?normr_gt0 ?subr_eq0 . ++ by close. ++ by rewrite (@creal_neq_always _ _ i neq_xy). ++ by close. +Qed. + +Lemma eq_creal_cst x y : reflect (cst_creal x == cst_creal y) (x == y). +Proof. +apply: (iffP idP)=> [|eq_xy]; first by move/eqP->. +by apply/negP=> /negP /neq_creal_cst; rewrite eq_xy; apply: eq_creal_refl. +Qed. + +Lemma lt_creal_cst x y : reflect (cst_creal x < cst_creal y)%CR (x < y). +Proof. +apply: (iffP idP)=> lt_xy; pose_big_enough i. ++ apply: (@lt_crealP (y - x) i i); rewrite ?subr_gt0 //=. + by rewrite addrCA subrr addr0. ++ by close. ++ by rewrite (@creal_lt_always _ _ i lt_xy). ++ by close. +Qed. + +Lemma le_creal_cst x y : reflect (cst_creal x <= cst_creal y)%CR (x <= y). +Proof. +apply: (iffP idP)=> [le_xy /lt_creal_cst|eq_xy]; first by rewrite ltrNge le_xy. +by rewrite lerNgt; apply/negP=> /lt_creal_cst. +Qed. + + +Lemma mul_creal_neq0 x y : x != 0 -> y != 0 -> x * y != 0. +Proof. +move=> x_neq0 y_neq0. +pose d := lbound x_neq0 * lbound y_neq0. +have d_gt0 : 0 < d by rewrite pmulr_rgt0 lbound_gt0. +pose_big_enough i. + apply: (@neq_crealP d i i)=> //; rewrite subr0 normrM. + rewrite (@ler_trans _ (`|x i| * lbound y_neq0)) //. + by rewrite ler_wpmul2r ?lbound_ge0 // lbound0P. + by rewrite ler_wpmul2l ?normr_ge0 // lbound0P. +by close. +Qed. + +Lemma mul_neq0_creal x y : x * y != 0 -> y != 0. +Proof. +move=> xy_neq0; pose_big_enough i. + apply: (@neq_crealP ((ubound x)^-1 * lbound xy_neq0) i i) => //. + by rewrite pmulr_rgt0 ?invr_gt0 ?lbound_gt0 ?ubound_gt0. + rewrite subr0 ler_pdivr_mull ?ubound_gt0 //. + have /(_ i)-> // := (ler_trans (lbound0_of xy_neq0)). + by rewrite normrM ler_wpmul2r ?normr_ge0 ?uboundP. +by close. +Qed. + +Lemma poly_mul_creal_eq0_coprime p q x : + coprimep p q -> + p.[x] * q.[x] == 0 -> {p.[x] == 0} + {q.[x] == 0}. +Proof. +move=> /Bezout_eq1_coprimepP /sig_eqW [[u v] /= hpq]; pose_big_enough i. + have := (erefl ((1 : {poly F}).[x i])). + rewrite -{1}hpq /= hornerD hornerC. + set upxi := (u * _).[_]. + move=> hpqi. + have [p_small|p_big] := lerP `|upxi| 2%:R^-1=> pqx0; [left|right]. + move=> px0; apply: pqx0; apply: mul_creal_neq0=> //. + apply: (@mul_neq0_creal v.[x]). + apply: (@neq_crealP 2%:R^-1 i i); rewrite ?gtr0E //. + rewrite /= subr0 -hornerM -(ler_add2l `|upxi|). + rewrite (ler_trans _ (ler_norm_add _ _)) // hpqi normr1. + rewrite (monoLR (addrNK _) (ler_add2r _)). + by rewrite {1}[1](splitf 2) /= mul1r addrK. + move=> qx0; apply: pqx0; apply: mul_creal_neq0=> //. + apply: (@mul_neq0_creal u.[x]). + apply: (@neq_crealP 2%:R^-1 i i); rewrite ?gtr0E //. + by rewrite /= subr0 -hornerM ltrW. +by close. +Qed. + +Lemma dvdp_creal_eq0 p q x : p %| q -> p.[x] == 0 -> q.[x] == 0. +Proof. +by move=> dpq px0; rewrite -[q](divpK dpq) horner_crealM px0 mul_creal0. +Qed. + +Lemma root_poly_expn_creal p k x : (0 < k)%N + -> (p ^+ k).[x] == 0 -> p.[x] == 0. +Proof. +move=> k_gt0 pkx_eq0; apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi; rewrite /= subr0. + rewrite -(@ltr_pexpn2r _ k) -?topredE /= ?normr_ge0 ?ltrW //. + by rewrite -normrX -horner_exp (@eq0_modP _ pkx_eq0) ?exprn_gt0 //. +by close. +Qed. + +Lemma horner_cst_creal c x : c%:P.[x] == c%:CR. +Proof. +apply: eq_crealP; exists (fun _ => 0%N)=> e i e_gt0 _. +by rewrite /= hornerC subrr normr0. +Qed. + +Lemma horner_creal_cst (p : {poly F}) (x : F) : p.[x%:CR] == p.[x]%:CR. +Proof. by apply: eq_crealP; exists m0=> *; rewrite /= subrr normr0. Qed. + + +Lemma poly_mul_creal_eq0 p q x : + p.[x] * q.[x] == 0 -> {p.[x] == 0} + {q.[x] == 0}. +Proof. +move=> mul_px_qx_eq0. +have [->|p_neq0] := altP (p =P 0); first by left; rewrite horner_cst_creal. +have [->|q_neq0] := altP (q =P 0); first by right; rewrite horner_cst_creal. +pose d := gcdp p q; pose p' := gdcop d p; pose q' := gdcop d q. +have cop_q'_d': coprimep p' q'. + rewrite /coprimep size_poly_eq1. + apply: (@coprimepP _ p' d _). + + by rewrite coprimep_gdco. + + by rewrite dvdp_gcdl. + rewrite dvdp_gcd (dvdp_trans (dvdp_gcdl _ _)) ?dvdp_gdco //. + by rewrite (dvdp_trans (dvdp_gcdr _ _)) ?dvdp_gdco. +suff : (p' * q').[x] * (d ^+ (size p + size q)).[x] == 0. + case/poly_mul_creal_eq0_coprime. + + by rewrite coprimep_expr // coprimep_mull ?coprimep_gdco. + + move=> p'q'x_eq0. + have : p'.[x] * q'.[x] == 0 by rewrite -horner_crealM. + case/poly_mul_creal_eq0_coprime=> // /dvdp_creal_eq0 hp'q'. + by left; apply: hp'q'; rewrite dvdp_gdco. + by right; apply: hp'q'; rewrite dvdp_gdco. + move/root_poly_expn_creal. + rewrite addn_gt0 lt0n size_poly_eq0 p_neq0=> /(_ isT) dx_eq0. + by left; apply: dvdp_creal_eq0 dx_eq0; rewrite dvdp_gcdl. +move: mul_px_qx_eq0; rewrite -!horner_crealM. +rewrite exprD mulrAC mulrA -mulrA [_ ^+ _ * _]mulrC. +apply: dvdp_creal_eq0; rewrite ?dvdp_mul // dvdp_gdcor //; +by rewrite gcdp_eq0 negb_and p_neq0. +Qed. + +Lemma coprimep_root (p q : {poly F}) x : + coprimep p q -> p.[x] == 0 -> q.[x] != 0. +Proof. +move=> /Bezout_eq1_coprimepP /sig_eqW [[u v] hpq] px0. +have upx_eq0 : u.[x] * p.[x] == 0 by rewrite px0 mul_creal0. +pose_big_enough i. + have := (erefl ((1 : {poly F}).[x i])). + rewrite -{1}hpq /= hornerD hornerC. + set upxi := (u * _).[_]; move=> hpqi. + apply: (@neq_crealP ((ubound v.[x])%CR^-1 / 2%:R) i i) => //. + by rewrite pmulr_rgt0 ?gtr0E // ubound_gt0. + rewrite /= subr0 ler_pdivr_mull ?ubound_gt0 //. + rewrite (@ler_trans _ `|(v * q).[x i]|) //; last first. + by rewrite hornerM normrM ler_wpmul2r ?normr_ge0 ?(uboundP v.[x]). + rewrite -(ler_add2l `|upxi|) (ler_trans _ (ler_norm_add _ _)) // hpqi normr1. + rewrite (monoLR (addrNK _) (ler_add2r _)). + rewrite {1}[1](splitf 2) /= mul1r addrK ltrW // /upxi hornerM. + by rewrite (@eq0_modP _ upx_eq0) ?gtr0E. +by close. +Qed. + +Lemma deriv_neq0_mono (p : {poly F}) (x : creal) : p^`().[x] != 0 -> + { r : F & 0 < r & + { i : nat & (cauchymod x r <= i)%N & (strong_mono p (x i) r)} }. +Proof. +move=> px_neq0. +wlog : p px_neq0 / (0 < p^`().[x])%CR. + case/neq_creal_ltVgt: (px_neq0)=> px_lt0; last exact. + case/(_ (- p)). + + pose_big_enough i. + apply: (@neq_crealP (lbound px_neq0) i i); do ?by rewrite ?lbound_gt0. + rewrite /= derivN hornerN subr0 normrN. + by rewrite (lbound0_of px_neq0). + by close. + + pose_big_enough i. + apply: (@lt_crealP (diff px_lt0) i i); do ?by rewrite ?diff_gt0. + rewrite /= add0r derivN hornerN -subr_le0 opprK addrC. + by rewrite (diff_of px_lt0) //. + by close. + move=> r r_ge0 [i hi]; move/strong_monoN; rewrite opprK=> sm. + by exists r=> //; exists i. +move=> px_gt0. +pose b1 := poly_accr_bound p^`() 0 (1 + ubound x). +pose b2 := poly_accr_bound2 p 0 (1 + ubound x). +pose r := minr 1 (minr + (diff px_gt0 / 4%:R / b1) + (diff px_gt0 / 4%:R / b2 / 2%:R)). +exists r. + rewrite !ltr_minr ?ltr01 ?pmulr_rgt0 ?gtr0E ?diff_gt0; + by rewrite ?poly_accr_bound2_gt0 ?poly_accr_bound_gt0. +pose_big_enough i. + exists i => //; left; split; last first. + move=> y hy; have := (@poly_accr_bound1P _ p^`() 0 (1 + ubound x) (x i) y). + rewrite ?subr0 ler_paddl ?ler01 ?uboundP //. + rewrite (@ler_trans _ (r + `|x i|)) ?subr0; last 2 first. + + rewrite (monoRL (addrNK _) (ler_add2r _)). + by rewrite (ler_trans (ler_sub_dist _ _)). + + by rewrite ler_add ?ler_minl ?lerr ?uboundP. + move=> /(_ isT isT). + rewrite ler_distl=> /andP[le_py ge_py]. + rewrite (ltr_le_trans _ le_py) // subr_gt0 -/b1. + rewrite (ltr_le_trans _ (diff0_of px_gt0)) //. + rewrite (@ler_lt_trans _ (r * b1)) //. + by rewrite ler_wpmul2r ?poly_accr_bound_ge0. + rewrite -ltr_pdivl_mulr ?poly_accr_bound_gt0 //. + rewrite !ltr_minl ltr_pmul2r ?invr_gt0 ?poly_accr_bound_gt0 //. + by rewrite gtr_pmulr ?diff_gt0 // invf_lt1 ?gtr0E ?ltr1n ?orbT. + exists (diff px_gt0 / 4%:R). + by rewrite pmulr_rgt0 ?gtr0E ?diff_gt0. + move=> y z neq_yz hy hz. + have := (@poly_accr_bound1P _ p^`() 0 (1 + ubound x) (x i) z). + have := @poly_accr_bound2P _ p 0 (1 + ubound x) z y; rewrite eq_sym !subr0. + rewrite neq_yz ?ler01 ?ubound_ge0=> // /(_ isT). + rewrite (@ler_trans _ (r + `|x i|)); last 2 first. + + rewrite (monoRL (addrNK _) (ler_add2r _)). + by rewrite (ler_trans (ler_sub_dist _ _)). + + rewrite ler_add ?ler_minl ?lerr ?uboundP //. + rewrite (@ler_trans _ (r + `|x i|)); last 2 first. + + rewrite (monoRL (addrNK _) (ler_add2r _)). + by rewrite (ler_trans (ler_sub_dist _ _)). + + rewrite ler_add ?ler_minl ?lerr ?uboundP //. + rewrite ler_paddl ?uboundP ?ler01 //. + move=> /(_ isT isT); rewrite ler_distl=> /andP [haccr _]. + move=> /(_ isT isT); rewrite ler_distl=> /andP [hp' _]. + rewrite (ltr_le_trans _ haccr) // (monoRL (addrK _) (ltr_add2r _)). + rewrite (ltr_le_trans _ hp') // (monoRL (addrK _) (ltr_add2r _)). + rewrite (ltr_le_trans _ (diff0_of px_gt0)) //. + rewrite {2}[diff _](splitf 4) /= -!addrA ltr_add2l ltr_spaddl //. + by rewrite pmulr_rgt0 ?gtr0E ?diff_gt0. + rewrite -/b1 -/b2 ler_add //. + rewrite -ler_pdivl_mulr ?poly_accr_bound2_gt0 //. + rewrite (@ler_trans _ (r * 2%:R)) //. + rewrite (ler_trans (ler_dist_add (x i) _ _)) //. + by rewrite mulrDr mulr1 ler_add // distrC. + by rewrite -ler_pdivl_mulr ?ltr0n // !ler_minl lerr !orbT. + rewrite -ler_pdivl_mulr ?poly_accr_bound_gt0 //. + by rewrite (@ler_trans _ r) // !ler_minl lerr !orbT. +by close. +Qed. + +Lemma smaller_factor (p q : {poly F}) x : + p \is monic-> p.[x] == 0 -> + ~~(p %| q) -> ~~ coprimep p q -> + {r : {poly F} | (size r < size p)%N && (r \is monic) & r.[x] == 0}. +Proof. +move=> monic_p px0 ndvd_pq. +rewrite /coprimep; set d := gcdp _ _=> sd_neq1. +pose r1 : {poly F} := (lead_coef d)^-1 *: d. +pose r2 := p %/ r1. +have ld_neq0 : lead_coef d != 0 :> F. + by rewrite lead_coef_eq0 gcdp_eq0 negb_and monic_neq0. +have monic_r1 : r1 \is monic. + by rewrite monicE /r1 -mul_polyC lead_coefM lead_coefC mulVf. +have eq_p_r2r1: p = r2 * r1. + by rewrite divpK // (@eqp_dvdl _ d) ?dvdp_gcdl // eqp_scale ?invr_eq0. +have monic_r2 : r2 \is monic by rewrite -(monicMr _ monic_r1) -eq_p_r2r1. +have eq_sr1_sd : size r1 = size d by rewrite size_scale ?invr_eq0. +have sr1 : (1 < size r1)%N. + by rewrite ltn_neqAle eq_sym lt0n size_poly_eq0 monic_neq0 ?andbT ?eq_sr1_sd. +have sr2 : (1 < size r2)%N. + rewrite size_divp ?size_dvdp ?monic_neq0 //. + rewrite ltn_subRL addn1 prednK ?(leq_trans _ sr1) // eq_sr1_sd. + rewrite ltn_neqAle dvdp_leq ?monic_neq0 ?andbT ?dvdp_size_eqp ?dvdp_gcdl //. + by apply: contra ndvd_pq=> /eqp_dvdl <-; rewrite dvdp_gcdr. +move: (px0); rewrite eq_p_r2r1=> r2r1x_eq0. +have : (r2.[x] * r1.[x] == 0) by rewrite -horner_crealM. +case/poly_mul_creal_eq0=> [r2x_eq0|r1x_eq0]. + exists r2; rewrite ?monic_r2 ?andbT // mulrC. + by rewrite -ltn_divpl ?divpp ?monic_neq0 // size_poly1. +exists r1; rewrite ?monic_r1 ?andbT //. +by rewrite -ltn_divpl ?divpp ?monic_neq0 // size_poly1. +Qed. + +Lemma root_cst_creal (x : F) : ('X - x%:P).[cst_creal x] == 0. +Proof. +apply: eq_crealP; exists_big_modulus m F. + by move=> e i e_gt0 hi; rewrite /= subr0 !hornerE subrr normr0. +by close. +Qed. + +Lemma has_root_creal_size_gt1 (x : creal) (p : {poly F}) : + (p != 0)%B -> p.[x] == 0 -> (1 < size p)%N. +Proof. +move=> p_neq0 rootpa. +rewrite ltnNge leq_eqVlt ltnS leqn0 size_poly_eq0 (negPf p_neq0) orbF. +apply/negP=> /size_poly1P [c c_neq0 eq_pc]; apply: rootpa. +by rewrite eq_pc horner_cst_creal; apply/neq_creal_cst. +Qed. + +Definition bound_poly_bound (z : creal) (q : {poly {poly F}}) (a r : F) i := + (1 + \sum_(j < sizeY q) + `|(norm_poly2 q).[(ubound z)%:P]^`N(i.+1)`_j| * (`|a| + `|r|) ^+ j). + +Lemma bound_poly_boundP (z : creal) i (q : {poly {poly F}}) (a r : F) j : + poly_bound q.[(z i)%:P]^`N(j.+1) a r <= bound_poly_bound z q a r j. +Proof. +rewrite /poly_bound. +pose f q (k : nat) := `|q^`N(j.+1)`_k| * (`|a| + `|r|) ^+ k. +rewrite ler_add //=. +rewrite (big_ord_widen (sizeY q) (f q.[(z i)%:P])); last first. + rewrite size_nderivn leq_subLR (leq_trans (max_size_evalC _ _)) //. + by rewrite leq_addl. +rewrite big_mkcond /= ler_sum // /f => k _. +case: ifP=> _; last by rewrite mulr_ge0 ?exprn_ge0 ?addr_ge0 ?normr_ge0. +rewrite ler_wpmul2r ?exprn_ge0 ?addr_ge0 ?normr_ge0 //. +rewrite !horner_coef. +rewrite !(@big_morph _ _ (fun p => p^`N(j.+1)) 0 +%R); + do ?[by rewrite raddf0|by move=> x y /=; rewrite raddfD]. +rewrite !coef_sum. +rewrite (ler_trans (ler_norm_sum _ _ _)) //. +rewrite ger0_norm; last first. + rewrite sumr_ge0=> //= l _. + rewrite coef_nderivn mulrn_wge0 ?natr_ge0 //. + rewrite -polyC_exp coefMC coef_norm_poly2 mulr_ge0 ?normr_ge0 //. + by rewrite exprn_ge0 ?ltrW ?ubound_gt0. +rewrite size_norm_poly2 ler_sum //= => l _. +rewrite !{1}coef_nderivn normrMn ler_pmuln2r ?bin_gt0 ?leq_addr //. +rewrite -!polyC_exp !coefMC coef_norm_poly2 normrM ler_wpmul2l ?normr_ge0 //. +rewrite normrX; case: (val l)=> // {l} l. +by rewrite ler_pexpn2r -?topredE //= ?uboundP ?ltrW ?ubound_gt0. +Qed. + +Lemma bound_poly_bound_ge0 z q a r i : 0 <= bound_poly_bound z q a r i. +Proof. +by rewrite (ler_trans _ (bound_poly_boundP _ 0%N _ _ _ _)) ?poly_bound_ge0. +Qed. + +Definition bound_poly_accr_bound (z : creal) (q : {poly {poly F}}) (a r : F) := + maxr 1 (2%:R * r) ^+ (sizeY q).-1 * + (1 + \sum_(i < (sizeY q).-1) bound_poly_bound z q a r i). + +Lemma bound_poly_accr_boundP (z : creal) i (q : {poly {poly F}}) (a r : F) : + poly_accr_bound q.[(z i)%:P] a r <= bound_poly_accr_bound z q a r. +Proof. +rewrite /poly_accr_bound /bound_poly_accr_bound /=. +set ui := _ ^+ _; set u := _ ^+ _; set vi := 1 + _. +rewrite (@ler_trans _ (u * vi)) //. + rewrite ler_wpmul2r //. + by rewrite addr_ge0 ?ler01 // sumr_ge0 //= => j _; rewrite poly_bound_ge0. + rewrite /ui /u; case: maxrP; first by rewrite !expr1n. + move=> r2_gt1; rewrite ler_eexpn2l //. + rewrite -subn1 leq_subLR add1n (leq_trans _ (leqSpred _)) //. + by rewrite max_size_evalC. +rewrite ler_wpmul2l ?exprn_ge0 ?ler_maxr ?ler01 // ler_add //. +pose f j := poly_bound q.[(z i)%:P]^`N(j.+1) a r. +rewrite (big_ord_widen (sizeY q).-1 f); last first. + rewrite -subn1 leq_subLR add1n (leq_trans _ (leqSpred _)) //. + by rewrite max_size_evalC. +rewrite big_mkcond /= ler_sum // /f => k _. +by case: ifP=> _; rewrite ?bound_poly_bound_ge0 ?bound_poly_boundP. +Qed. + +Lemma bound_poly_accr_bound_gt0 (z : creal) (q : {poly {poly F}}) (a r : F) : + 0 < bound_poly_accr_bound z q a r. +Proof. +rewrite (ltr_le_trans _ (bound_poly_accr_boundP _ 0%N _ _ _)) //. +by rewrite poly_accr_bound_gt0. +Qed. + +Lemma horner2_crealP (p : {poly {poly F}}) (x y : creal) : + creal_axiom (fun i => p.[x i, y i]). +Proof. +set a := x (cauchymod x 1). +exists_big_modulus m F. + move=> e i j e_gt0 hi hj; rewrite (@split_dist_add p.[x i, y j]) //. + rewrite (ler_lt_trans (@poly_accr_bound1P _ _ 0 (ubound y) _ _ _ _)) //; + do ?by rewrite ?subr0 ?uboundP. + rewrite (@ler_lt_trans _ (`|y i - y j| + * bound_poly_accr_bound x p 0 (ubound y))) //. + by rewrite ler_wpmul2l ?normr_ge0 // bound_poly_accr_boundP. + rewrite -ltr_pdivl_mulr ?bound_poly_accr_bound_gt0 //. + by rewrite cauchymodP // !pmulr_rgt0 ?gtr0E ?bound_poly_accr_bound_gt0. + rewrite -[p]swapXYK ![(swapXY (swapXY _)).[_, _]]horner2_swapXY. + rewrite (ler_lt_trans (@poly_accr_bound1P _ _ 0 (ubound x) _ _ _ _)) //; + do ?by rewrite ?subr0 ?uboundP. + rewrite (@ler_lt_trans _ (`|x i - x j| + * bound_poly_accr_bound y (swapXY p) 0 (ubound x))) //. + by rewrite ler_wpmul2l ?normr_ge0 // bound_poly_accr_boundP. + rewrite -ltr_pdivl_mulr ?bound_poly_accr_bound_gt0 //. + by rewrite cauchymodP // !pmulr_rgt0 ?gtr0E ?bound_poly_accr_bound_gt0. +by close. +Qed. + +Definition horner2_creal (p : {poly {poly F}}) (x y : creal) := + CReal (horner2_crealP p x y). +Notation "p .[ x , y ]" := (horner2_creal p x y) + (at level 2, left associativity) : creal_scope. + +Lemma root_monic_from_neq0 (p : {poly F}) (x : creal) : + p.[x] == 0 -> ((lead_coef p) ^-1 *: p).[x] == 0. +Proof. by rewrite -mul_polyC horner_crealM; move->; rewrite mul_creal0. Qed. + +Lemma root_sub_annihilant_creal (x y : creal) (p q : {poly F}) : + (p != 0)%B -> (q != 0)%B -> p.[x] == 0 -> q.[y] == 0 -> + (sub_annihilant p q).[x - y] == 0. +Proof. +move=> p_neq0 q_neq0 px_eq0 qy_eq0. +have [||[u v] /= [hu hv] hpq] := @sub_annihilant_in_ideal _ p q. ++ by rewrite (@has_root_creal_size_gt1 x). ++ by rewrite (@has_root_creal_size_gt1 y). +apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi /=; rewrite subr0. + rewrite (hpq (y i)) addrCA subrr addr0 split_norm_add // normrM. + rewrite (@ler_lt_trans _ ((ubound u.[y, x - y]) * `|p.[x i]|)) //. + by rewrite ler_wpmul2r ?normr_ge0 // (uboundP u.[y, x - y] i). + rewrite -ltr_pdivl_mull ?ubound_gt0 //. + by rewrite (@eq0_modP _ px_eq0) // !pmulr_rgt0 ?gtr0E ?ubound_gt0. + rewrite (@ler_lt_trans _ ((ubound v.[y, x - y]) * `|q.[y i]|)) //. + by rewrite ler_wpmul2r ?normr_ge0 // (uboundP v.[y, x - y] i). + rewrite -ltr_pdivl_mull ?ubound_gt0 //. + by rewrite (@eq0_modP _ qy_eq0) // !pmulr_rgt0 ?gtr0E ?ubound_gt0. +by close. +Qed. + +Lemma root_div_annihilant_creal (x y : creal) (p q : {poly F}) (y_neq0 : y != 0) : + (p != 0)%B -> (q != 0)%B -> p.[x] == 0 -> q.[y] == 0 -> + (div_annihilant p q).[(x / y_neq0)%CR] == 0. +Proof. +move=> p_neq0 q_neq0 px_eq0 qy_eq0. +have [||[u v] /= [hu hv] hpq] := @div_annihilant_in_ideal _ p q. ++ by rewrite (@has_root_creal_size_gt1 x). ++ by rewrite (@has_root_creal_size_gt1 y). +apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi /=; rewrite subr0. + rewrite (hpq (y i)) mulrCA divff ?mulr1; last first. + by rewrite -normr_gt0 (ltr_le_trans _ (lbound0_of y_neq0)) ?lbound_gt0. + rewrite split_norm_add // normrM. + rewrite (@ler_lt_trans _ ((ubound u.[y, x / y_neq0]) * `|p.[x i]|)) //. + by rewrite ler_wpmul2r ?normr_ge0 // (uboundP u.[y, x / y_neq0] i). + rewrite -ltr_pdivl_mull ?ubound_gt0 //. + by rewrite (@eq0_modP _ px_eq0) // !pmulr_rgt0 ?gtr0E ?ubound_gt0. + rewrite (@ler_lt_trans _ ((ubound v.[y, x / y_neq0]) * `|q.[y i]|)) //. + by rewrite ler_wpmul2r ?normr_ge0 // (uboundP v.[y, x / y_neq0] i). + rewrite -ltr_pdivl_mull ?ubound_gt0 //. + by rewrite (@eq0_modP _ qy_eq0) // !pmulr_rgt0 ?gtr0E ?ubound_gt0. +by close. +Qed. + +Definition exp_creal x n := (iterop n *%CR x 1%:CR). +Notation "x ^+ n" := (exp_creal x n) : creal_scope. + +Add Morphism exp_creal with + signature eq_creal ==> (@eq _) ==> eq_creal as exp_creal_morph. +Proof. +move=> x y eq_xy [//|n]; rewrite /exp_creal !iteropS. +by elim: n=> //= n ->; rewrite eq_xy. +Qed. +Global Existing Instance exp_creal_morph_Proper. + +Lemma horner_coef_creal p x : + p.[x] == \big[+%CR/0%:CR]_(i < size p) ((p`_i)%:CR * (x ^+ i))%CR. +Proof. +apply: eq_crealP; exists m0=> e n e_gt0 hn /=; rewrite horner_coef. +rewrite (@big_morph _ _ (fun u : creal => u n) 0%R +%R) //. +rewrite -sumrB /= big1 ?normr0=> //= i _. +apply/eqP; rewrite subr_eq0; apply/eqP; congr (_ * _). +case: val=> {i} // i; rewrite exprS /exp_creal iteropS. +by elim: i=> [|i ihi]; rewrite ?expr0 ?mulr1 //= exprS ihi. +Qed. + +End CauchyReals. + +Notation "x == y" := (eq_creal x y) : creal_scope. +Notation "!=%CR" := neq_creal : creal_scope. +Notation "x != y" := (neq_creal x y) : creal_scope. + +Notation "x %:CR" := (cst_creal x) + (at level 2, left associativity, format "x %:CR") : creal_scope. +Notation "0" := (0 %:CR)%CR : creal_scope. + +Notation "<%CR" := lt_creal : creal_scope. +Notation "x < y" := (lt_creal x y) : creal_scope. + +Notation "<=%CR" := le_creal : creal_scope. +Notation "x <= y" := (le_creal x y) : creal_scope. + +Notation "-%CR" := opp_creal : creal_scope. +Notation "- x" := (opp_creal x) : creal_scope. + +Notation "+%CR" := add_creal : creal_scope. +Notation "x + y" := (add_creal x y) : creal_scope. +Notation "x - y" := (x + - y)%CR : creal_scope. + +Notation "*%CR" := mul_creal : creal_scope. +Notation "x * y" := (mul_creal x y) : creal_scope. + +Notation "x_neq0 ^-1" := (inv_creal x_neq0) : creal_scope. +Notation "x / y_neq0" := (x * (y_neq0 ^-1))%CR : creal_scope. +Notation "p .[ x ]" := (horner_creal p x) : creal_scope. +Notation "p .[ x , y ]" := (horner2_creal p x y) + (at level 2, left associativity) : creal_scope. +Notation "x ^+ n" := (exp_creal x n) : creal_scope. + +Notation "`| x |" := (norm_creal x) : creal_scope. + +Hint Resolve eq_creal_refl. +Hint Resolve le_creal_refl. + +Notation lbound_of p := (@lboundP _ _ _ p _ _ _). +Notation lbound0_of p := (@lbound0P _ _ p _ _ _). +Notation diff_of p := (@diffP _ _ _ p _ _ _). +Notation diff0_of p := (@diff0P _ _ p _ _ _). + +Notation "{ 'asympt' e : i / P }" := (asympt1 (fun e i => P)) + (at level 0, e ident, i ident, format "{ 'asympt' e : i / P }") : type_scope. +Notation "{ 'asympt' e : i j / P }" := (asympt2 (fun e i j => P)) + (at level 0, e ident, i ident, j ident, format "{ 'asympt' e : i j / P }") : type_scope. diff --git a/mathcomp/real_closed/complex.v b/mathcomp/real_closed/complex.v new file mode 100644 index 0000000..1c26a9d --- /dev/null +++ b/mathcomp/real_closed/complex.v @@ -0,0 +1,1252 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import bigop ssralg ssrint div ssrnum rat poly closed_field polyrcf. +Require Import matrix mxalgebra tuple mxpoly zmodp binomial realalg. + +(**********************************************************************) +(* This files defines the extension R[i] of a real field R, *) +(* and provide it a structure of numeric field with a norm operator. *) +(* When R is a real closed field, it also provides a structure of *) +(* algebraically closed field for R[i], using a proof by Derksen *) +(* (cf comments below, thanks to Pierre Lairez for finding the paper) *) +(**********************************************************************) + +Import GRing.Theory Num.Theory. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. + +Reserved Notation "x +i* y" + (at level 40, left associativity, format "x +i* y"). +Reserved Notation "x -i* y" + (at level 40, left associativity, format "x -i* y"). +Reserved Notation "R [i]" + (at level 2, left associativity, format "R [i]"). + +Local Notation sgr := Num.sg. +Local Notation sqrtr := Num.sqrt. + +CoInductive complex (R : Type) : Type := Complex { Re : R; Im : R }. + +Definition real_complex_def (F : ringType) (phF : phant F) (x : F) := + Complex x 0. +Notation real_complex F := (@real_complex_def _ (Phant F)). +Notation "x %:C" := (real_complex _ x) + (at level 2, left associativity, format "x %:C") : ring_scope. +Notation "x +i* y" := (Complex x y) : ring_scope. +Notation "x -i* y" := (Complex x (- y)) : ring_scope. +Notation "x *i " := (Complex 0 x) (at level 8, format "x *i") : ring_scope. +Notation "''i'" := (Complex 0 1) : ring_scope. +Notation "R [i]" := (complex R) + (at level 2, left associativity, format "R [i]"). + +Module ComplexEqChoice. +Section ComplexEqChoice. + +Variable R : Type. + +Definition sqR_of_complex (x : R[i]) := let: a +i* b := x in [::a; b]. +Definition complex_of_sqR (x : seq R) := + if x is [:: a; b] then Some (a +i* b) else None. + +Lemma complex_of_sqRK : pcancel sqR_of_complex complex_of_sqR. +Proof. by case. Qed. + +End ComplexEqChoice. +End ComplexEqChoice. + +Definition complex_eqMixin (R : eqType) := + PcanEqMixin (@ComplexEqChoice.complex_of_sqRK R). +Definition complex_choiceMixin (R : choiceType) := + PcanChoiceMixin (@ComplexEqChoice.complex_of_sqRK R). +Definition complex_countMixin (R : countType) := + PcanCountMixin (@ComplexEqChoice.complex_of_sqRK R). + +Canonical Structure complex_eqType (R : eqType) := + EqType R[i] (complex_eqMixin R). +Canonical Structure complex_choiceType (R : choiceType) := + ChoiceType R[i] (complex_choiceMixin R). +Canonical Structure complex_countType (R : countType) := + CountType R[i] (complex_countMixin R). + +Lemma eq_complex : forall (R : eqType) (x y : complex R), + (x == y) = (Re x == Re y) && (Im x == Im y). +Proof. +move=> R [a b] [c d] /=. +apply/eqP/andP; first by move=> [-> ->]; split. +by case; move/eqP->; move/eqP->. +Qed. + +Lemma complexr0 : forall (R : ringType) (x : R), x +i* 0 = x%:C. Proof. by []. Qed. + +Module ComplexField. +Section ComplexField. + +Variable R : rcfType. +Local Notation C := R[i]. +Local Notation C0 := ((0 : R)%:C). +Local Notation C1 := ((1 : R)%:C). + +Definition addc (x y : R[i]) := let: a +i* b := x in let: c +i* d := y in + (a + c) +i* (b + d). +Definition oppc (x : R[i]) := let: a +i* b := x in (- a) +i* (- b). + +Lemma addcC : commutative addc. +Proof. by move=> [a b] [c d] /=; congr (_ +i* _); rewrite addrC. Qed. +Lemma addcA : associative addc. +Proof. by move=> [a b] [c d] [e f] /=; rewrite !addrA. Qed. + +Lemma add0c : left_id C0 addc. +Proof. by move=> [a b] /=; rewrite !add0r. Qed. + +Lemma addNc : left_inverse C0 oppc addc. +Proof. by move=> [a b] /=; rewrite !addNr. Qed. + +Definition complex_ZmodMixin := ZmodMixin addcA addcC add0c addNc. +Canonical Structure complex_ZmodType := ZmodType R[i] complex_ZmodMixin. + +Definition mulc (x y : R[i]) := let: a +i* b := x in let: c +i* d := y in + ((a * c) - (b * d)) +i* ((a * d) + (b * c)). + +Lemma mulcC : commutative mulc. +Proof. +move=> [a b] [c d] /=. +by rewrite [c * b + _]addrC ![_ * c]mulrC ![_ * d]mulrC. +Qed. + +Lemma mulcA : associative mulc. +Proof. +move=> [a b] [c d] [e f] /=. +rewrite !mulrDr !mulrDl !mulrN !mulNr !mulrA !opprD -!addrA. +by congr ((_ + _) +i* (_ + _)); rewrite !addrA addrAC; + congr (_ + _); rewrite addrC. +Qed. + +Definition invc (x : R[i]) := let: a +i* b := x in let n2 := (a ^+ 2 + b ^+ 2) in + (a / n2) -i* (b / n2). + +Lemma mul1c : left_id C1 mulc. +Proof. by move=> [a b] /=; rewrite !mul1r !mul0r subr0 addr0. Qed. + +Lemma mulc_addl : left_distributive mulc addc. +Proof. +move=> [a b] [c d] [e f] /=; rewrite !mulrDl !opprD -!addrA. +by congr ((_ + _) +i* (_ + _)); rewrite addrCA. +Qed. + +Lemma nonzero1c : C1 != C0. Proof. by rewrite eq_complex /= oner_eq0. Qed. + +Definition complex_comRingMixin := + ComRingMixin mulcA mulcC mul1c mulc_addl nonzero1c. +Canonical Structure complex_Ring := + Eval hnf in RingType R[i] complex_comRingMixin. +Canonical Structure complex_comRing := Eval hnf in ComRingType R[i] mulcC. + +Lemma mulVc : forall x, x != C0 -> mulc (invc x) x = C1. +Proof. +move=> [a b]; rewrite eq_complex => /= hab; rewrite !mulNr opprK. +rewrite ![_ / _ * _]mulrAC [b * a]mulrC subrr complexr0 -mulrDl mulfV //. +by rewrite paddr_eq0 -!expr2 ?expf_eq0 ?sqr_ge0. +Qed. + +Lemma invc0 : invc C0 = C0. Proof. by rewrite /= !mul0r oppr0. Qed. + +Definition ComplexFieldUnitMixin := FieldUnitMixin mulVc invc0. +Canonical Structure complex_unitRing := + Eval hnf in UnitRingType C ComplexFieldUnitMixin. +Canonical Structure complex_comUnitRing := + Eval hnf in [comUnitRingType of R[i]]. + +Lemma field_axiom : GRing.Field.mixin_of complex_unitRing. +Proof. by []. Qed. + +Definition ComplexFieldIdomainMixin := (FieldIdomainMixin field_axiom). +Canonical Structure complex_iDomain := + Eval hnf in IdomainType R[i] (FieldIdomainMixin field_axiom). +Canonical Structure complex_fieldMixin := FieldType R[i] field_axiom. + +Ltac simpc := do ? + [ rewrite -[(_ +i* _) - (_ +i* _)]/(_ +i* _) + | rewrite -[(_ +i* _) + (_ +i* _)]/(_ +i* _) + | rewrite -[(_ +i* _) * (_ +i* _)]/(_ +i* _)]. + +Lemma real_complex_is_rmorphism : rmorphism (real_complex R). +Proof. +split; [|split=> //] => a b /=; simpc; first by rewrite subrr. +by rewrite !mulr0 !mul0r addr0 subr0. +Qed. + +Canonical Structure real_complex_rmorphism := + RMorphism real_complex_is_rmorphism. +Canonical Structure real_complex_additive := + Additive real_complex_is_rmorphism. + +Lemma Re_is_additive : additive (@Re R). +Proof. by case=> a1 b1; case=> a2 b2. Qed. + +Canonical Structure Re_additive := Additive Re_is_additive. + +Lemma Im_is_additive : additive (@Im R). +Proof. by case=> a1 b1; case=> a2 b2. Qed. + +Canonical Structure Im_additive := Additive Im_is_additive. + +Definition lec (x y : R[i]) := + let: a +i* b := x in let: c +i* d := y in + (d == b) && (a <= c). + +Definition ltc (x y : R[i]) := + let: a +i* b := x in let: c +i* d := y in + (d == b) && (a < c). + +Definition normc (x : R[i]) : R := + let: a +i* b := x in sqrtr (a ^+ 2 + b ^+ 2). + +Notation normC x := (normc x)%:C. + +Lemma ltc0_add : forall x y, ltc 0 x -> ltc 0 y -> ltc 0 (x + y). +Proof. +move=> [a b] [c d] /= /andP [/eqP-> ha] /andP [/eqP-> hc]. +by rewrite addr0 eqxx addr_gt0. +Qed. + +Lemma eq0_normc x : normc x = 0 -> x = 0. +Proof. +case: x => a b /= /eqP; rewrite sqrtr_eq0 ler_eqVlt => /orP [|]; last first. + by rewrite ltrNge addr_ge0 ?sqr_ge0. +by rewrite paddr_eq0 ?sqr_ge0 ?expf_eq0 //= => /andP[/eqP -> /eqP ->]. +Qed. + +Lemma eq0_normC x : normC x = 0 -> x = 0. Proof. by case=> /eq0_normc. Qed. + +Lemma ge0_lec_total x y : lec 0 x -> lec 0 y -> lec x y || lec y x. +Proof. +move: x y => [a b] [c d] /= /andP[/eqP -> a_ge0] /andP[/eqP -> c_ge0]. +by rewrite eqxx ler_total. +Qed. + +(* :TODO: put in ssralg ? *) +Lemma exprM (a b : R) : (a * b) ^+ 2 = a ^+ 2 * b ^+ 2. +Proof. by rewrite mulrACA. Qed. + +Lemma normcM x y : normc (x * y) = normc x * normc y. +Proof. +move: x y => [a b] [c d] /=; rewrite -sqrtrM ?addr_ge0 ?sqr_ge0 //. +rewrite sqrrB sqrrD mulrDl !mulrDr -!exprM. +rewrite mulrAC [b * d]mulrC !mulrA. +suff -> : forall (u v w z t : R), (u - v + w) + (z + v + t) = u + w + (z + t). + by rewrite addrAC !addrA. +by move=> u v w z t; rewrite [_ - _ + _]addrAC [z + v]addrC !addrA addrNK. +Qed. + +Lemma normCM x y : normC (x * y) = normC x * normC y. +Proof. by rewrite -rmorphM normcM. Qed. + +Lemma subc_ge0 x y : lec 0 (y - x) = lec x y. +Proof. by move: x y => [a b] [c d] /=; simpc; rewrite subr_ge0 subr_eq0. Qed. + +Lemma lec_def x y : lec x y = (normC (y - x) == y - x). +Proof. +rewrite -subc_ge0; move: (_ - _) => [a b]; rewrite eq_complex /= eq_sym. +have [<- /=|_] := altP eqP; last by rewrite andbF. +by rewrite [0 ^+ _]mul0r addr0 andbT sqrtr_sqr ger0_def. +Qed. + +Lemma ltc_def x y : ltc x y = (y != x) && lec x y. +Proof. +move: x y => [a b] [c d] /=; simpc; rewrite eq_complex /=. +by have [] := altP eqP; rewrite ?(andbF, andbT) //= ltr_def. +Qed. + +Lemma lec_normD x y : lec (normC (x + y)) (normC x + normC y). +Proof. +move: x y => [a b] [c d] /=; simpc; rewrite addr0 eqxx /=. +rewrite -(@ler_pexpn2r _ 2) -?topredE /= ?(ler_paddr, sqrtr_ge0) //. +rewrite [X in _ <= X] sqrrD ?sqr_sqrtr; + do ?by rewrite ?(ler_paddr, sqrtr_ge0, sqr_ge0, mulr_ge0) //. +rewrite -addrA addrCA (monoRL (addrNK _) (ler_add2r _)) !sqrrD. +set u := _ *+ 2; set v := _ *+ 2. +rewrite [a ^+ _ + _ + _]addrAC [b ^+ _ + _ + _]addrAC -addrA. +rewrite [u + _] addrC [X in _ - X]addrAC [b ^+ _ + _]addrC. +rewrite [u]lock [v]lock !addrA; set x := (a ^+ 2 + _ + _ + _). +rewrite -addrA addrC addKr -!lock addrC. +have [huv|] := ger0P (u + v); last first. + by move=> /ltrW /ler_trans -> //; rewrite pmulrn_lge0 // mulr_ge0 ?sqrtr_ge0. +rewrite -(@ler_pexpn2r _ 2) -?topredE //=; last first. + by rewrite ?(pmulrn_lge0, mulr_ge0, sqrtr_ge0) //. +rewrite -mulr_natl !exprM !sqr_sqrtr ?(ler_paddr, sqr_ge0) //. +rewrite -mulrnDl -mulr_natl !exprM ler_pmul2l ?exprn_gt0 ?ltr0n //. +rewrite sqrrD mulrDl !mulrDr -!exprM addrAC. +rewrite [_ + (b * d) ^+ 2]addrC [X in _ <= X]addrAC -!addrA !ler_add2l. +rewrite mulrAC mulrA -mulrA mulrACA mulrC. +by rewrite -subr_ge0 addrAC -sqrrB sqr_ge0. +Qed. + +Definition complex_POrderedMixin := NumMixin lec_normD ltc0_add eq0_normC + ge0_lec_total normCM lec_def ltc_def. +Canonical Structure complex_numDomainType := + NumDomainType R[i] complex_POrderedMixin. + +End ComplexField. +End ComplexField. + +Canonical complex_ZmodType (R : rcfType) := + ZmodType R[i] (ComplexField.complex_ZmodMixin R). +Canonical complex_Ring (R : rcfType) := + Eval hnf in RingType R[i] (ComplexField.complex_comRingMixin R). +Canonical complex_comRing (R : rcfType) := + Eval hnf in ComRingType R[i] (@ComplexField.mulcC R). +Canonical complex_unitRing (R : rcfType) := + Eval hnf in UnitRingType R[i] (ComplexField.ComplexFieldUnitMixin R). +Canonical complex_comUnitRing (R : rcfType) := + Eval hnf in [comUnitRingType of R[i]]. +Canonical complex_iDomain (R : rcfType) := + Eval hnf in IdomainType R[i] (FieldIdomainMixin (@ComplexField.field_axiom R)). +Canonical complex_fieldType (R : rcfType) := + FieldType R[i] (@ComplexField.field_axiom R). +Canonical complex_numDomainType (R : rcfType) := + NumDomainType R[i] (ComplexField.complex_POrderedMixin R). +Canonical complex_numFieldType (R : rcfType) := + [numFieldType of complex R]. + +Canonical ComplexField.real_complex_rmorphism. +Canonical ComplexField.real_complex_additive. +Canonical ComplexField.Re_additive. +Canonical ComplexField.Im_additive. + +Definition conjc {R : ringType} (x : R[i]) := let: a +i* b := x in a -i* b. +Notation "x ^*" := (conjc x) (at level 2, format "x ^*"). + +Ltac simpc := do ? + [ rewrite -[(_ +i* _) - (_ +i* _)]/(_ +i* _) + | rewrite -[(_ +i* _) + (_ +i* _)]/(_ +i* _) + | rewrite -[(_ +i* _) * (_ +i* _)]/(_ +i* _) + | rewrite -[(_ +i* _) <= (_ +i* _)]/((_ == _) && (_ <= _)) + | rewrite -[(_ +i* _) < (_ +i* _)]/((_ == _) && (_ < _)) + | rewrite -[`|_ +i* _|]/(sqrtr (_ + _))%:C + | rewrite (mulrNN, mulrN, mulNr, opprB, opprD, mulr0, mul0r, + subr0, sub0r, addr0, add0r, mulr1, mul1r, subrr, opprK, oppr0, + eqxx) ]. + + +Section ComplexTheory. + +Variable R : rcfType. + +Lemma ReiNIm : forall x : R[i], Re (x * 'i) = - Im x. +Proof. by case=> a b; simpc. Qed. + +Lemma ImiRe : forall x : R[i], Im (x * 'i) = Re x. +Proof. by case=> a b; simpc. Qed. + +Lemma complexE x : x = (Re x)%:C + 'i * (Im x)%:C :> R[i]. +Proof. by case: x => *; simpc. Qed. + +Lemma real_complexE x : x%:C = x +i* 0 :> R[i]. Proof. done. Qed. + +Lemma sqr_i : 'i ^+ 2 = -1 :> R[i]. +Proof. by rewrite exprS; simpc; rewrite -real_complexE rmorphN. Qed. + +Lemma complexI : injective (real_complex R). Proof. by move=> x y []. Qed. + +Lemma ler0c (x : R) : (0 <= x%:C) = (0 <= x). Proof. by simpc. Qed. + +Lemma lecE : forall x y : R[i], (x <= y) = (Im y == Im x) && (Re x <= Re y). +Proof. by move=> [a b] [c d]. Qed. + +Lemma ltcE : forall x y : R[i], (x < y) = (Im y == Im x) && (Re x < Re y). +Proof. by move=> [a b] [c d]. Qed. + +Lemma lecR : forall x y : R, (x%:C <= y%:C) = (x <= y). +Proof. by move=> x y; simpc. Qed. + +Lemma ltcR : forall x y : R, (x%:C < y%:C) = (x < y). +Proof. by move=> x y; simpc. Qed. + +Lemma conjc_is_rmorphism : rmorphism (@conjc R). +Proof. +split=> [[a b] [c d]|] /=; first by simpc; rewrite [d - _]addrC. +by split=> [[a b] [c d]|] /=; simpc. +Qed. + +Canonical conjc_rmorphism := RMorphism conjc_is_rmorphism. +Canonical conjc_additive := Additive conjc_is_rmorphism. + +Lemma conjcK : involutive (@conjc R). +Proof. by move=> [a b] /=; rewrite opprK. Qed. + +Lemma mulcJ_ge0 (x : R[i]) : 0 <= x * x ^*. +Proof. +by move: x=> [a b]; simpc; rewrite mulrC addNr eqxx addr_ge0 ?sqr_ge0. +Qed. + +Lemma conjc_real (x : R) : x%:C^* = x%:C. +Proof. by rewrite /= oppr0. Qed. + +Lemma ReJ_add (x : R[i]) : (Re x)%:C = (x + x^*) / 2%:R. +Proof. +case: x => a b; simpc; rewrite [0 ^+ 2]mul0r addr0 /=. +rewrite -!mulr2n -mulr_natr -mulrA [_ * (_ / _)]mulrA. +by rewrite divff ?mulr1 // -natrM pnatr_eq0. +Qed. + +Lemma ImJ_sub (x : R[i]) : (Im x)%:C = (x^* - x) / 2%:R * 'i. +Proof. +case: x => a b; simpc; rewrite [0 ^+ 2]mul0r addr0 /=. +rewrite -!mulr2n -mulr_natr -mulrA [_ * (_ / _)]mulrA. +by rewrite divff ?mulr1 ?opprK // -natrM pnatr_eq0. +Qed. + +Lemma ger0_Im (x : R[i]) : 0 <= x -> Im x = 0. +Proof. by move: x=> [a b] /=; simpc => /andP [/eqP]. Qed. + +(* Todo : extend theory of : *) +(* - signed exponents *) + +Lemma conj_ge0 : forall x : R[i], (0 <= x ^*) = (0 <= x). +Proof. by move=> [a b] /=; simpc; rewrite oppr_eq0. Qed. + +Lemma conjc_nat : forall n, (n%:R : R[i])^* = n%:R. +Proof. exact: rmorph_nat. Qed. + +Lemma conjc0 : (0 : R[i]) ^* = 0. +Proof. exact: (conjc_nat 0). Qed. + +Lemma conjc1 : (1 : R[i]) ^* = 1. +Proof. exact: (conjc_nat 1). Qed. + +Lemma conjc_eq0 : forall x : R[i], (x ^* == 0) = (x == 0). +Proof. by move=> [a b]; rewrite !eq_complex /= eqr_oppLR oppr0. Qed. + +Lemma conjc_inv: forall x : R[i], (x^-1)^* = (x^* )^-1. +Proof. exact: fmorphV. Qed. + +Lemma complex_root_conj (p : {poly R[i]}) (x : R[i]) : + root (map_poly conjc p) x = root p x^*. +Proof. by rewrite /root -{1}[x]conjcK horner_map /= conjc_eq0. Qed. + +Lemma complex_algebraic_trans (T : comRingType) (toR : {rmorphism T -> R}) : + integralRange toR -> integralRange (real_complex R \o toR). +Proof. +set f := _ \o _ => R_integral [a b]. +have integral_real x : integralOver f (x%:C) by apply: integral_rmorph. +rewrite [_ +i* _]complexE. +apply: integral_add => //; apply: integral_mul => //=. +exists ('X^2 + 1). + by rewrite monicE lead_coefDl ?size_polyXn ?size_poly1 ?lead_coefXn. +by rewrite rmorphD rmorph1 /= ?map_polyXn rootE !hornerE -expr2 sqr_i addNr. +Qed. + +Lemma normc_def (z : R[i]) : `|z| = (sqrtr ((Re z)^+2 + (Im z)^+2))%:C. +Proof. by case: z. Qed. + +Lemma add_Re2_Im2 (z : R[i]) : ((Re z)^+2 + (Im z)^+2)%:C = `|z|^+2. +Proof. by rewrite normc_def -rmorphX sqr_sqrtr ?addr_ge0 ?sqr_ge0. Qed. + +Lemma addcJ (z : R[i]) : z + z^* = 2%:R * (Re z)%:C. +Proof. by rewrite ReJ_add mulrC mulfVK ?pnatr_eq0. Qed. + +Lemma subcJ (z : R[i]) : z - z^* = 2%:R * (Im z)%:C * 'i. +Proof. +rewrite ImJ_sub mulrCA mulrA mulfVK ?pnatr_eq0 //. +by rewrite -mulrA ['i * _]sqr_i mulrN1 opprB. +Qed. + +End ComplexTheory. + +(* Section RcfDef. *) + +(* Variable R : realFieldType. *) +(* Notation C := (complex R). *) + +(* Definition rcf_odd := forall (p : {poly R}), *) +(* ~~odd (size p) -> {x | p.[x] = 0}. *) +(* Definition rcf_square := forall x : R, *) +(* {y | (0 <= y) && if 0 <= x then (y ^ 2 == x) else y == 0}. *) + +(* Lemma rcf_odd_sqr_from_ivt : rcf_axiom R -> rcf_odd * rcf_square. *) +(* Proof. *) +(* move=> ivt. *) +(* split. *) +(* move=> p sp. *) +(* move: (ivt p). *) +(* admit. *) +(* move=> x. *) +(* case: (boolP (0 <= x)) (@ivt ('X^2 - x%:P) 0 (1 + x))=> px; last first. *) +(* by move=> _; exists 0; rewrite lerr eqxx. *) +(* case. *) +(* * by rewrite ler_paddr ?ler01. *) +(* * rewrite !horner_lin oppr_le0 px /=. *) +(* rewrite subr_ge0 (@ler_trans _ (1 + x)) //. *) +(* by rewrite ler_paddl ?ler01 ?lerr. *) +(* by rewrite ler_pemulr // addrC -subr_ge0 ?addrK // subr0 ler_paddl ?ler01. *) +(* * move=> y hy; rewrite /root !horner_lin; move/eqP. *) +(* move/(canRL (@addrNK _ _)); rewrite add0r=> <-. *) +(* by exists y; case/andP: hy=> -> _; rewrite eqxx. *) +(* Qed. *) + +(* Lemma ivt_from_closed : GRing.ClosedField.axiom [ringType of C] -> rcf_axiom R. *) +(* Proof. *) +(* rewrite /GRing.ClosedField.axiom /= => hclosed. *) +(* move=> p a b hab. *) +(* Admitted. *) + +(* Lemma closed_form_rcf_odd_sqr : rcf_odd -> rcf_square *) +(* -> GRing.ClosedField.axiom [ringType of C]. *) +(* Proof. *) +(* Admitted. *) + +(* Lemma closed_form_ivt : rcf_axiom R -> GRing.ClosedField.axiom [ringType of C]. *) +(* Proof. *) +(* move/rcf_odd_sqr_from_ivt; case. *) +(* exact: closed_form_rcf_odd_sqr. *) +(* Qed. *) + +(* End RcfDef. *) + +Section ComplexClosed. + +Variable R : rcfType. + +Definition sqrtc (x : R[i]) : R[i] := + let: a +i* b := x in + let sgr1 b := if b == 0 then 1 else sgr b in + let r := sqrtr (a^+2 + b^+2) in + (sqrtr ((r + a)/2%:R)) +i* (sgr1 b * sqrtr ((r - a)/2%:R)). + +Lemma sqr_sqrtc : forall x, (sqrtc x) ^+ 2 = x. +Proof. +have sqr: forall x : R, x ^+ 2 = x * x. + by move=> x; rewrite exprS expr1. +case=> a b; rewrite exprS expr1; simpc. +have F0: 2%:R != 0 :> R by rewrite pnatr_eq0. +have F1: 0 <= 2%:R^-1 :> R by rewrite invr_ge0 ler0n. +have F2: `|a| <= sqrtr (a^+2 + b^+2). + rewrite -sqrtr_sqr ler_wsqrtr //. + by rewrite addrC -subr_ge0 addrK exprn_even_ge0. +have F3: 0 <= (sqrtr (a ^+ 2 + b ^+ 2) - a) / 2%:R. + rewrite mulr_ge0 // subr_ge0 (ler_trans _ F2) //. + by rewrite -(maxrN a) ler_maxr lerr. +have F4: 0 <= (sqrtr (a ^+ 2 + b ^+ 2) + a) / 2%:R. + rewrite mulr_ge0 // -{2}[a]opprK subr_ge0 (ler_trans _ F2) //. + by rewrite -(maxrN a) ler_maxr lerr orbT. +congr (_ +i* _); set u := if _ then _ else _. + rewrite mulrCA !mulrA. + have->: (u * u) = 1. + rewrite /u; case: (altP (_ =P _)); rewrite ?mul1r //. + by rewrite -expr2 sqr_sg => ->. + rewrite mul1r -!sqr !sqr_sqrtr //. + rewrite [_+a]addrC -mulrBl opprD addrA addrK. + by rewrite opprK -mulr2n -mulr_natl [_*a]mulrC mulfK. +rewrite mulrCA -mulrA -mulrDr [sqrtr _ * _]mulrC. +rewrite -mulr2n -sqrtrM // mulrAC !mulrA ?[_ * (_ - _)]mulrC -subr_sqr. +rewrite sqr_sqrtr; last first. + by rewrite ler_paddr // exprn_even_ge0. +rewrite [_^+2 + _]addrC addrK -mulrA -expr2 sqrtrM ?exprn_even_ge0 //. +rewrite !sqrtr_sqr -mulr_natr. +rewrite [`|_^-1|]ger0_norm // -mulrA [_ * _%:R]mulrC divff //. +rewrite mulr1 /u; case: (_ =P _)=>[->|]. + by rewrite normr0 mulr0. +by rewrite mulr_sg_norm. +Qed. + +Lemma sqrtc_sqrtr : + forall (x : R[i]), 0 <= x -> sqrtc x = (sqrtr (Re x))%:C. +Proof. +move=> [a b] /andP [/eqP->] /= a_ge0. +rewrite eqxx mul1r [0 ^+ _]exprS mul0r addr0 sqrtr_sqr. +rewrite ger0_norm // subrr mul0r sqrtr0 -mulr2n. +by rewrite -[_*+2]mulr_natr mulfK // pnatr_eq0. +Qed. + +Lemma sqrtc0 : sqrtc 0 = 0. +Proof. by rewrite sqrtc_sqrtr ?lerr // sqrtr0. Qed. + +Lemma sqrtc1 : sqrtc 1 = 1. +Proof. by rewrite sqrtc_sqrtr ?ler01 // sqrtr1. Qed. + +Lemma sqrtN1 : sqrtc (-1) = 'i. +Proof. +rewrite /sqrtc /= oppr0 eqxx [0^+_]exprS mulr0 addr0. +rewrite exprS expr1 mulN1r opprK sqrtr1 subrr mul0r sqrtr0. +by rewrite mul1r -mulr2n divff ?sqrtr1 // pnatr_eq0. +Qed. + +Lemma sqrtc_ge0 (x : R[i]) : (0 <= sqrtc x) = (0 <= x). +Proof. +apply/idP/idP=> [psx|px]; last first. + by rewrite sqrtc_sqrtr // lecR sqrtr_ge0. +by rewrite -[x]sqr_sqrtc exprS expr1 mulr_ge0. +Qed. + +Lemma sqrtc_eq0 (x : R[i]) : (sqrtc x == 0) = (x == 0). +Proof. +apply/eqP/eqP=> [eqs|->]; last by rewrite sqrtc0. +by rewrite -[x]sqr_sqrtc eqs exprS mul0r. +Qed. + +Lemma normcE x : `|x| = sqrtc (x * x^*). +Proof. +case: x=> a b; simpc; rewrite [b * a]mulrC addNr sqrtc_sqrtr //. +by simpc; rewrite /= addr_ge0 ?sqr_ge0. +Qed. + +Lemma sqr_normc (x : R[i]) : (`|x| ^+ 2) = x * x^*. +Proof. by rewrite normcE sqr_sqrtc. Qed. + +Lemma normc_ge_Re (x : R[i]) : `|Re x|%:C <= `|x|. +Proof. +by case: x => a b; simpc; rewrite -sqrtr_sqr ler_wsqrtr // ler_addl sqr_ge0. +Qed. + +Lemma normcJ (x : R[i]) : `|x^*| = `|x|. +Proof. by case: x => a b; simpc; rewrite /= sqrrN. Qed. + +Lemma invc_norm (x : R[i]) : x^-1 = `|x|^-2 * x^*. +Proof. +case: (altP (x =P 0)) => [->|dx]; first by rewrite rmorph0 mulr0 invr0. +apply: (mulIf dx); rewrite mulrC divff // -mulrA [_^* * _]mulrC -(sqr_normc x). +by rewrite mulVf // expf_neq0 ?normr_eq0. +Qed. + +Lemma canonical_form (a b c : R[i]) : + a != 0 -> + let d := b ^+ 2 - 4%:R * a * c in + let r1 := (- b - sqrtc d) / 2%:R / a in + let r2 := (- b + sqrtc d) / 2%:R / a in + a *: 'X^2 + b *: 'X + c%:P = a *: (('X - r1%:P) * ('X - r2%:P)). +Proof. +move=> a_neq0 d r1 r2. +rewrite !(mulrDr, mulrDl, mulNr, mulrN, opprK, scalerDr). +rewrite [_ * _%:P]mulrC !mul_polyC !scalerN !scalerA -!addrA; congr (_ + _). +rewrite addrA; congr (_ + _). + rewrite -opprD -scalerDl -scaleNr; congr(_ *: _). + rewrite ![a * _]mulrC !divfK // !mulrDl addrACA !mulNr addNr addr0. + by rewrite -opprD opprK -mulrDr -mulr2n -mulr_natl divff ?mulr1 ?pnatr_eq0. +symmetry; rewrite -!alg_polyC scalerA; congr (_%:A). +rewrite [a * _]mulrC divfK // /r2 mulrA mulrACA -invfM -natrM -subr_sqr. +rewrite sqr_sqrtc sqrrN /d opprB addrC addrNK -2!mulrA. +by rewrite mulrACA -natf_div // mul1r mulrAC divff ?mul1r. +Qed. + +Lemma monic_canonical_form (b c : R[i]) : + let d := b ^+ 2 - 4%:R * c in + let r1 := (- b - sqrtc d) / 2%:R in + let r2 := (- b + sqrtc d) / 2%:R in + 'X^2 + b *: 'X + c%:P = (('X - r1%:P) * ('X - r2%:P)). +Proof. +by rewrite /= -['X^2]scale1r canonical_form ?oner_eq0 // scale1r mulr1 !divr1. +Qed. + +Section extramx. +(* missing lemmas from matrix.v or mxalgebra.v *) + +Lemma mul_mx_rowfree_eq0 (K : fieldType) (m n p: nat) + (W : 'M[K]_(m,n)) (V : 'M[K]_(n,p)) : + row_free V -> (W *m V == 0) = (W == 0). +Proof. by move=> free; rewrite -!mxrank_eq0 mxrankMfree ?mxrank_eq0. Qed. + +Lemma sub_sums_genmxP (F : fieldType) (I : finType) (P : pred I) (m n : nat) + (A : 'M[F]_(m, n)) (B_ : I -> 'M_(m, n)) : +reflect (exists u_ : I -> 'M_m, A = \sum_(i | P i) u_ i *m B_ i) + (A <= \sum_(i | P i) <<B_ i>>)%MS. +Proof. +apply: (iffP idP); last first. + by move=> [u_ ->]; rewrite summx_sub_sums // => i _; rewrite genmxE submxMl. +move=> /sub_sumsmxP [u_ hA]. +have Hu i : exists v, u_ i *m <<B_ i>>%MS = v *m B_ i. + by apply/submxP; rewrite (submx_trans (submxMl _ _)) ?genmxE. +exists (fun i => projT1 (sig_eqW (Hu i))); rewrite hA. +by apply: eq_bigr => i /= P_i; case: sig_eqW. +Qed. + +Lemma mulmxP (K : fieldType) (m n : nat) (A B : 'M[K]_(m, n)) : + reflect (forall u : 'rV__, u *m A = u *m B) (A == B). +Proof. +apply: (iffP eqP) => [-> //|eqAB]. +apply: (@row_full_inj _ _ _ _ 1%:M); first by rewrite row_full_unit unitmx1. +by apply/row_matrixP => i; rewrite !row_mul eqAB. +Qed. + +Section Skew. + +Variable (K : numFieldType). + +Implicit Types (phK : phant K) (n : nat). + +Definition skew_vec n i j : 'rV[K]_(n * n) := + (mxvec ((delta_mx i j)) - (mxvec (delta_mx j i))). + +Definition skew_def phK n : 'M[K]_(n * n) := + (\sum_(i | ((i.2 : 'I__) < (i.1 : 'I__))%N) <<skew_vec i.1 i.2>>)%MS. + +Variable (n : nat). +Local Notation skew := (@skew_def (Phant K) n). + + +Lemma skew_direct_sum : mxdirect skew. +Proof. +apply/mxdirect_sumsE => /=; split => [i _|]; first exact: mxdirect_trivial. +apply/mxdirect_sumsP => [] [i j] /= lt_ij; apply/eqP; rewrite -submx0. +apply/rV_subP => v; rewrite sub_capmx => /andP []; rewrite !genmxE. +move=> /submxP [w ->] /sub_sums_genmxP [/= u_]. +move/matrixP => /(_ 0 (mxvec_index i j)); rewrite !mxE /= big_ord1. +rewrite /skew_vec /= !mxvec_delta !mxE !eqxx /=. +have /(_ _ _ (_, _) (_, _)) /= eq_mviE := + inj_eq (bij_inj (onT_bij (curry_mxvec_bij _ _))). +rewrite eq_mviE xpair_eqE -!val_eqE /= eq_sym andbb. +rewrite ltn_eqF // subr0 mulr1 summxE big1. + rewrite [w as X in X *m _]mx11_scalar => ->. + by rewrite mul_scalar_mx scale0r submx0. +move=> [i' j'] /= /andP[lt_j'i']. +rewrite xpair_eqE /= => neq'_ij. +rewrite /= !mxvec_delta !mxE big_ord1 !mxE !eqxx !eq_mviE. +rewrite !xpair_eqE /= [_ == i']eq_sym [_ == j']eq_sym (negPf neq'_ij) /=. +set z := (_ && _); suff /negPf -> : ~~ z by rewrite subrr mulr0. +by apply: contraL lt_j'i' => /andP [/eqP <- /eqP <-]; rewrite ltnNge ltnW. +Qed. +Hint Resolve skew_direct_sum. + +Lemma rank_skew : \rank skew = (n * n.-1)./2. +Proof. +rewrite /skew (mxdirectP _) //= -bin2 -triangular_sum big_mkord. +rewrite (eq_bigr (fun _ => 1%N)); last first. + move=> [i j] /= lt_ij; rewrite genmxE. + apply/eqP; rewrite eqn_leq rank_leq_row /= lt0n mxrank_eq0. + rewrite /skew_vec /= !mxvec_delta /= subr_eq0. + set j1 := mxvec_index _ _. + apply/negP => /eqP /matrixP /(_ 0 j1) /=; rewrite !mxE eqxx /=. + have /(_ _ _ (_, _) (_, _)) -> := + inj_eq (bij_inj (onT_bij (curry_mxvec_bij _ _))). + rewrite xpair_eqE -!val_eqE /= eq_sym andbb ltn_eqF //. + by move/eqP; rewrite oner_eq0. +transitivity (\sum_(i < n) (\sum_(j < n | j < i) 1))%N. + by rewrite pair_big_dep. +apply: eq_bigr => [] [[|i] Hi] _ /=; first by rewrite big1. +rewrite (eq_bigl _ _ (fun _ => ltnS _ _)). +have [n_eq0|n_gt0] := posnP n; first by move: Hi (Hi); rewrite {1}n_eq0. +rewrite -[n]prednK // big_ord_narrow_leq /=. + by rewrite -ltnS prednK // (leq_trans _ Hi). +by rewrite sum_nat_const card_ord muln1. +Qed. + +Lemma skewP (M : 'rV_(n * n)) : + reflect ((vec_mx M)^T = - vec_mx M) (M <= skew)%MS. +Proof. +apply: (iffP idP). + move/sub_sumsmxP => [v ->]; rewrite !linear_sum /=. + apply: eq_bigr => [] [i j] /= lt_ij; rewrite !mulmx_sum_row !linear_sum /=. + apply: eq_bigr => k _; rewrite !linearZ /=; congr (_ *: _) => {v}. + set r := << _ >>%MS; move: (row _ _) (row_sub k r) => v. + move: @r; rewrite /= genmxE => /sub_rVP [a ->]; rewrite !linearZ /=. + by rewrite /skew_vec !linearB /= !mxvecK !scalerN opprK addrC !trmx_delta. +move=> skewM; pose M' := vec_mx M. +pose xM i j := (M' i j - M' j i) *: skew_vec i j. +suff -> : M = 2%:R^-1 *: + (\sum_(i | true && ((i.2 : 'I__) < (i.1 : 'I__))%N) xM i.1 i.2). + rewrite scalemx_sub // summx_sub_sums // => [] [i j] /= lt_ij. + by rewrite scalemx_sub // genmxE. +rewrite /xM /= /skew_vec (eq_bigr _ (fun _ _ => scalerBr _ _ _)). +rewrite big_split /= sumrN !(eq_bigr _ (fun _ _ => scalerBl _ _ _)). +rewrite !big_split /= !sumrN opprD ?opprK addrACA [- _ + _]addrC. +rewrite -!sumrN -2!big_split /=. +rewrite /xM /= /skew_vec -!(eq_bigr _ (fun _ _ => scalerBr _ _ _)). +apply: (can_inj vec_mxK); rewrite !(linearZ, linearB, linearD, linear_sum) /=. +have -> /= : vec_mx M = 2%:R^-1 *: (M' - M'^T). + by rewrite skewM opprK -mulr2n -scaler_nat scalerA mulVf ?pnatr_eq0 ?scale1r. +rewrite {1 2}[M']matrix_sum_delta; congr (_ *: _). +rewrite pair_big /= !linear_sum /= -big_split /=. +rewrite (bigID (fun ij => (ij.2 : 'I__) < (ij.1 : 'I__))%N) /=; congr (_ + _). + apply: eq_bigr => [] [i j] /= lt_ij. + by rewrite !linearZ linearB /= ?mxvecK trmx_delta scalerN scalerBr. +rewrite (bigID (fun ij => (ij.1 : 'I__) == (ij.2 : 'I__))%N) /=. +rewrite big1 ?add0r; last first. + by move=> [i j] /= /andP[_ /eqP ->]; rewrite linearZ /= trmx_delta subrr. +rewrite (@reindex_inj _ _ _ _ (fun ij => (ij.2, ij.1))) /=; last first. + by move=> [? ?] [? ?] [] -> ->. +apply: eq_big => [] [i j] /=; first by rewrite -leqNgt ltn_neqAle andbC. +by rewrite !linearZ linearB /= ?mxvecK trmx_delta scalerN scalerBr. +Qed. + +End Skew. + +Notation skew K n := (@skew_def _ (Phant K) n). + +Section Companion. + +Variable (K : fieldType). + +Lemma companion_subproof (p : {poly K}) : + {M : 'M[K]_((size p).-1)| p \is monic -> char_poly M = p}. +Proof. +have simp := (castmxE, mxE, castmx_id, cast_ord_id). +case Hsp: (size p) => [|sp] /=. + move/eqP: Hsp; rewrite size_poly_eq0 => /eqP ->. + by exists 0; rewrite qualifE lead_coef0 eq_sym oner_eq0. +case: sp => [|sp] in Hsp *. + move: Hsp => /eqP/size_poly1P/sig2_eqW [c c_neq0 ->]. + by exists ((-c)%:M); rewrite monicE lead_coefC => /eqP ->; apply: det_mx00. +have addn1n n : (n + 1 = 1 + n)%N by rewrite addn1. +exists (castmx (erefl _, addn1n _) + (block_mx (\row_(i < sp) - p`_(sp - i)) (-p`_0)%:M + 1%:M 0)). +elim/poly_ind: p sp Hsp (addn1n _) => [|p c IHp] sp; first by rewrite size_poly0. +rewrite size_MXaddC. +have [->|p_neq0] //= := altP eqP; first by rewrite size_poly0; case: ifP. +move=> [Hsp] eq_cast. +rewrite monicE lead_coefDl ?size_polyC ?size_mul ?polyX_eq0 //; last first. + by rewrite size_polyX addn2 Hsp ltnS (leq_trans (leq_b1 _)). +rewrite lead_coefMX -monicE => p_monic. +rewrite -/_`_0 coefD coefMX coefC eqxx add0r. +case: sp => [|sp] in Hsp p_neq0 p_monic eq_cast *. + move: Hsp p_monic => /eqP/size_poly1P [l l_neq0 ->]. + rewrite monicE lead_coefC => /eqP ->; rewrite mul1r. + rewrite /char_poly /char_poly_mx thinmx0 flatmx0 castmx_id. + set b := (block_mx _ _ _ _); rewrite [map_mx _ b]map_block_mx => {b}. + rewrite !map_mx0 map_scalar_mx (@opp_block_mx _ 1 0 0 1) !oppr0. + set b := block_mx _ _ _ _; rewrite (_ : b = c%:P%:M); last first. + apply/matrixP => i j; rewrite !mxE; case: splitP => k /= Hk; last first. + by move: (ltn_ord i); rewrite Hk. + rewrite !ord1 !mxE; case: splitP => {k Hk} k /= Hk; first by move: (ltn_ord k). + by rewrite ord1 !mxE mulr1n rmorphN opprK. + by rewrite -rmorphD det_scalar. +rewrite /char_poly /char_poly_mx (expand_det_col _ ord_max). +rewrite big_ord_recr /= big_ord_recl //= big1 ?simp; last first. + move=> i _; rewrite !simp. + case: splitP => k /=; first by rewrite /bump leq0n ord1. + rewrite /bump leq0n => [] [Hik]; rewrite !simp. + case: splitP => l /=; first by move/eqP; rewrite gtn_eqF. + rewrite !ord1 addn0 => _ {l}; rewrite !simp -!val_eqE /=. + by rewrite /bump leq0n ltn_eqF ?ltnS ?add1n // mulr0n subrr mul0r. +case: splitP => i //=; rewrite !ord1 !simp => _ {i}. +case: splitP => i //=; first by move/eqP; rewrite gtn_eqF. +rewrite ord1 !simp => {i}. +case: splitP => i //=; rewrite ?ord1 ?simp // => /esym [eq_i_sp] _. +case: splitP => j //=; first by move/eqP; rewrite gtn_eqF. +rewrite ord1 !simp => {j} _. +rewrite eqxx mulr0n ?mulr1n rmorphN ?opprK !add0r !addr0 subr0 /=. +rewrite -[c%:P in X in _ = X]mulr1 addrC mulrC. +rewrite /cofactor -signr_odd addnn odd_double expr0 mul1r /=. +rewrite !linearB /= -!map_col' -!map_row'. +congr (_ * 'X + c%:P * _). + have coefE := (coefD, coefMX, coefC, eqxx, add0r, addr0). + rewrite -[X in _ = X](IHp sp Hsp _ p_monic) /char_poly /char_poly_mx. + congr (\det (_ - _)). + apply/matrixP => k l; rewrite !simp -val_eqE /=; + by rewrite /bump ![(sp < _)%N]ltnNge ?leq_ord. + apply/matrixP => k l; rewrite !simp. + case: splitP => k' /=; rewrite ?ord1 /bump ltnNge leq_ord add0n. + case: splitP => [k'' /= |k'' -> //]; rewrite ord1 !simp => k_eq0 _. + case: splitP => l' /=; rewrite ?ord1 /bump ltnNge leq_ord add0n !simp; + last by move/eqP; rewrite ?addn0 ltn_eqF. + move<-; case: splitP => l'' /=; rewrite ?ord1 ?addn0 !simp. + by move<-; rewrite subSn ?leq_ord ?coefE. + move->; rewrite eqxx mulr1n ?coefE subSn ?subrr //=. + by rewrite !rmorphN ?subnn addr0. + case: splitP => k'' /=; rewrite ?ord1 => -> // []; rewrite !simp. + case: splitP => l' /=; rewrite /bump ltnNge leq_ord add0n !simp -?val_eqE /=; + last by rewrite ord1 addn0 => /eqP; rewrite ltn_eqF. + by case: splitP => l'' /= -> <- <-; rewrite !simp // ?ord1 ?addn0 ?ltn_eqF. +move=> {IHp Hsp p_neq0 p_monic}; rewrite add0n; set s := _ ^+ _; +apply: (@mulfI _ s); first by rewrite signr_eq0. +rewrite mulrA -expr2 sqrr_sign mulr1 mul1r /s. +pose fix D n : 'M[{poly K}]_n.+1 := + if n is n'.+1 then block_mx (-1 :'M_1) ('X *: pid_mx 1) + 0 (D n') else -1. +pose D' n : 'M[{poly K}]_n.+1 := \matrix_(i, j) ('X *+ (i.+1 == j) - (i == j)%:R). +set M := (_ - _); have -> : M = D' sp. + apply/matrixP => k l; rewrite !simp. + case: splitP => k' /=; rewrite ?ord1 !simp // /bump leq0n add1n; case. + case: splitP => l' /=; rewrite /bump ltnNge leq_ord add0n; last first. + by move/eqP; rewrite ord1 addn0 ltn_eqF. + rewrite !simp -!val_eqE /= /bump leq0n ltnNge leq_ord [(true + _)%N]add1n ?add0n. + by move=> -> ->; rewrite polyC_muln. +have -> n : D' n = D n. + clear -simp; elim: n => [|n IHn] //=; apply/matrixP => i j; rewrite !simp. + by rewrite !ord1 /= ?mulr0n sub0r. + case: splitP => i' /=; rewrite -!val_eqE /= ?ord1 !simp => -> /=. + case: splitP => j' /=; rewrite ?ord1 !simp => -> /=; first by rewrite sub0r. + by rewrite eqSS andbT subr0 mulr_natr. + by case: splitP => j' /=; rewrite ?ord1 -?IHn ?simp => -> //=; rewrite subr0. +elim: sp {eq_cast i M eq_i_sp s} => [|n IHn]. + by rewrite /= (_ : -1 = (-1)%:M) ?det_scalar // rmorphN. +rewrite /= (@det_ublock _ 1 n.+1) IHn. +by rewrite (_ : -1 = (-1)%:M) ?det_scalar // rmorphN. +Qed. + +Definition companion (p : {poly K}) : 'M[K]_((size p).-1) := + projT1 (companion_subproof p). + +Lemma companionK (p : {poly K}) : p \is monic -> char_poly (companion p) = p. +Proof. exact: projT2 (companion_subproof _). Qed. + +End Companion. + +Section Restriction. + +Variable K : fieldType. +Variable m : nat. +Variables (V : 'M[K]_m). + +Implicit Types f : 'M[K]_m. + +Definition restrict f : 'M_(\rank V) := row_base V *m f *m (pinvmx (row_base V)). + +Lemma stable_row_base f : + (row_base V *m f <= row_base V)%MS = (V *m f <= V)%MS. +Proof. +rewrite eq_row_base. +by apply/idP/idP=> /(submx_trans _) ->; rewrite ?submxMr ?eq_row_base. +Qed. + +Lemma eigenspace_restrict f : (V *m f <= V)%MS -> + forall n a (W : 'M_(n, \rank V)), + (W <= eigenspace (restrict f) a)%MS = + (W *m row_base V <= eigenspace f a)%MS. +Proof. +move=> f_stabV n a W; apply/eigenspaceP/eigenspaceP; rewrite scalemxAl. + by move<-; rewrite -mulmxA -[X in _ = X]mulmxA mulmxKpV ?stable_row_base. +move/(congr1 (mulmx^~ (pinvmx (row_base V)))). +rewrite -2!mulmxA [_ *m (f *m _)]mulmxA => ->. +by apply: (row_free_inj (row_base_free V)); rewrite mulmxKpV ?submxMl. +Qed. + +Lemma eigenvalue_restrict f : (V *m f <= V)%MS -> + {subset eigenvalue (restrict f) <= eigenvalue f}. +Proof. +move=> f_stabV a /eigenvalueP [x /eigenspaceP]; rewrite eigenspace_restrict //. +move=> /eigenspaceP Hf x_neq0; apply/eigenvalueP. +by exists (x *m row_base V); rewrite ?mul_mx_rowfree_eq0 ?row_base_free. +Qed. + +Lemma restrictM : {in [pred f | (V *m f <= V)%MS] &, + {morph restrict : f g / f *m g}}. +Proof. +move=> f g; rewrite !inE => Vf Vg /=. +by rewrite /restrict 2!mulmxA mulmxA mulmxKpV ?stable_row_base. +Qed. + +End Restriction. + +End extramx. +Notation skew K n := (@skew_def _ (Phant K) n). + +Section Paper_HarmDerksen. + +(* Following http://www.math.lsa.umich.edu/~hderksen/preprints/linalg.pdf *) +(* quite literally except for Lemma5 where we don't use hermitian matrices. *) +(* Instead we encode the morphism by hand in 'M[R]_(n * n), which turns out *) +(* to be very clumsy for formalizing commutation and the end of Lemma 4. *) +(* Moreover, the Qed takes time, so it would be far much better to formalize *) +(* Herm C n and use it instead ! *) + +Implicit Types (K : fieldType). + +Definition CommonEigenVec_def K (phK : phant K) (d r : nat) := + forall (m : nat) (V : 'M[K]_m), ~~ (d %| \rank V) -> + forall (sf : seq 'M_m), size sf = r -> + {in sf, forall f, (V *m f <= V)%MS} -> + {in sf &, forall f g, f *m g = g *m f} -> + exists2 v : 'rV_m, (v != 0) & forall f, f \in sf -> + exists a, (v <= eigenspace f a)%MS. +Notation CommonEigenVec K d r := (@CommonEigenVec_def _ (Phant K) d r). + +Definition Eigen1Vec_def K (phK : phant K) (d : nat) := + forall (m : nat) (V : 'M[K]_m), ~~ (d %| \rank V) -> + forall (f : 'M_m), (V *m f <= V)%MS -> exists a, eigenvalue f a. +Notation Eigen1Vec K d := (@Eigen1Vec_def _ (Phant K) d). + +Lemma Eigen1VecP (K : fieldType) (d : nat) : + CommonEigenVec K d 1%N <-> Eigen1Vec K d. +Proof. +split=> [Hd m V HV f|Hd m V HV [] // f [] // _ /(_ _ (mem_head _ _))] f_stabV. + have [] := Hd _ _ HV [::f] (erefl _). + + by move=> ?; rewrite in_cons orbF => /eqP ->. + + by move=> ? ?; rewrite /= !in_cons !orbF => /eqP -> /eqP ->. + move=> v v_neq0 /(_ f (mem_head _ _)) [a /eigenspaceP]. + by exists a; apply/eigenvalueP; exists v. +have [a /eigenvalueP [v /eigenspaceP v_eigen v_neq0]] := Hd _ _ HV _ f_stabV. +by exists v => // ?; rewrite in_cons orbF => /eqP ->; exists a. +Qed. + +Lemma Lemma3 K d : Eigen1Vec K d -> forall r, CommonEigenVec K d r.+1. +Proof. +move=> E1V_K_d; elim => [|r IHr m V]; first exact/Eigen1VecP. +move: (\rank V) {-2}V (leqnn (\rank V)) => n {V}. +elim: n m => [|n IHn] m V. + by rewrite leqn0 => /eqP ->; rewrite dvdn0. +move=> le_rV_Sn HrV [] // f sf /= [] ssf f_sf_stabV f_sf_comm. +have [->|f_neq0] := altP (f =P 0). + have [||v v_neq0 Hsf] := (IHr _ _ HrV _ ssf). + + by move=> g f_sf /=; rewrite f_sf_stabV // in_cons f_sf orbT. + + move=> g h g_sf h_sf /=. + by apply: f_sf_comm; rewrite !in_cons ?g_sf ?h_sf ?orbT. + exists v => // g; rewrite in_cons => /orP [/eqP->|]; last exact: Hsf. + by exists 0; apply/eigenspaceP; rewrite mulmx0 scale0r. +have f_stabV : (V *m f <= V)%MS by rewrite f_sf_stabV ?mem_head. +have sf_stabV : {in sf, forall f, (V *m f <= V)%MS}. + by move=> g g_sf /=; rewrite f_sf_stabV // in_cons g_sf orbT. +pose f' := restrict V f; pose sf' := map (restrict V) sf. +have [||a a_eigen_f'] := E1V_K_d _ 1%:M _ f'; do ?by rewrite ?mxrank1 ?submx1. +pose W := (eigenspace f' a)%MS; pose Z := (f' - a%:M). +have rWZ : (\rank W + \rank Z)%N = \rank V. + by rewrite (mxrank_ker (f' - a%:M)) subnK // rank_leq_row. +have f'_stabW : (W *m f' <= W)%MS. + by rewrite (eigenspaceP (submx_refl _)) scalemx_sub. +have f'_stabZ : (Z *m f' <= Z)%MS. + rewrite (submx_trans _ (submxMl f' _)) //. + by rewrite mulmxDl mulmxDr mulmxN mulNmx scalar_mxC. +have sf'_comm : {in [::f' & sf'] &, forall f g, f *m g = g *m f}. + move=> g' h' /=; rewrite -!map_cons. + move=> /mapP [g g_s_sf -> {g'}] /mapP [h h_s_sf -> {h'}]. + by rewrite -!restrictM ?inE /= ?f_sf_stabV // f_sf_comm. +have sf'_stabW : {in sf', forall f, (W *m f <= W)%MS}. + move=> g g_sf /=; apply/eigenspaceP. + rewrite -mulmxA -[g *m _]sf'_comm ?(mem_head, in_cons, g_sf, orbT) //. + by rewrite mulmxA scalemxAl (eigenspaceP (submx_refl _)). +have sf'_stabZ : {in sf', forall f, (Z *m f <= Z)%MS}. + move=> g g_sf /=. + rewrite mulmxBl sf'_comm ?(mem_head, in_cons, g_sf, orbT) //. + by rewrite -scalar_mxC -mulmxBr submxMl. +have [eqWV|neqWV] := altP (@eqmxP _ _ _ _ W 1%:M). + have [] // := IHr _ W _ sf'; do ?by rewrite ?eqWV ?mxrank1 ?size_map. + move=> g h g_sf' h_sf'; apply: sf'_comm; + by rewrite in_cons (g_sf', h_sf') orbT. + move=> v v_neq0 Hv; exists (v *m row_base V). + by rewrite mul_mx_rowfree_eq0 ?row_base_free. + move=> g; rewrite in_cons => /orP [/eqP ->|g_sf]; last first. + have [|b] := Hv (restrict V g); first by rewrite map_f. + by rewrite eigenspace_restrict // ?sf_stabV //; exists b. + by exists a; rewrite -eigenspace_restrict // eqWV submx1. +have lt_WV : (\rank W < \rank V)%N. + rewrite -[X in (_ < X)%N](@mxrank1 K) rank_ltmx //. + by rewrite ltmxEneq neqWV // submx1. +have ltZV : (\rank Z < \rank V)%N. + rewrite -[X in (_ < X)%N]rWZ -subn_gt0 addnK lt0n mxrank_eq0 -lt0mx. + move: a_eigen_f' => /eigenvalueP [v /eigenspaceP] sub_vW v_neq0. + by rewrite (ltmx_sub_trans _ sub_vW) // lt0mx. +have [] // := IHn _ (if d %| \rank Z then W else Z) _ _ [:: f' & sf']. ++ by rewrite -ltnS (@leq_trans (\rank V)) //; case: ifP. ++ by apply: contra HrV; case: ifP => [*|-> //]; rewrite -rWZ dvdn_add. ++ by rewrite /= size_map ssf. ++ move=> g; rewrite in_cons => /= /orP [/eqP -> {g}|g_sf']; case: ifP => _ //; + by rewrite (sf'_stabW, sf'_stabZ). +move=> v v_neq0 Hv; exists (v *m row_base V). + by rewrite mul_mx_rowfree_eq0 ?row_base_free. +move=> g Hg; have [|b] := Hv (restrict V g); first by rewrite -map_cons map_f. +rewrite eigenspace_restrict //; first by exists b. +by move: Hg; rewrite in_cons => /orP [/eqP -> //|/sf_stabV]. +Qed. + +Lemma Lemma4 r : CommonEigenVec R 2 r.+1. +Proof. +apply: Lemma3=> m V hV f f_stabV. +have [|a] := @odd_poly_root _ (char_poly (restrict V f)). + by rewrite size_char_poly /= -dvdn2. +rewrite -eigenvalue_root_char => /eigenvalueP [v] /eigenspaceP v_eigen v_neq0. +exists a; apply/eigenvalueP; exists (v *m row_base V). + by apply/eigenspaceP; rewrite -eigenspace_restrict. +by rewrite mul_mx_rowfree_eq0 ?row_base_free. +Qed. + +Notation toC := (real_complex R). +Notation MtoC := (map_mx toC). + +Lemma Lemma5 : Eigen1Vec R[i] 2. +Proof. +move=> m V HrV f f_stabV. +suff: exists a, eigenvalue (restrict V f) a. + by move=> [a /eigenvalue_restrict Hf]; exists a; apply: Hf. +move: (\rank V) (restrict V f) => {f f_stabV V m} n f in HrV *. +pose u := map_mx (@Re R) f; pose v := map_mx (@Im R) f. +have fE : f = MtoC u + 'i *: MtoC v. + rewrite /u /v [f]lock; apply/matrixP => i j; rewrite !mxE /=. + by case: (locked f i j) => a b; simpc. +move: u v => u v in fE *. +pose L1fun : 'M[R]_n -> _ := + 2%:R^-1 \*: (mulmxr u \+ (mulmxr v \o trmx) + \+ ((mulmx (u^T)) \- (mulmx (v^T) \o trmx))). +pose L1 := lin_mx [linear of L1fun]. +pose L2fun : 'M[R]_n -> _ := + 2%:R^-1 \*: (((@GRing.opp _) \o (mulmxr u \o trmx) \+ mulmxr v) + \+ ((mulmx (u^T) \o trmx) \+ (mulmx (v^T)))). +pose L2 := lin_mx [linear of L2fun]. +have [] := @Lemma4 _ _ 1%:M _ [::L1; L2] (erefl _). ++ by move: HrV; rewrite mxrank1 !dvdn2 ?negbK odd_mul andbb. ++ by move=> ? _ /=; rewrite submx1. ++ suff {f fE}: L1 *m L2 = L2 *m L1. + move: L1 L2 => L1 L2 commL1L2 La Lb. + rewrite !{1}in_cons !{1}in_nil !{1}orbF. + by move=> /orP [] /eqP -> /orP [] /eqP -> //; symmetry. + apply/eqP/mulmxP => x; rewrite [X in X = _]mulmxA [X in _ = X]mulmxA. + rewrite 4!mul_rV_lin !mxvecK /= /L1fun /L2fun /=; congr (mxvec (_ *: _)). + move=> {L1 L2 L1fun L2fun}. + case: n {x} (vec_mx x) => [//|n] x in HrV u v *. + do ?[rewrite -(scalemxAl, scalemxAr, scalerN, scalerDr) + |rewrite (mulmxN, mulNmx, trmxK, trmx_mul) + |rewrite ?[(_ *: _)^T]linearZ ?[(_ + _)^T]linearD ?[(- _)^T]linearN /=]. + congr (_ *: _). + rewrite !(mulmxDr, mulmxDl, mulNmx, mulmxN, mulmxA, opprD, opprK). + do ![move: (_ *m _ *m _)] => t1 t2 t3 t4 t5 t6 t7 t8 t9 t10 t11 t12. + rewrite [X in X + _ + _]addrC [X in X + _ = _]addrACA. + rewrite [X in _ = (_ + _ + X) + _]addrC [X in _ = X + _]addrACA. + rewrite [X in _ + (_ + _ + X)]addrC [X in _ + X = _]addrACA. + rewrite [X in _ = _ + (X + _)]addrC [X in _ = _ + X]addrACA. + rewrite [X in X = _]addrACA [X in _ = X]addrACA; congr (_ + _). + by rewrite addrC [X in X + _ = _]addrACA [X in _ + X = _]addrACA. +move=> g g_neq0 Hg; have [] := (Hg L1, Hg L2). +rewrite !(mem_head, in_cons, orbT) => []. +move=> [//|a /eigenspaceP g_eigenL1] [//|b /eigenspaceP g_eigenL2]. +rewrite !mul_rV_lin /= /L1fun /L2fun /= in g_eigenL1 g_eigenL2. +do [move=> /(congr1 vec_mx); rewrite mxvecK linearZ /=] in g_eigenL1. +do [move=> /(congr1 vec_mx); rewrite mxvecK linearZ /=] in g_eigenL2. +move=> {L1 L2 L1fun L2fun Hg HrV}. +set vg := vec_mx g in g_eigenL1 g_eigenL2. +exists (a +i* b); apply/eigenvalueP. +pose w := (MtoC vg - 'i *: MtoC vg^T). +exists (nz_row w); last first. + rewrite nz_row_eq0 subr_eq0; apply: contraNneq g_neq0 => Hvg. + rewrite -vec_mx_eq0; apply/eqP/matrixP => i j; rewrite !mxE /=. + move: Hvg => /matrixP /(_ i j); rewrite !mxE /=; case. + by rewrite !(mul0r, mulr0, add0r, mul1r, oppr0) => ->. +apply/eigenspaceP. +case: n f => [|n] f in u v g g_neq0 vg w fE g_eigenL1 g_eigenL2 *. + by rewrite thinmx0 eqxx in g_neq0. +rewrite (submx_trans (nz_row_sub _)) //; apply/eigenspaceP. +rewrite fE [a +i* b]complexE /=. +rewrite !(mulmxDr, mulmxBl, =^~scalemxAr, =^~scalemxAl) -!map_mxM. +rewrite !(scalerDl, scalerDr, scalerN, =^~scalemxAr, =^~scalemxAl). +rewrite !scalerA /= mulrAC ['i * _]sqr_i ?mulN1r scaleN1r scaleNr !opprK. +rewrite [_ * 'i]mulrC -!scalerA -!map_mxZ /=. +do 2!rewrite [X in (_ - _) + X]addrC [_ - 'i *: _ + _]addrACA. +rewrite ![- _ + _]addrC -!scalerBr -!(rmorphB, rmorphD) /=. +congr (_ + 'i *: _); congr map_mx; rewrite -[_ *: _^T]linearZ /=; +rewrite -g_eigenL1 -g_eigenL2 linearZ -(scalerDr, scalerBr); +do ?rewrite ?trmxK ?trmx_mul ?[(_ + _)^T]linearD ?[(- _)^T]linearN /=; +rewrite -[in X in _ *: (_ + X)]addrC 1?opprD 1?opprB ?mulmxN ?mulNmx; +rewrite [X in _ *: X]addrACA. + rewrite -mulr2n [X in _ *: (_ + X)]addrACA subrr addNr !addr0. + by rewrite -scaler_nat scalerA mulVf ?pnatr_eq0 // scale1r. +rewrite subrr addr0 addrA addrAC -addrA -mulr2n addrC. +by rewrite -scaler_nat scalerA mulVf ?pnatr_eq0 // scale1r. +Qed. + +Lemma Lemma6 k r : CommonEigenVec R[i] (2^k.+1) r.+1. +Proof. +elim: k {-2}k (leqnn k) r => [|k IHk] l. + by rewrite leqn0 => /eqP ->; apply: Lemma3; apply: Lemma5. +rewrite leq_eqVlt ltnS => /orP [/eqP ->|/IHk //] r {l}. +apply: Lemma3 => m V Hn f f_stabV {r}. +have [dvd2n|Ndvd2n] := boolP (2 %| \rank V); last first. + exact: @Lemma5 _ _ Ndvd2n _ f_stabV. +suff: exists a, eigenvalue (restrict V f) a. + by move=> [a /eigenvalue_restrict Hf]; exists a; apply: Hf. +case: (\rank V) (restrict V f) => {f f_stabV V m} [|n] f in Hn dvd2n *. + by rewrite dvdn0 in Hn. +pose L1 := lin_mx [linear of mulmxr f \+ (mulmx f^T)]. +pose L2 := lin_mx [linear of mulmxr f \o mulmx f^T]. +have [] /= := IHk _ (leqnn _) _ _ (skew R[i] n.+1) _ [::L1; L2] (erefl _). ++ rewrite rank_skew; apply: contra Hn. + rewrite -(@dvdn_pmul2r 2) //= -expnSr muln2 -[_.*2]add0n. + have n_odd : odd n by rewrite dvdn2 /= ?negbK in dvd2n *. + have {2}<- : odd (n.+1 * n) = 0%N :> nat by rewrite odd_mul /= andNb. + by rewrite odd_double_half Gauss_dvdl // coprime_pexpl // coprime2n. ++ move=> L; rewrite 2!in_cons in_nil orbF => /orP [] /eqP ->; + apply/rV_subP => v /submxP [s -> {v}]; rewrite mulmxA; apply/skewP; + set u := _ *m skew _ _; + do [have /skewP : (u <= skew R[i] n.+1)%MS by rewrite submxMl]; + rewrite mul_rV_lin /= !mxvecK => skew_u. + by rewrite opprD linearD /= !trmx_mul skew_u mulmxN mulNmx addrC trmxK. + by rewrite !trmx_mul trmxK skew_u mulNmx mulmxN mulmxA. ++ suff commL1L2: L1 *m L2 = L2 *m L1. + move=> La Lb; rewrite !in_cons !in_nil !orbF. + by move=> /orP [] /eqP -> /orP [] /eqP -> //; symmetry. + apply/eqP/mulmxP => u; rewrite !mulmxA !mul_rV_lin ?mxvecK /=. + by rewrite !(mulmxDr, mulmxDl, mulmxA). +move=> v v_neq0 HL1L2; have [] := (HL1L2 L1, HL1L2 L2). +rewrite !(mem_head, in_cons) orbT => [] [] // a vL1 [] // b vL2 {HL1L2}. +move/eigenspaceP in vL1; move/eigenspaceP in vL2. +move: vL2 => /(congr1 vec_mx); rewrite linearZ mul_rV_lin /= mxvecK. +move: vL1 => /(congr1 vec_mx); rewrite linearZ mul_rV_lin /= mxvecK. +move=> /(canRL (addKr _)) ->; rewrite mulmxDl mulNmx => Hv. +pose p := 'X^2 + (- a) *: 'X + b%:P. +have : vec_mx v *m (horner_mx f p) = 0. + rewrite !(rmorphN, rmorphB, rmorphD, rmorphM) /= linearZ /=. + rewrite horner_mx_X horner_mx_C !mulmxDr mul_mx_scalar -Hv. + rewrite addrAC addrA mulmxA addrN add0r. + by rewrite -scalemxAl -scalemxAr scaleNr addrN. +rewrite [p]monic_canonical_form; move: (_ / 2%:R) (_ / 2%:R). +move=> r2 r1 {Hv p a b L1 L2 Hn}. +rewrite rmorphM !rmorphB /= horner_mx_X !horner_mx_C mulmxA => Hv. +have: exists2 w : 'M_n.+1, w != 0 & exists a, (w <= eigenspace f a)%MS. + move: Hv; set w := vec_mx _ *m _. + have [w_eq0 _|w_neq0 r2_eigen] := altP (w =P 0). + exists (vec_mx v); rewrite ?vec_mx_eq0 //; exists r1. + apply/eigenspaceP/eqP. + by rewrite -mul_mx_scalar -subr_eq0 -mulmxBr -/w w_eq0. + exists w => //; exists r2; apply/eigenspaceP/eqP. + by rewrite -mul_mx_scalar -subr_eq0 -mulmxBr r2_eigen. +move=> [w w_neq0 [a /(submx_trans (nz_row_sub _)) /eigenspaceP Hw]]. +by exists a; apply/eigenvalueP; exists (nz_row w); rewrite ?nz_row_eq0. +Qed. + +(* We enunciate a corollary of Theorem 7 *) +Corollary Theorem7' (m : nat) (f : 'M[R[i]]_m) : (0 < m)%N -> exists a, eigenvalue f a. +Proof. +case: m f => // m f _; have /Eigen1VecP := @Lemma6 m 0. +move=> /(_ m.+1 1 _ f) []; last by move=> a; exists a. ++ by rewrite mxrank1 (contra (dvdn_leq _)) // -ltnNge ltn_expl. ++ by rewrite submx1. +Qed. + +Lemma C_acf_axiom : GRing.ClosedField.axiom [ringType of R[i]]. +Proof. +move=> n c n_gt0; pose p := 'X^n - \poly_(i < n) c i. +suff [x rpx] : exists x, root p x. + exists x; move: rpx; rewrite /root /p hornerD hornerN hornerXn subr_eq0. + by move=> /eqP ->; rewrite horner_poly. +have p_monic : p \is monic. + rewrite qualifE lead_coefDl ?lead_coefXn //. + by rewrite size_opp size_polyXn ltnS size_poly. +have sp_gt1 : (size p > 1)%N. + by rewrite size_addl size_polyXn // size_opp ltnS size_poly. +case: n n_gt0 p => //= n _ p in p_monic sp_gt1 *. +have [] := Theorem7' (companion p); first by rewrite -(subnK sp_gt1) addn2. +by move=> x; rewrite eigenvalue_root_char companionK //; exists x. +Qed. + +Definition C_decFieldMixin := closed_fields_QEMixin C_acf_axiom. +Canonical C_decField := DecFieldType R[i] C_decFieldMixin. +Canonical C_closedField := ClosedFieldType R[i] C_acf_axiom. + +End Paper_HarmDerksen. + +End ComplexClosed. + +Definition complexalg := realalg[i]. + +Canonical complexalg_eqType := [eqType of complexalg]. +Canonical complexalg_choiceType := [choiceType of complexalg]. +Canonical complexalg_countype := [choiceType of complexalg]. +Canonical complexalg_zmodType := [zmodType of complexalg]. +Canonical complexalg_ringType := [ringType of complexalg]. +Canonical complexalg_comRingType := [comRingType of complexalg]. +Canonical complexalg_unitRingType := [unitRingType of complexalg]. +Canonical complexalg_comUnitRingType := [comUnitRingType of complexalg]. +Canonical complexalg_idomainType := [idomainType of complexalg]. +Canonical complexalg_fieldType := [fieldType of complexalg]. +Canonical complexalg_decDieldType := [decFieldType of complexalg]. +Canonical complexalg_closedFieldType := [closedFieldType of complexalg]. +Canonical complexalg_numDomainType := [numDomainType of complexalg]. +Canonical complexalg_numFieldType := [numFieldType of complexalg]. +Canonical complexalg_numClosedFieldType := [numClosedFieldType of complexalg]. + +Lemma complexalg_algebraic : integralRange (@ratr [unitRingType of complexalg]). +Proof. +move=> x; suff [p p_monic] : integralOver (real_complex _ \o realalg_of _) x. + by rewrite (eq_map_poly (fmorph_eq_rat _)); exists p. +by apply: complex_algebraic_trans; apply: realalg_algebraic. +Qed. diff --git a/mathcomp/real_closed/ordered_qelim.v b/mathcomp/real_closed/ordered_qelim.v new file mode 100644 index 0000000..c718d74 --- /dev/null +++ b/mathcomp/real_closed/ordered_qelim.v @@ -0,0 +1,1180 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div choice fintype. +Require Import bigop ssralg finset fingroup zmodp. +Require Import poly ssrnum. + + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. +Import GRing. + +Reserved Notation "p <% q" (at level 70, no associativity). +Reserved Notation "p <=% q" (at level 70, no associativity). + +(* Set Printing Width 30. *) + +Module ord. + +Section Formulas. + +Variable T : Type. + +Inductive formula : Type := +| Bool of bool +| Equal of (term T) & (term T) +| Lt of (term T) & (term T) +| Le of (term T) & (term T) +| Unit of (term T) +| And of formula & formula +| Or of formula & formula +| Implies of formula & formula +| Not of formula +| Exists of nat & formula +| Forall of nat & formula. + +End Formulas. + +Fixpoint term_eq (T : eqType)(t1 t2 : term T) := + match t1, t2 with + | Var n1, Var n2 => n1 == n2 + | Const r1, Const r2 => r1 == r2 + | NatConst n1, NatConst n2 => n1 == n2 + | Add r1 s1, Add r2 s2 => (term_eq r1 r2) && (term_eq s1 s2) + | Opp r1, Opp r2 => term_eq r1 r2 + | NatMul r1 n1, NatMul r2 n2 => (term_eq r1 r2) && (n1 == n2) + | Mul r1 s1, Mul r2 s2 => (term_eq r1 r2) && (term_eq s1 s2) + | Inv r1, Inv r2 => term_eq r1 r2 + | Exp r1 n1, Exp r2 n2 => (term_eq r1 r2) && (n1 == n2) + |_, _ => false + end. + +Lemma term_eqP (T : eqType) : Equality.axiom (@term_eq T). +Proof. +move=> t1 t2; apply: (iffP idP) => [|<-]; last first. + by elim: t1 {t2} => //= t -> // n; rewrite eqxx. +elim: t1 t2. +- by move=> n1 /= [] // n2 /eqP ->. +- by move=> r1 /= [] // r2 /eqP ->. +- by move=> n1 /= [] // n2 /eqP ->. +- by move=> r1 hr1 r2 hr2 [] //= s1 s2 /andP [] /hr1 -> /hr2 ->. +- by move=> r1 hr1 [] //= s1 /hr1 ->. +- by move=> s1 hs1 n1 [] //= s2 n2 /andP [] /hs1 -> /eqP ->. +- by move=> r1 hr1 r2 hr2 [] //= s1 s2 /andP [] /hr1 -> /hr2 ->. +- by move=> r1 hr1 [] //= s1 /hr1 ->. +- by move=> s1 hs1 n1 [] //= s2 n2 /andP [] /hs1 -> /eqP ->. +Qed. + +Canonical term_eqMixin (T : eqType) := EqMixin (@term_eqP T). +Canonical term_eqType (T : eqType) := + Eval hnf in EqType (term T) (@term_eqMixin T). + +Implicit Arguments term_eqP [x y]. +Prenex Implicits term_eq. + + +Bind Scope oterm_scope with term. +Bind Scope oterm_scope with formula. +Arguments Scope Add [_ oterm_scope oterm_scope]. +Arguments Scope Opp [_ oterm_scope]. +Arguments Scope NatMul [_ oterm_scope nat_scope]. +Arguments Scope Mul [_ oterm_scope oterm_scope]. +Arguments Scope Mul [_ oterm_scope oterm_scope]. +Arguments Scope Inv [_ oterm_scope]. +Arguments Scope Exp [_ oterm_scope nat_scope]. +Arguments Scope Equal [_ oterm_scope oterm_scope]. +Arguments Scope Unit [_ oterm_scope]. +Arguments Scope And [_ oterm_scope oterm_scope]. +Arguments Scope Or [_ oterm_scope oterm_scope]. +Arguments Scope Implies [_ oterm_scope oterm_scope]. +Arguments Scope Not [_ oterm_scope]. +Arguments Scope Exists [_ nat_scope oterm_scope]. +Arguments Scope Forall [_ nat_scope oterm_scope]. + +Implicit Arguments Bool [T]. +Prenex Implicits Const Add Opp NatMul Mul Exp Bool Unit And Or Implies Not. +Prenex Implicits Exists Forall Lt. + +Notation True := (Bool true). +Notation False := (Bool false). + +Delimit Scope oterm_scope with oT. +Notation "''X_' i" := (Var _ i) : oterm_scope. +Notation "n %:R" := (NatConst _ n) : oterm_scope. +Notation "x %:T" := (Const x) : oterm_scope. +Notation "0" := 0%:R%oT : oterm_scope. +Notation "1" := 1%:R%oT : oterm_scope. +Infix "+" := Add : oterm_scope. +Notation "- t" := (Opp t) : oterm_scope. +Notation "t - u" := (Add t (- u)) : oterm_scope. +Infix "*" := Mul : oterm_scope. +Infix "*+" := NatMul : oterm_scope. +Notation "t ^-1" := (Inv t) : oterm_scope. +Notation "t / u" := (Mul t u^-1) : oterm_scope. +Infix "^+" := Exp : oterm_scope. +Notation "t ^- n" := (t^-1 ^+ n)%oT : oterm_scope. +Infix "==" := Equal : oterm_scope. +Infix "<%" := Lt : oterm_scope. +Infix "<=%" := Le : oterm_scope. +Infix "/\" := And : oterm_scope. +Infix "\/" := Or : oterm_scope. +Infix "==>" := Implies : oterm_scope. +Notation "~ f" := (Not f) : oterm_scope. +Notation "x != y" := (Not (x == y)) : oterm_scope. +Notation "''exists' ''X_' i , f" := (Exists i f) : oterm_scope. +Notation "''forall' ''X_' i , f" := (Forall i f) : oterm_scope. + +Section Substitution. + +Variable T : Type. + + +Fixpoint fsubst (f : formula T) (s : nat * term T) := + match f with + | Bool _ => f + | (t1 == t2) => (tsubst t1 s == tsubst t2 s) + | (t1 <% t2) => (tsubst t1 s <% tsubst t2 s) + | (t1 <=% t2) => (tsubst t1 s <=% tsubst t2 s) + | (Unit t1) => Unit (tsubst t1 s) + | (f1 /\ f2) => (fsubst f1 s /\ fsubst f2 s) + | (f1 \/ f2) => (fsubst f1 s \/ fsubst f2 s) + | (f1 ==> f2) => (fsubst f1 s ==> fsubst f2 s) + | (~ f1) => (~ fsubst f1 s) + | ('exists 'X_i, f1) => ('exists 'X_i, if i == s.1 then f1 else fsubst f1 s) + | ('forall 'X_i, f1) => ('forall 'X_i, if i == s.1 then f1 else fsubst f1 s) + end%oT. + +End Substitution. + +Section OrderedClause. + +Inductive oclause (R : Type) : Type := + Oclause : seq (term R) -> seq (term R) -> seq (term R) -> seq (term R) -> oclause R. + +Definition eq_of_oclause (R : Type)(x : oclause R) := + let: Oclause y _ _ _ := x in y. +Definition neq_of_oclause (R : Type)(x : oclause R) := + let: Oclause _ y _ _ := x in y. +Definition lt_of_oclause (R : Type) (x : oclause R) := + let: Oclause _ _ y _ := x in y. +Definition le_of_oclause (R : Type) (x : oclause R) := + let: Oclause _ _ _ y := x in y. + +End OrderedClause. + +Delimit Scope oclause_scope with OCLAUSE. +Open Scope oclause_scope. + +Notation "p .1" := (@eq_of_oclause _ p) + (at level 2, left associativity, format "p .1") : oclause_scope. +Notation "p .2" := (@neq_of_oclause _ p) + (at level 2, left associativity, format "p .2") : oclause_scope. + +Notation "p .3" := (@lt_of_oclause _ p) + (at level 2, left associativity, format "p .3") : oclause_scope. +Notation "p .4" := (@le_of_oclause _ p) + (at level 2, left associativity, format "p .4") : oclause_scope. + +Definition oclause_eq (T : eqType)(t1 t2 : oclause T) := + let: Oclause eq_l1 neq_l1 lt_l1 leq_l1 := t1 in + let: Oclause eq_l2 neq_l2 lt_l2 leq_l2 := t2 in + [&& eq_l1 == eq_l2, neq_l1 == neq_l2, lt_l1 == lt_l2 & leq_l1 == leq_l2]. + +Lemma oclause_eqP (T : eqType) : Equality.axiom (@oclause_eq T). +Proof. +move=> t1 t2; apply: (iffP idP) => [|<-] /= ; last first. + by rewrite /oclause_eq; case: t1=> l1 l2 l3 l4; rewrite !eqxx. +case: t1 => [l1 l2 l3 l4]; case: t2 => m1 m2 m3 m4 /=; case/and4P. +by move/eqP=> -> /eqP -> /eqP -> /eqP ->. +Qed. + +Canonical oclause_eqMixin (T : eqType) := EqMixin (@oclause_eqP T). +Canonical oclause_eqType (T : eqType) := + Eval hnf in EqType (oclause T) (@oclause_eqMixin T). + +Implicit Arguments oclause_eqP [x y]. +Prenex Implicits oclause_eq. + +Section EvalTerm. + +Variable R : realDomainType. + +(* Evaluation of a reified formula *) + +Fixpoint holds (e : seq R) (f : ord.formula R) {struct f} : Prop := + match f with + | Bool b => b + | (t1 == t2)%oT => eval e t1 = eval e t2 + | (t1 <% t2)%oT => eval e t1 < eval e t2 + | (t1 <=% t2)%oT => eval e t1 <= eval e t2 + | Unit t1 => eval e t1 \in unit + | (f1 /\ f2)%oT => holds e f1 /\ holds e f2 + | (f1 \/ f2)%oT => holds e f1 \/ holds e f2 + | (f1 ==> f2)%oT => holds e f1 -> holds e f2 + | (~ f1)%oT => ~ holds e f1 + | ('exists 'X_i, f1)%oT => exists x, holds (set_nth 0 e i x) f1 + | ('forall 'X_i, f1)%oT => forall x, holds (set_nth 0 e i x) f1 + end. + + +(* Extensionality of formula evaluation *) +Lemma eq_holds e e' f : same_env e e' -> holds e f -> holds e' f. +Proof. +pose sv := set_nth (0 : R). +have eq_i i v e1 e2: same_env e1 e2 -> same_env (sv e1 i v) (sv e2 i v). + by move=> eq_e j; rewrite !nth_set_nth /= eq_e. +elim: f e e' => //=. +- by move=> t1 t2 e e' eq_e; rewrite !(eq_eval _ eq_e). +- by move=> t1 t2 e e' eq_e; rewrite !(eq_eval _ eq_e). +- by move=> t1 t2 e e' eq_e; rewrite !(eq_eval _ eq_e). +- by move=> t e e' eq_e; rewrite (eq_eval _ eq_e). +- by move=> f1 IH1 f2 IH2 e e' eq_e; move/IH2: (eq_e); move/IH1: eq_e; tauto. +- by move=> f1 IH1 f2 IH2 e e' eq_e; move/IH2: (eq_e); move/IH1: eq_e; tauto. +- by move=> f1 IH1 f2 IH2 e e' eq_e f12; move/IH1: (same_env_sym eq_e); eauto. +- by move=> f1 IH1 e e'; move/same_env_sym; move/IH1; tauto. +- by move=> i f1 IH1 e e'; move/(eq_i i)=> eq_e [x f_ex]; exists x; eauto. +by move=> i f1 IH1 e e'; move/(eq_i i); eauto. +Qed. + +(* Evaluation and substitution by a constant *) +Lemma holds_fsubst e f i v : + holds e (fsubst f (i, v%:T)%T) <-> holds (set_nth 0 e i v) f. +Proof. +elim: f e => //=; do [ + by move=> *; rewrite !eval_tsubst +| move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto +| move=> f IHf e; move: (IHf e); tauto +| move=> j f IHf e]. +- case eq_ji: (j == i); first rewrite (eqP eq_ji). + by split=> [] [x f_x]; exists x; rewrite set_set_nth eqxx in f_x *. + split=> [] [x f_x]; exists x; move: f_x; rewrite set_set_nth eq_sym eq_ji; + have:= IHf (set_nth 0 e j x); tauto. +case eq_ji: (j == i); first rewrite (eqP eq_ji). + by split=> [] f_ x; move: (f_ x); rewrite set_set_nth eqxx. +split=> [] f_ x; move: (IHf (set_nth 0 e j x)) (f_ x); + by rewrite set_set_nth eq_sym eq_ji; tauto. +Qed. + +(* Boolean test selecting formulas in the theory of rings *) +Fixpoint rformula (f : formula R) := + match f with + | Bool _ => true + | t1 == t2 => rterm t1 && rterm t2 + | t1 <% t2 => rterm t1 && rterm t2 + | t1 <=% t2 => rterm t1 && rterm t2 + | Unit t1 => false + | (f1 /\ f2) | (f1 \/ f2) | (f1 ==> f2) => rformula f1 && rformula f2 + | (~ f1) | ('exists 'X__, f1) | ('forall 'X__, f1) => rformula f1 + end%oT. + +(* An oformula stating that t1 is equal to 0 in the ring theory. *) +Definition eq0_rform t1 := + let m := @ub_var R t1 in + let: (t1', r1) := to_rterm t1 [::] m in + let fix loop r i := match r with + | [::] => t1' == 0 + | t :: r' => + let f := ('X_i * t == 1 /\ t * 'X_i == 1) in + 'forall 'X_i, (f \/ 'X_i == t /\ ~ ('exists 'X_i, f)) ==> loop r' i.+1 + end%oT + in loop r1 m. + +(* An oformula stating that t1 is less than 0 in the equational ring theory. +Definition leq0_rform t1 := + let m := @ub_var R t1 in + let: (t1', r1) := to_rterm t1 [::] m in + let fix loop r i := match r with + | [::] => 'exists 'X_m.+1, t1' == 'X_m.+1 * 'X_m.+1 + | t :: r' => + let f := ('X_i * t == 1 /\ t * 'X_i == 1) in + 'forall 'X_i, (f \/ 'X_i == t /\ ~ ('exists 'X_i, f)) ==> loop r' i.+1 + end%oT + in loop r1 m. +*) +Definition leq0_rform t1 := + let m := @ub_var R t1 in + let: (t1', r1) := to_rterm t1 [::] m in + let fix loop r i := match r with + | [::] => t1' <=% 0 + | t :: r' => + let f := ('X_i * t == 1 /\ t * 'X_i == 1) in + 'forall 'X_i, (f \/ 'X_i == t /\ ~ ('exists 'X_i, f)) ==> loop r' i.+1 + end%oT + in loop r1 m. + + +(* Definition lt0_rform t1 := *) +(* let m := @ub_var R t1 in *) +(* let: (t1', r1) := to_rterm t1 [::] m in *) +(* let fix loop r i := match r with *) +(* | [::] => 'exists 'X_m.+1, (t1' == 'X_m.+1 * 'X_m.+1 /\ ~ ('X_m.+1 == 0)) *) +(* | t :: r' => *) +(* let f := ('X_i * t == 1 /\ t * 'X_i == 1) in *) +(* 'forall 'X_i, (f \/ 'X_i == t /\ ~ ('exists 'X_i, f)) ==> loop r' i.+1 *) +(* end%oT *) +(* in loop r1 m. *) + +Definition lt0_rform t1 := + let m := @ub_var R t1 in + let: (t1', r1) := to_rterm t1 [::] m in + let fix loop r i := match r with + | [::] => t1' <% 0 + | t :: r' => + let f := ('X_i * t == 1 /\ t * 'X_i == 1) in + 'forall 'X_i, (f \/ 'X_i == t /\ ~ ('exists 'X_i, f)) ==> loop r' i.+1 + end%oT + in loop r1 m. + +(* Transformation of a formula in the theory of rings with units into an *) + +(* equivalent formula in the sub-theory of rings. *) +Fixpoint to_rform f := + match f with + | Bool b => f + | t1 == t2 => eq0_rform (t1 - t2) + | t1 <% t2 => lt0_rform (t1 - t2) + | t1 <=% t2 => leq0_rform (t1 - t2) + | Unit t1 => eq0_rform (t1 * t1^-1 - 1) + | f1 /\ f2 => to_rform f1 /\ to_rform f2 + | f1 \/ f2 => to_rform f1 \/ to_rform f2 + | f1 ==> f2 => to_rform f1 ==> to_rform f2 + | ~ f1 => ~ to_rform f1 + | ('exists 'X_i, f1) => 'exists 'X_i, to_rform f1 + | ('forall 'X_i, f1) => 'forall 'X_i, to_rform f1 + end%oT. + +(* The transformation gives a ring formula. *) +(* the last part of the proof consists in 3 cases that are exactly the same. + how to factorize ? *) +Lemma to_rform_rformula f : rformula (to_rform f). +Proof. +suffices [h1 h2 h3]: + [/\ forall t1, rformula (eq0_rform t1), + forall t1, rformula (lt0_rform t1) & + forall t1, rformula (leq0_rform t1)]. + by elim: f => //= => f1 ->. +split=> t1. +- rewrite /eq0_rform; move: (ub_var t1) => m. + set tr := _ m. + suffices: all (@rterm R) (tr.1 :: tr.2)%PAIR. + case: tr => {t1} t1 r /= /andP[t1_r]. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + have: all (@rterm R) [::] by []. + rewrite {}/tr; elim: t1 [::] => //=. + + move=> t1 IHt1 t2 IHt2 r. + move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r]. + move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r]. + by rewrite t1_r t2_r. + + by move=> t1 IHt1 r /IHt1; case: to_rterm. + + by move=> t1 IHt1 n r /IHt1; case: to_rterm. + + move=> t1 IHt1 t2 IHt2 r. + move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r]. + move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r]. + by rewrite t1_r t2_r. + + move=> t1 IHt1 r. + by move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /=; rewrite all_rcons. + + by move=> t1 IHt1 n r /IHt1; case: to_rterm. +- rewrite /lt0_rform; move: (ub_var t1) => m; set tr := _ m. + suffices: all (@rterm R) (tr.1 :: tr.2)%PAIR. + case: tr => {t1} t1 r /= /andP[t1_r]. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + have: all (@rterm R) [::] by []. + rewrite {}/tr; elim: t1 [::] => //=. + + move=> t1 IHt1 t2 IHt2 r. + move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r]. + move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r]. + by rewrite t1_r t2_r. + + by move=> t1 IHt1 r /IHt1; case: to_rterm. + + by move=> t1 IHt1 n r /IHt1; case: to_rterm. + + move=> t1 IHt1 t2 IHt2 r. + move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r]. + move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r]. + by rewrite t1_r t2_r. + + move=> t1 IHt1 r. + by move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /=; rewrite all_rcons. + + by move=> t1 IHt1 n r /IHt1; case: to_rterm. +- rewrite /leq0_rform; move: (ub_var t1) => m; set tr := _ m. + suffices: all (@rterm R) (tr.1 :: tr.2)%PAIR. + case: tr => {t1} t1 r /= /andP[t1_r]. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + have: all (@rterm R) [::] by []. + rewrite {}/tr; elim: t1 [::] => //=. + + move=> t1 IHt1 t2 IHt2 r. + move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r]. + move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r]. + by rewrite t1_r t2_r. + + by move=> t1 IHt1 r /IHt1; case: to_rterm. + + by move=> t1 IHt1 n r /IHt1; case: to_rterm. + + move=> t1 IHt1 t2 IHt2 r. + move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /= /andP[t1_r]. + move/IHt2; case: to_rterm => {t2 r IHt2} t2 r /= /andP[t2_r]. + by rewrite t1_r t2_r. + + move=> t1 IHt1 r. + by move/IHt1; case: to_rterm => {t1 r IHt1} t1 r /=; rewrite all_rcons. + + by move=> t1 IHt1 n r /IHt1; case: to_rterm. +Qed. + +Import Num.Theory. + +(* Correctness of the transformation. *) +Lemma to_rformP e f : holds e (to_rform f) <-> holds e f. +Proof. +suffices{e f} [equal0_equiv lt0_equiv le0_equiv]: + [/\ forall e t1 t2, holds e (eq0_rform (t1 - t2)) <-> (eval e t1 == eval e t2), + forall e t1 t2, holds e (lt0_rform (t1 - t2)) <-> (eval e t1 < eval e t2) & + forall e t1 t2, holds e (leq0_rform (t1 - t2)) <-> (eval e t1 <= eval e t2)]. +- elim: f e => /=; try tauto. + + move=> t1 t2 e. + by split; [move/equal0_equiv/eqP | move/eqP/equal0_equiv]. + + by move=> t1 t2 e; split; move/lt0_equiv. + + by move=> t1 t2 e; split; move/le0_equiv. + + move=> t1 e; rewrite unitrE; exact: equal0_equiv. + + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + move=> f1 IHf1 e; move: (IHf1 e); tauto. + + by move=> n f1 IHf1 e; split=> [] [x] /IHf1; exists x. + + by move=> n f1 IHf1 e; split=> Hx x; apply/IHf1. +suffices h e t1 t2 : + [/\ holds e (eq0_rform (t1 - t2)) <-> (eval e t1 == eval e t2), + holds e (lt0_rform (t1 - t2)) <-> (eval e t1 < eval e t2) & + holds e (leq0_rform (t1 - t2)) <-> (eval e t1 <= eval e t2)]. + by split => e t1 t2; case: (h e t1 t2). +rewrite -{1}(add0r (eval e t2)) -(can2_eq (subrK _) (addrK _)). +rewrite -subr_lt0 -subr_le0 -/(eval e (t1 - t2)); move: (t1 - t2)%T => {t1 t2} t. +have sub_var_tsubst s t0: (s.1%PAIR >= ub_var t0)%N -> tsubst t0 s = t0. + elim: t0 {t} => //=. + - by move=> n; case: ltngtP. + - by move=> t1 IHt1 t2 IHt2; rewrite geq_max => /andP[/IHt1-> /IHt2->]. + - by move=> t1 IHt1 /IHt1->. + - by move=> t1 IHt1 n /IHt1->. + - by move=> t1 IHt1 t2 IHt2; rewrite geq_max => /andP[/IHt1-> /IHt2->]. + - by move=> t1 IHt1 /IHt1->. + - by move=> t1 IHt1 n /IHt1->. +pose fix rsub t' m r : term R := + if r is u :: r' then tsubst (rsub t' m.+1 r') (m, u^-1)%T else t'. +pose fix ub_sub m r : Prop := + if r is u :: r' then (ub_var u <= m)%N /\ ub_sub m.+1 r' else true. +suffices{t} rsub_to_r t r0 m: (m >= ub_var t)%N -> ub_sub _ m r0 -> + let: (t', r) := to_rterm t r0 m in + [/\ take (size r0) r = r0, + ub_var t' <= m + size r, ub_sub _ m r & rsub t' m r = t]%N. +- have:= rsub_to_r t [::] _ (leqnn _); rewrite /eq0_rform /lt0_rform /leq0_rform. + case: (to_rterm _ _ _) => [t1' r1] /= [//| _ _ ub_r1 def_t]. + rewrite -{2 4 6}def_t {def_t}. + elim: r1 (ub_var t) e ub_r1 => [|u r1 IHr1] m e /= => [_|[ub_u ub_r1]]. + by split => //; split=> /eqP. + rewrite eval_tsubst /=; set y := eval e u; split; split => //= t_h0. + + case: (IHr1 m.+1 (set_nth 0 e m y^-1) ub_r1) => h _ _; apply/h. + apply: t_h0. + rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)). + rewrite sub_var_tsubst //= -/y. + case Uy: (y \in unit); [left | right]; first by rewrite mulVr ?divrr. + split=> [|[z]]; first by rewrite invr_out ?Uy. + rewrite nth_set_nth /= eqxx. + rewrite -!(eval_tsubst _ _ (m, Const _)) !sub_var_tsubst // -/y => yz1. + by case/unitrP: Uy; exists z. + + move=> x def_x. + case: (IHr1 m.+1 (set_nth 0 e m x) ub_r1) => h _ _. apply/h. + suff ->: x = y^-1 by []; move: def_x. + rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)). + rewrite sub_var_tsubst //= -/y; case=> [[xy1 yx1] | [xy nUy]]. + by rewrite -[y^-1]mul1r -[1]xy1 mulrK //; apply/unitrP; exists x. + rewrite invr_out //; apply/unitrP=> [[z yz1]]; case: nUy; exists z. + rewrite nth_set_nth /= eqxx -!(eval_tsubst _ _ (m, _%:T)%T). + by rewrite !sub_var_tsubst. + + case: (IHr1 m.+1 (set_nth 0 e m y^-1) ub_r1) => _ h _. apply/h. + apply: t_h0. + rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)). + rewrite sub_var_tsubst //= -/y. + case Uy: (y \in unit); [left | right]; first by rewrite mulVr ?divrr. + split=> [|[z]]; first by rewrite invr_out ?Uy. + rewrite nth_set_nth /= eqxx. + rewrite -!(eval_tsubst _ _ (m, Const _)) !sub_var_tsubst // -/y => yz1. + by case/unitrP: Uy; exists z. + + move=> x def_x. + case: (IHr1 m.+1 (set_nth 0 e m x) ub_r1) => _ h _. apply/h. + suff ->: x = y^-1 by []; move: def_x. + rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)). + rewrite sub_var_tsubst //= -/y; case=> [[xy1 yx1] | [xy nUy]]. + by rewrite -[y^-1]mul1r -[1]xy1 mulrK //; apply/unitrP; exists x. + rewrite invr_out //; apply/unitrP=> [[z yz1]]; case: nUy; exists z. + rewrite nth_set_nth /= eqxx -!(eval_tsubst _ _ (m, _%:T)%T). + by rewrite !sub_var_tsubst. + + case: (IHr1 m.+1 (set_nth 0 e m y^-1) ub_r1) => _ _ h. apply/h. + apply: t_h0. + rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)). + rewrite sub_var_tsubst //= -/y. + case Uy: (y \in unit); [left | right]; first by rewrite mulVr ?divrr. + split=> [|[z]]; first by rewrite invr_out ?Uy. + rewrite nth_set_nth /= eqxx. + rewrite -!(eval_tsubst _ _ (m, Const _)) !sub_var_tsubst // -/y => yz1. + by case/unitrP: Uy; exists z. + + move=> x def_x. + case: (IHr1 m.+1 (set_nth 0 e m x) ub_r1) => _ _ h. apply/h. + suff ->: x = y^-1 by []; move: def_x. + rewrite nth_set_nth /= eqxx -(eval_tsubst e u (m, Const _)). + rewrite sub_var_tsubst //= -/y; case=> [[xy1 yx1] | [xy nUy]]. + by rewrite -[y^-1]mul1r -[1]xy1 mulrK //; apply/unitrP; exists x. + rewrite invr_out //; apply/unitrP=> [[z yz1]]; case: nUy; exists z. + rewrite nth_set_nth /= eqxx -!(eval_tsubst _ _ (m, _%:T)%T). + by rewrite !sub_var_tsubst. +have rsub_id r t0 n: (ub_var t0 <= n)%N -> rsub t0 n r = t0. + by elim: r n => //= t1 r IHr n let0n; rewrite IHr ?sub_var_tsubst ?leqW. +have rsub_acc r s t1 m1: + (ub_var t1 <= m1 + size r)%N -> rsub t1 m1 (r ++ s) = rsub t1 m1 r. + elim: r t1 m1 => [|t1 r IHr] t2 m1 /=; first by rewrite addn0; apply: rsub_id. + by move=> letmr; rewrite IHr ?addSnnS. +elim: t r0 m => /=; try do [ + by move=> n r m hlt hub; rewrite take_size (ltn_addr _ hlt) rsub_id +| by move=> n r m hlt hub; rewrite leq0n take_size rsub_id +| move=> t1 IHt1 t2 IHt2 r m; rewrite geq_max; case/andP=> hub1 hub2 hmr; + case: to_rterm {IHt1 hub1 hmr}(IHt1 r m hub1 hmr) => t1' r1; + case=> htake1 hub1' hsub1 <-; + case: to_rterm {IHt2 hub2 hsub1}(IHt2 r1 m hub2 hsub1) => t2' r2 /=; + rewrite geq_max; case=> htake2 -> hsub2 /= <-; + rewrite -{1 2}(cat_take_drop (size r1) r2) htake2; set r3 := drop _ _; + rewrite size_cat addnA (leq_trans _ (leq_addr _ _)) //; + split=> {hsub2}//; + first by [rewrite takel_cat // -htake1 size_take geq_minr]; + rewrite -(rsub_acc r1 r3 t1') {hub1'}// -{htake1}htake2 {r3}cat_take_drop; + by elim: r2 m => //= u r2 IHr2 m; rewrite IHr2 +| do [ move=> t1 IHt1 r m; do 2!move/IHt1=> {IHt1}IHt1 + | move=> t1 IHt1 n r m; do 2!move/IHt1=> {IHt1}IHt1]; + case: to_rterm IHt1 => t1' r1 [-> -> hsub1 <-]; split=> {hsub1}//; + by elim: r1 m => //= u r1 IHr1 m; rewrite IHr1]. +move=> t1 IH r m letm /IH {IH} /(_ letm) {letm}. +case: to_rterm => t1' r1 /= [def_r ub_t1' ub_r1 <-]. +rewrite size_rcons addnS leqnn -{1}cats1 takel_cat ?def_r; last first. + by rewrite -def_r size_take geq_minr. +elim: r1 m ub_r1 ub_t1' {def_r} => /= [|u r1 IHr1] m => [_|[->]]. + by rewrite addn0 eqxx. +by rewrite -addSnnS => /IHr1 IH /IH[_ _ ub_r1 ->]. +Qed. + +(* The above proof is ugly but is in fact copypaste *) + +(* Boolean test selecting formulas which describe a constructible set, *) +(* i.e. formulas without quantifiers. *) + +(* The quantifier elimination check. *) +Fixpoint qf_form (f : formula R) := + match f with + | Bool _ | _ == _ | Unit _ | Lt _ _ | Le _ _ => true + | f1 /\ f2 | f1 \/ f2 | f1 ==> f2 => qf_form f1 && qf_form f2 + | ~ f1 => qf_form f1 + | _ => false + end%oT. + +(* Boolean holds predicate for quantifier free formulas *) +Definition qf_eval e := fix loop (f : formula R) : bool := + match f with + | Bool b => b + | t1 == t2 => (eval e t1 == eval e t2)%bool + | t1 <% t2 => (eval e t1 < eval e t2)%bool + | t1 <=% t2 => (eval e t1 <= eval e t2)%bool + | Unit t1 => eval e t1 \in unit + | f1 /\ f2 => loop f1 && loop f2 + | f1 \/ f2 => loop f1 || loop f2 + | f1 ==> f2 => (loop f1 ==> loop f2)%bool + | ~ f1 => ~~ loop f1 + |_ => false + end%oT. + +(* qf_eval is equivalent to holds *) +Lemma qf_evalP e f : qf_form f -> reflect (holds e f) (qf_eval e f). +Proof. +elim: f => //=; try by move=> *; exact: idP. +- move=> t1 t2 _; exact: eqP. +- move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1T]; last by right; case. + by case/IHf2; [left | right; case]. +- move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1F]; first by do 2 left. + by case/IHf2; [left; right | right; case]. +- move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1T]; last by left. + by case/IHf2; [left | right; move/(_ f1T)]. +by move=> f1 IHf1 /IHf1[]; [right | left]. +Qed. + +(* Quantifier-free formula are normalized into DNF. A DNF is *) +(* represented by the type seq (seq (term R) * seq (term R)), where we *) +(* separate positive and negative literals *) + + +(* DNF preserving conjunction *) + +Definition and_odnf (bcs1 bcs2 : seq (oclause R)) := + \big[cat/nil]_(bc1 <- bcs1) + map (fun bc2 : oclause R => + (Oclause (bc1.1 ++ bc2.1) (bc1.2 ++ bc2.2) (bc1.3 ++ bc2.3) (bc1.4 ++ bc2.4)))%OCLAUSE bcs2. + +(* Computes a DNF from a qf ring formula *) +Fixpoint qf_to_odnf (f : formula R) (neg : bool) {struct f} : seq (oclause R) := + match f with + | Bool b => if b (+) neg then [:: (Oclause [::] [::] [::] [::])] else [::] + | t1 == t2 => + [:: if neg then (Oclause [::] [:: t1 - t2] [::] [::]) else (Oclause [:: t1 - t2] [::] [::] [::])] + | t1 <% t2 => + [:: if neg then (Oclause [::] [::] [::] [:: t1 - t2]) else (Oclause [::] [::] [:: t2 - t1] [::])] + | t1 <=% t2 => + [:: if neg then (Oclause [::] [::] [:: t1 - t2] [::]) else (Oclause [::] [::] [::] [:: t2 - t1])] + | f1 /\ f2 => (if neg then cat else and_odnf) [rec f1, neg] [rec f2, neg] + | f1 \/ f2 => (if neg then and_odnf else cat) [rec f1, neg] [rec f2, neg] + | f1 ==> f2 => (if neg then and_odnf else cat) [rec f1, ~~ neg] [rec f2, neg] + | ~ f1 => [rec f1, ~~ neg] + | _ => if neg then [:: (Oclause [::] [::] [::] [::])] else [::] + end%oT where "[ 'rec' f , neg ]" := (qf_to_odnf f neg). + +(* Conversely, transforms a DNF into a formula *) +Definition odnf_to_oform := + let pos_lit t := And (t == 0)%oT in let neg_lit t := And (t != 0)%oT in + let lt_lit t := And (0 <% t)%oT in let le_lit t := And (0 <=% t)%oT in + let ocls (bc : oclause R) := + Or + (foldr pos_lit True bc.1 /\ foldr neg_lit True bc.2 /\ + foldr lt_lit True bc.3 /\ foldr le_lit True bc.4) in + foldr ocls False. + +(* Catenation of dnf is the Or of formulas *) +Lemma cat_dnfP e bcs1 bcs2 : + qf_eval e (odnf_to_oform (bcs1 ++ bcs2)) + = qf_eval e (odnf_to_oform bcs1 \/ odnf_to_oform bcs2). +Proof. +by elim: bcs1 => //= bc1 bcs1 IH1; rewrite -orbA; congr orb; rewrite IH1. +Qed. + + + +(* and_dnf is the And of formulas *) +Lemma and_odnfP e bcs1 bcs2 : + qf_eval e (odnf_to_oform (and_odnf bcs1 bcs2)) + = qf_eval e (odnf_to_oform bcs1 /\ odnf_to_oform bcs2). +Proof. +elim: bcs1 => [|bc1 bcs1 IH1] /=; first by rewrite /and_odnf big_nil. +rewrite /and_odnf big_cons -/(and_odnf bcs1 bcs2) cat_dnfP /=. +rewrite {}IH1 /= andb_orl; congr orb. +elim: bcs2 bc1 {bcs1} => [|bc2 bcs2 IH] bc1 /=; first by rewrite andbF. +rewrite {}IH /= andb_orr; congr orb => {bcs2}. +suffices aux (l1 l2 : seq (term R)) g : let redg := foldr (And \o g) True in + qf_eval e (redg (l1 ++ l2)) = qf_eval e (redg l1 /\ redg l2)%oT. ++ rewrite !aux /= !andbA; congr (_ && _); rewrite -!andbA; congr (_ && _). + rewrite -andbCA; congr (_ && _); bool_congr; rewrite andbCA; bool_congr. + by rewrite andbA andbC !andbA. +by elim: l1 => [| t1 l1 IHl1] //=; rewrite -andbA IHl1. +Qed. + +Lemma qf_to_dnfP e : + let qev f b := qf_eval e (odnf_to_oform (qf_to_odnf f b)) in + forall f, qf_form f && rformula f -> qev f false = qf_eval e f. +Proof. +move=> qev; have qevT f: qev f true = ~~ qev f false. + rewrite {}/qev; elim: f => //=; do [by case | move=> f1 IH1 f2 IH2 | ]. + - by move=> t1 t2; rewrite !andbT !orbF. + - by move=> t1 t2; rewrite !andbT !orbF; rewrite !subr_gte0 -lerNgt. + - by move=> t1 t2; rewrite !andbT !orbF; rewrite !subr_gte0 -ltrNge. + - by rewrite and_odnfP cat_dnfP negb_and -IH1 -IH2. + - by rewrite and_odnfP cat_dnfP negb_or -IH1 -IH2. + - by rewrite and_odnfP cat_dnfP /= negb_or IH1 -IH2 negbK. + by move=> t1 ->; rewrite negbK. +rewrite /qev; elim=> //=; first by case. +- by move=> t1 t2 _; rewrite subr_eq0 !andbT orbF. +- by move=> t1 t2 _; rewrite orbF !andbT subr_gte0. +- by move=> t1 t2 _; rewrite orbF !andbT subr_gte0. +- move=> f1 IH1 f2 IH2; rewrite andbCA -andbA andbCA andbA; case/andP. + by rewrite and_odnfP /= => /IH1-> /IH2->. +- move=> f1 IH1 f2 IH2; rewrite andbCA -andbA andbCA andbA; case/andP. + by rewrite cat_dnfP /= => /IH1-> => /IH2->. +- move=> f1 IH1 f2 IH2; rewrite andbCA -andbA andbCA andbA; case/andP. + by rewrite cat_dnfP /= [qf_eval _ _]qevT -implybE => /IH1 <- /IH2->. +by move=> f1 IH1 /IH1 <-; rewrite -qevT. +Qed. + +Lemma dnf_to_form_qf bcs : qf_form (odnf_to_oform bcs). +Proof. +elim: bcs => //= [[clT clF] clLt clLe ? ->] /=; elim: clT => //=. +by rewrite andbT; elim: clF; elim: clLt => //; elim: clLe. +Qed. + +Definition dnf_rterm (cl : oclause R) := + [&& all (@rterm R) cl.1, all (@rterm R) cl.2, + all (@rterm R) cl.3 & all (@rterm R) cl.4]. + +Lemma qf_to_dnf_rterm f b : rformula f -> all dnf_rterm (qf_to_odnf f b). +Proof. +set ok := all dnf_rterm. +have cat_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (bcs1 ++ bcs2). + by move=> ok1 ok2; rewrite [ok _]all_cat; exact/andP. +have and_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (and_odnf bcs1 bcs2). + rewrite /and_odnf unlock; elim: bcs1 => //= cl1 bcs1 IH1; rewrite -andbA. + case/and3P=> ok11 ok12 ok1 ok2; rewrite cat_ok ?{}IH1 {bcs1 ok1}//. + elim: bcs2 ok2 => //= cl2 bcs2 IH2 /andP[ok2 /IH2->]. + by rewrite /dnf_rterm /= !all_cat andbT ok11; case/and3P: ok12=> -> -> ->. +elim: f b => //=; try by [move=> _ ? ? [] | move=> ? ? ? ? [] /= /andP[]; auto]. +- by do 2!case. +- by rewrite /dnf_rterm => ? ? [] /= ->. +- by rewrite /dnf_rterm => ? ? [] /=; rewrite andbC !andbT. +- by rewrite /dnf_rterm => ? ? [] /=; rewrite andbC !andbT. +by auto. +Qed. + +Lemma dnf_to_rform bcs : rformula (odnf_to_oform bcs) = all dnf_rterm bcs. +Proof. +elim: bcs => //= [[cl1 cl2 cl3 cl4] bcs ->]; rewrite {2}/dnf_rterm /=; congr (_ && _). +congr andb; first by elim: cl1 => //= t cl ->; rewrite andbT. +congr andb; first by elim: cl2 => //= t cl ->; rewrite andbT. +congr andb; first by elim: cl3 => //= t cl ->. +by elim: cl4 => //= t cl ->. +Qed. + +Implicit Type f : formula R. + +Fixpoint leq_elim_aux (eq_l lt_l le_l : seq (term R)) := + match le_l with + [::] => [:: (eq_l, lt_l)] + |le1 :: le_l' => + let res := leq_elim_aux eq_l lt_l le_l' in + let as_eq := map (fun x => (le1 :: x.1%PAIR, x.2%PAIR)) res in + let as_lt := map (fun x => (x.1%PAIR, le1 :: x.2%PAIR)) res in + as_eq ++ as_lt + end. + +Definition oclause_leq_elim oc : seq (oclause R) := + let: Oclause eq_l neq_l lt_l le_l := oc in + map (fun x => Oclause x.1%PAIR neq_l x.2%PAIR [::]) + (leq_elim_aux eq_l lt_l le_l). + +Definition terms_of_oclause (oc : oclause R) := + let: Oclause eq_l neq_l lt_l le_l := oc in + eq_l ++ neq_l ++ lt_l ++ le_l. + +Lemma terms_of_leq_elim oc1 oc2: + oc2 \in (oclause_leq_elim oc1) -> + (terms_of_oclause oc2) =i (terms_of_oclause oc1). +case: oc1 => eq1 neq1 lt1 leq1 /=. +elim: leq1 eq1 lt1 oc2 => [|t1 leq1 ih] eq1 lt1 [eq2 neq2 lt2 leq2] /=. + by rewrite inE; case/eqP=> -> -> -> -> ?. +rewrite map_cat /= mem_cat -!map_comp; set f := fun _ => _. +rewrite -/f in ih; case/orP. + case/mapP=> [[y1 y2]] yin ye. + move: (ih eq1 lt1 (f (y1, y2))); rewrite mem_map //; last first. + by move=> [u1 u2] [v1 v2]; rewrite /f /=; case=> -> ->. + move/(_ yin); move: ye; rewrite /f /=; case=> -> -> -> -> /= h. + move=> u; rewrite in_cons (h u) !mem_cat in_cons. + by rewrite orbC !orbA; set x := _ || (u \in lt1); rewrite orbAC. +case/mapP=> [[y1 y2]] yin ye. +move: (ih eq1 lt1 (f (y1, y2))); rewrite mem_map //; last first. + by move=> [u1 u2] [v1 v2]; rewrite /f /=; case=> -> ->. +move/(_ yin); move: ye; rewrite /f /=; case=> -> -> -> -> /= h u. +rewrite !mem_cat !in_cons orbA orbCA -!orbA; move: (h u); rewrite !mem_cat=> ->. +by rewrite orbC !orbA; set x := _ || (u \in lt1); rewrite orbAC. +Qed. + +Lemma odnf_to_oform_cat e c d : holds e (odnf_to_oform (c ++ d)) + <-> holds e ((odnf_to_oform c) \/ (odnf_to_oform d))%oT. +Proof. +elim: c d => [| tc c ihc] d /=; first by split => // hd; [right | case: hd]. +rewrite ihc /=; split. + case; first by case=> ?; case => ?; case => ? ?; left; left. + case; first by move=> ?; left; right. + by move=> ?; right. +case; last by move=> ?; right; right. +case; last by move=> ?; right; left. +by do 3!case=> ?; move=> ?; left. +Qed. + +Lemma oclause_leq_elimP oc e : + holds e (odnf_to_oform [:: oc]) <-> + holds e (odnf_to_oform (oclause_leq_elim oc)). +Proof. +case: oc => eq_l neq_l lt_l le_l; rewrite /oclause_leq_elim. +elim: le_l eq_l neq_l lt_l => [|t le_l ih] eq_l neq_l lt_l //=. +move: (ih eq_l neq_l lt_l) => /= {ih}. +set x1 := foldr _ _ _; set x2 := foldr _ _ _; set x3 := foldr _ _ _. +set x4 := foldr _ _ _ => h. +have -> : (holds e x1 /\ holds e x2 /\ holds e x3 /\ 0%:R <= eval e t /\ + holds e x4 \/ false) <-> + (0%:R <= eval e t) /\ (holds e x1 /\ holds e x2 /\ holds e x3 /\ + holds e x4 \/ false). + split; first by case=> //; do 4! (case => ?); move=> ?; split => //; left. + by case=> ?; case => //; do 3! (case=> ?); move=> ?; left. +rewrite h {h} /= !map_cat /= -!map_comp. +set s1 := [seq _ | _ <- _]; set s2 := [seq _ | _ <- _]. +set s3 := [seq _ | _ <- _]. rewrite odnf_to_oform_cat. +suff {x1 x2 x3 x4} /= -> : + holds e (odnf_to_oform s2) <-> eval e t == 0%:R /\ holds e (odnf_to_oform s1). + suff /= -> : + holds e (odnf_to_oform s3) <-> 0%:R < eval e t /\ holds e (odnf_to_oform s1). + rewrite ler_eqVlt eq_sym; split; first by case; case/orP=> -> ?; [left|right]. + by case; [case=> -> ? /= |case=> ->; rewrite orbT]. + rewrite /s1 /s3. + elim: (leq_elim_aux eq_l lt_l le_l) => /= [| t1 l ih]; first by split=> // [[]]. + rewrite /= ih; split. + case; last by case=> -> ?; split=> //; right. + by do 2!case=> ?; case; case=> -> ? _; split => //; auto. + by case=> ->; case; [do 3!case=> ?; move=> _; left | right]. +rewrite /s2 /s1. +elim: (leq_elim_aux eq_l lt_l le_l) => /= [| t1 l ih]; first by split=> // [[]]. +rewrite /= ih; split. + case; last by case=> -> ?; split=> //; right. + by case; case=> /eqP ? ?; do 2! case=> ?; move=> _; split=>//; left. +case=> /eqP ?; case; first by do 3!case=> ?; move=> _; left. +by right; split=> //; apply/eqP. +Qed. + +Fixpoint neq_elim_aux (lt_l neq_l : seq (term R)) := + match neq_l with + [::] => [:: lt_l] + |neq1 :: neq_l' => + let res := neq_elim_aux lt_l neq_l' in + let as_pos := map (fun x => neq1 :: x) res in + let as_neg := map (fun x => Opp neq1 :: x) res in + as_pos ++ as_neg + end. + +Definition oclause_neq_elim oc : seq (oclause R) := + let: Oclause eq_l neq_l lt_l le_l := oc in + map (fun x => Oclause eq_l [::] x le_l) (neq_elim_aux lt_l neq_l). + +Lemma terms_of_neq_elim oc1 oc2: + oc2 \in (oclause_neq_elim oc1) -> + {subset (terms_of_oclause oc2) <= (terms_of_oclause oc1) ++ (map Opp oc1.2)}. +Proof. +case: oc1 => eq1 neq1 lt1 leq1 /=. +elim: neq1 lt1 oc2 => [|t1 neq1 ih] lt1 [eq2 neq2 lt2 leq2] /=. + by rewrite inE; case/eqP=> -> -> -> ->; rewrite !cats0 !cat0s. +rewrite map_cat /= mem_cat -!map_comp; set f := fun _ => _. +rewrite -/f in ih; case/orP. + case/mapP=> y yin ye. + move: (ih lt1 (f y)); rewrite mem_map //; last first. + by move=> u v; rewrite /f /=; case. + move/(_ yin); move: ye; rewrite /f /=; case=> -> -> -> -> /= h. + move=> u. rewrite !mem_cat !in_cons orbAC orbC mem_cat -!orbA. + case/orP; first by move->; rewrite !orbT. + rewrite !orbA [_ || (_ \in eq1)]orbC; move: (h u); rewrite !mem_cat=> hu. + by move/hu; do 2! (case/orP; last by move->; rewrite !orbT); move->. +case/mapP=> y yin ye. +move: (ih lt1 (f y)); rewrite mem_map //; last first. + by move=> u v; rewrite /f /=; case. +move/(_ yin); move: ye; rewrite /f /=; case=> -> -> -> -> /= h. +move=> u; rewrite !mem_cat !in_cons orbAC orbC mem_cat -!orbA. +case/orP; first by move->; rewrite !orbT. +rewrite !orbA [_ || (_ \in eq1)]orbC; move: (h u); rewrite !mem_cat=> hu. +by move/hu; do 2! (case/orP; last by move->; rewrite !orbT); move->. +Qed. + + +Lemma oclause_neq_elimP oc e : + holds e (odnf_to_oform [:: oc]) <-> + holds e (odnf_to_oform (oclause_neq_elim oc)). +Proof. +case: oc => eq_l neq_l lt_l le_l; rewrite /oclause_neq_elim. +elim: neq_l lt_l => [|t neq_l ih] lt_l //=. +move: (ih lt_l) => /= {ih}. +set x1 := foldr _ _ _; set x2 := foldr _ _ _; set x3 := foldr _ _ _. +set x4 := foldr _ _ _ => h /=. +have -> : holds e x1 /\ + (eval e t <> 0%:R /\ + holds e x2) /\ + holds e x3 /\ holds e x4 \/ + false <-> + (eval e t <> 0%:R) /\ (holds e x1 /\ holds e x2 /\ holds e x3 /\ + holds e x4 \/ false). + split; case => //. + - by case=> ?; case; case=> ? ? [] ? ?; split=> //; left. + - by move=> ?; case => //; do 3! case => ?; move=> ?; left. +rewrite h {h} /= !map_cat /= -!map_comp. +set s1 := [seq _ | _ <- _]; set s2 := [seq _ | _ <- _]. +set s3 := [seq _ | _ <- _]; rewrite odnf_to_oform_cat. +suff {x1 x2 x3 x4} /= -> : + holds e (odnf_to_oform s2) <-> 0%:R < eval e t/\ holds e (odnf_to_oform s1). + suff /= -> : + holds e (odnf_to_oform s3) <-> 0%:R < - eval e t /\ holds e (odnf_to_oform s1). + rewrite oppr_gt0; split. + by case; move/eqP; rewrite neqr_lt; case/orP=> -> h1; [right | left]. + by case; case=> h ?; split=> //; apply/eqP; rewrite neqr_lt h ?orbT. + rewrite /s1 /s3. + elim: (neq_elim_aux lt_l neq_l) => /= [| t1 l ih] /=; first by split => //; case. + set y1 := foldr _ _ _; set y2 := foldr _ _ _; set y3 := foldr _ _ _. + rewrite ih; split. + case; first by case=> ?; case=> _; case; case=> -> ? ?; split=> //; left. + by case=> ? ?; split=> //; right. + by case=> ->; case; [case=> ?; case=> _; case=> ? ?; left| move=> ? ; right]. +rewrite /s1 /s2. +elim: (neq_elim_aux lt_l neq_l) => /= [| t1 l ih] /=; first by split => //; case. +set y1 := foldr _ _ _; set y2 := foldr _ _ _; set y3 := foldr _ _ _. +rewrite ih; split. + case; first by case=> ? [] _ [] [] ? ? ?; split=> //; left. + by case=> ? ?; split=> //; right. +case=> ? []; last by right. +by case=> ? [] _ [] ? ?; left. +Qed. + +Definition oclause_neq_leq_elim oc := + flatten (map oclause_neq_elim (oclause_leq_elim oc)). + +Lemma terms_of_neq_leq_elim oc1 oc2: + oc2 \in (oclause_neq_leq_elim oc1) -> + {subset (terms_of_oclause oc2) <= (terms_of_oclause oc1) ++ map Opp oc1.2}. +Proof. +rewrite /oclause_neq_leq_elim /flatten; rewrite foldr_map. +suff : forall oc3, + oc3 \in (oclause_leq_elim oc1) -> + (terms_of_oclause oc3 =i terms_of_oclause oc1) /\ oc3.2 = oc1.2. + elim: (oclause_leq_elim oc1) => [| t l ih] //= h1. + rewrite mem_cat; case/orP. + - move/terms_of_neq_elim=> h u; move/(h u); rewrite !mem_cat. + by case: (h1 t (mem_head _ _)); move/(_ u)=> -> ->. + - by move=> h; apply: (ih _ h) => ? loc3; apply: h1; rewrite in_cons loc3 orbT. +move=> {oc2} oc3 hoc3; split; first exact: terms_of_leq_elim. +case: oc3 hoc3=> eq2 neq2 lt2 leq2 /=; case: oc1=> eq1 neq1 lt1 leq1 /=. +elim: leq1 => [| t1 le1 ih] //=; first by rewrite inE; case/eqP=> _ ->. +rewrite map_cat mem_cat; move: ih. +elim: (leq_elim_aux eq1 lt1 le1) => [| t2 l2 ih2] //=; rewrite !in_cons. +move=> h1; case/orP=> /=. + case/orP; first by case/eqP. + by move=> h2; apply: ih2; rewrite ?h2 //; move=> h3; apply: h1; rewrite h3 orbT. +case/orP; first by case/eqP. +move=> h3; apply: ih2; last by rewrite h3 orbT. +by move=> h2; apply: h1; rewrite h2 orbT. +Qed. + +Lemma oclause_neq_leq_elimP oc e : + holds e (odnf_to_oform [:: oc]) <-> + holds e (odnf_to_oform (oclause_neq_leq_elim oc)). +Proof. +rewrite /oclause_neq_leq_elim. +rewrite oclause_leq_elimP; elim: (oclause_leq_elim oc) => [| t l ih] //=. +rewrite odnf_to_oform_cat /= ih -oclause_neq_elimP /=. +suff -> : forall A, A \/ false <-> A by []. +by intuition. +Qed. + +Definition oclause_to_w oc := + let s := oclause_neq_leq_elim oc in + map (fun x => let: Oclause eq_l neq_l lt_l leq_l := x in (eq_l, lt_l)) s. + +Definition w_to_oclause (t : seq (term R) * seq (term R)) := + Oclause t.1%PAIR [::] t.2%PAIR [::]. + +Lemma oclause_leq_elim4 bc oc : oc \in (oclause_leq_elim bc) -> oc.4 == [::]. +Proof. +case: bc => bc1 bc2 bc3 bc4; elim: bc4 bc1 bc3 oc => [|t bc4 ih] bc1 bc3 /= oc. + by rewrite inE; move/eqP; case: oc => ? ? ? oc4 /=; case=> _ _ _ /eqP. +rewrite map_cat; move: (ih bc1 bc3 oc) => /= {ih}. +elim: (leq_elim_aux bc1 bc3 bc4) => [| t2 l2 ih2] //= ih1. +rewrite in_cons; case/orP. + by move/eqP; case: oc {ih1 ih2} => ? ? ? ? [] /= _ _ _ /eqP. +rewrite mem_cat; case/orP=> [hoc1|]. + apply: ih2; first by move=> hoc2; apply: ih1; rewrite in_cons hoc2 orbT. + by rewrite mem_cat hoc1. +rewrite in_cons; case/orP=> [| hoc1]. + by move/eqP; case: {ih1 ih2} oc=> ? ? ? ? [] /= _ _ _ /eqP. +apply: ih2; first by move=> hoc2; apply: ih1; rewrite in_cons hoc2 orbT. +by rewrite mem_cat hoc1 orbT. +Qed. + +Lemma oclause_neq_elim2 bc oc : + oc \in (oclause_neq_elim bc) -> (oc.2 == [::]) && (oc.4 == bc.4). +Proof. +case: bc => bc1 bc2 bc3 bc4; elim: bc2 bc4 oc => [|t bc2 /= ih] bc4 /= oc. + by rewrite inE; move/eqP; case: oc => ? ? ? oc4 /=; case=> _ /eqP -> _ /eqP. +rewrite map_cat; move: (ih bc4 oc) => /= {ih}. +elim: (neq_elim_aux bc3 bc2) => [| t2 l2 ih2] //= ih1. +rewrite in_cons; case/orP. + by move/eqP; case: oc {ih1 ih2} => ? ? ? ? [] /= _ -> _ ->; rewrite !eqxx. +rewrite mem_cat; case/orP=> [hoc1|]. + apply: ih2; first by move=> hoc2; apply: ih1; rewrite in_cons hoc2 orbT. + by rewrite mem_cat hoc1. +rewrite in_cons; case/orP=> [| hoc1]. + by move/eqP; case: {ih1 ih2} oc=> ? ? ? ? [] /= _ -> _ ->; rewrite !eqxx. +apply: ih2; first by move=> hoc2; apply: ih1; rewrite in_cons hoc2 orbT. +by rewrite mem_cat hoc1 orbT. +Qed. + +Lemma oclause_to_wP e bc : + holds e (odnf_to_oform (oclause_neq_leq_elim bc)) <-> + holds e (odnf_to_oform (map w_to_oclause (oclause_to_w bc))). +Proof. +rewrite /oclause_to_w /oclause_neq_leq_elim. +move: (@oclause_leq_elim4 bc). +elim: (oclause_leq_elim bc) => [| t1 l1 ih1] //= h4. +rewrite !map_cat !odnf_to_oform_cat. +rewrite -[holds e (_ \/ _)]/(holds e _ \/ holds e _). +suff <- : (oclause_neq_elim t1) = map w_to_oclause + [seq (let: Oclause eq_l _ lt_l _ := x in (eq_l, lt_l)) + | x <- oclause_neq_elim t1]. + by rewrite ih1 //; move=> oc hoc; apply: h4; rewrite in_cons hoc orbT. +have : forall oc, oc \in (oclause_neq_elim t1) -> oc.2 = [::] /\ oc.4 = [::]. + move=> oc hoc; move/oclause_neq_elim2: (hoc); case/andP=> /eqP -> /eqP ->. + by move/eqP: (h4 _ (mem_head _ _))->. +elim: (oclause_neq_elim t1) => [| [teq1 tneq1 tleq1 tlt1] l2 ih2] h24 //=. +rewrite /w_to_oclause /=; move: (h24 _ (mem_head _ _ ))=> /= [] -> ->. +by congr (_ :: _); apply: ih2 => oc hoc; apply: h24; rewrite in_cons hoc orbT. +Qed. + +Variable wproj : nat -> (seq (term R) * seq (term R)) -> formula R. + +Definition proj (n : nat)(oc : oclause R) := + foldr Or False (map (wproj n) (oclause_to_w oc)). + +Hypothesis wf_QE_wproj : forall i bc (bc_i := wproj i bc), + dnf_rterm (w_to_oclause bc) -> qf_form bc_i && rformula bc_i. + +Lemma dnf_rterm_subproof bc : dnf_rterm bc -> + all (dnf_rterm \o w_to_oclause) (oclause_to_w bc). +Proof. +case: bc => leq lneql llt lle; rewrite /dnf_rterm /=; case/and4P=> req rneq rlt rle. +rewrite /oclause_to_w /= !all_map. +apply/allP => [] [oc_eq oc_neq oc_le oc_lt] hoc; rewrite /dnf_rterm /= andbT. +rewrite -all_cat; apply/allP=> u hu; move/terms_of_neq_leq_elim: hoc => /=. +move/(_ u); rewrite !mem_cat. +have {hu} hu : [|| u \in oc_eq, u \in oc_neq, u \in oc_le | u \in oc_lt]. + by move: hu; rewrite mem_cat; case/orP=> ->; rewrite ?orbT. +move/(_ hu); case/orP; last first. + move: rneq. + have <- : (all (@rterm R) (map Opp lneql)) = all (@rterm R) lneql. + by elim: lneql => [| t l] //= ->. + by move/allP; apply. +case/orP; first by apply: (allP req). +case/orP; first by apply: (allP rneq). +case/orP; first by apply: (allP rlt). +exact: (allP rle). +Qed. + + +Lemma wf_QE_proj i : forall bc (bc_i := proj i bc), + dnf_rterm bc -> qf_form bc_i && rformula bc_i. +Proof. +case=> leq lneql llt lle /= hdnf; move: (hdnf). +rewrite /dnf_rterm /=; case/and4P=> req rneq rlt rle; rewrite /proj; apply/andP. +move: (dnf_rterm_subproof hdnf). +elim: (oclause_to_w _ ) => //= [a t] ih /andP [h1 h2]. +by case: (ih h2)=> -> ->; case/andP: (wf_QE_wproj i h1) => -> ->. +Qed. + +Hypothesis valid_QE_wproj : + forall i bc (bc' := w_to_oclause bc) + (ex_i_bc := ('exists 'X_i, odnf_to_oform [:: bc'])%oT) e, + dnf_rterm bc' -> + reflect (holds e ex_i_bc) (qf_eval e (wproj i bc)). + +Lemma valid_QE_proj e i : forall bc (bc_i := proj i bc) + (ex_i_bc := ('exists 'X_i, odnf_to_oform [:: bc])%oT), + dnf_rterm bc -> reflect (holds e ex_i_bc) (qf_eval e (proj i bc)). +Proof. +move=> bc; rewrite /dnf_rterm => hdnf; rewrite /proj; apply: (equivP idP). +have -> : holds e ('exists 'X_i, odnf_to_oform [:: bc]) <-> + (exists x : R, holds (set_nth 0 e i x) + (odnf_to_oform (oclause_neq_leq_elim bc))). + split; case=> x h; exists x; first by rewrite -oclause_neq_leq_elimP. + by rewrite oclause_neq_leq_elimP. +have -> : + (exists x : R, + holds (set_nth 0 e i x) (odnf_to_oform (oclause_neq_leq_elim bc))) <-> + (exists x : R, + holds (set_nth 0 e i x) (odnf_to_oform (map w_to_oclause (oclause_to_w bc)))). + by split; case=> x; move/oclause_to_wP=> h; exists x. +move: (dnf_rterm_subproof hdnf). +rewrite /oclause_to_w; elim: (oclause_neq_leq_elim bc) => /= [|a l ih]. + by split=> //; case. +case/andP=> h1 h2; have {ih h2} ih := (ih h2); split. +- case/orP. + move/(valid_QE_wproj i e h1)=> /= [x /=] [] // [] h2 [] _ [] h3 _; exists x. + by left. + by case/ih => x h; exists x; right. +- case=> x [] /=. + + case=> h2 [] _ h3; apply/orP; left; apply/valid_QE_wproj => //=. + by exists x; left. + + by move=> hx; apply/orP; right; apply/ih; exists x. +Qed. + +Let elim_aux f n := foldr Or False (map (proj n) (qf_to_odnf f false)). + +Fixpoint quantifier_elim f := + match f with + | f1 /\ f2 => (quantifier_elim f1) /\ (quantifier_elim f2) + | f1 \/ f2 => (quantifier_elim f1) \/ (quantifier_elim f2) + | f1 ==> f2 => (~ quantifier_elim f1) \/ (quantifier_elim f2) + | ~ f => ~ quantifier_elim f + | ('exists 'X_n, f) => elim_aux (quantifier_elim f) n + | ('forall 'X_n, f) => ~ elim_aux (~ quantifier_elim f) n + | _ => f + end%oT. + +Lemma quantifier_elim_wf f : + let qf := quantifier_elim f in rformula f -> qf_form qf && rformula qf. +Proof. +suffices aux_wf f0 n : let qf := elim_aux f0 n in + rformula f0 -> qf_form qf && rformula qf. +- by elim: f => //=; do ?[ move=> f1 IH1 f2 IH2; + case/andP=> rf1 rf2; + case/andP:(IH1 rf1)=> -> ->; + case/andP:(IH2 rf2)=> -> -> // + | move=> n f1 IH rf1; + case/andP: (IH rf1)=> qff rf; + rewrite aux_wf ]. +rewrite /elim_aux => rf. +suffices or_wf fs : let ofs := foldr Or False fs in + all qf_form fs && all rformula fs -> qf_form ofs && rformula ofs. +- apply: or_wf. + suffices map_proj_wf bcs: let mbcs := map (proj n) bcs in + all dnf_rterm bcs -> all qf_form mbcs && all rformula mbcs. + by apply: map_proj_wf; exact: qf_to_dnf_rterm. + elim: bcs => [|bc bcs ihb] bcsr //= /andP[rbc rbcs]. + by rewrite andbAC andbA wf_QE_proj //= andbC ihb. +elim: fs => //= g gs ihg; rewrite -andbA => /and4P[-> qgs -> rgs] /=. +by apply: ihg; rewrite qgs rgs. +Qed. + +Lemma quantifier_elim_rformP e f : + rformula f -> reflect (holds e f) (qf_eval e (quantifier_elim f)). +Proof. +pose rc e n f := exists x, qf_eval (set_nth 0 e n x) f. +have auxP f0 e0 n0: qf_form f0 && rformula f0 -> + reflect (rc e0 n0 f0) (qf_eval e0 (elim_aux f0 n0)). ++ rewrite /elim_aux => cf; set bcs := qf_to_odnf f0 false. + apply: (@iffP (rc e0 n0 (odnf_to_oform bcs))); last first. + - by case=> x; rewrite -qf_to_dnfP //; exists x. + - by case=> x; rewrite qf_to_dnfP //; exists x. + have: all dnf_rterm bcs by case/andP: cf => _; exact: qf_to_dnf_rterm. + elim: {f0 cf}bcs => [|bc bcs IHbcs] /=; first by right; case. + case/andP=> r_bc /IHbcs {IHbcs}bcsP. + have f_qf := dnf_to_form_qf [:: bc]. + case: valid_QE_proj => //= [ex_x|no_x]. + left; case: ex_x => x /(qf_evalP _ f_qf); rewrite /= orbF => bc_x. + by exists x; rewrite /= bc_x. + apply: (iffP bcsP) => [[x bcs_x] | [x]] /=. + by exists x; rewrite /= bcs_x orbT. + case/orP => [bc_x|]; last by exists x. + by case: no_x; exists x; apply/(qf_evalP _ f_qf); rewrite /= bc_x. +elim: f e => //. +- move=> b e _; exact: idP. +- move=> t1 t2 e _; exact: eqP. +- move=> t1 t2 e _; exact: idP. +- move=> t1 t2 e _; exact: idP. +- move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; last by right; case. + by case/IH2; [left | right; case]. +- move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; first by do 2!left. + by case/IH2; [left; right | right; case]. +- move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; last by left. + by case/IH2; [left | right; move/(_ f1e)]. +- by move=> f IHf e /= /IHf[]; [right | left]. +- move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf. + by apply: (iffP (auxP _ _ _ rqf)) => [] [x]; exists x; exact/IHf. +move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf. +case: auxP => // [f_x|no_x]; first by right=> no_x; case: f_x => x /IHf[]. +by left=> x; apply/IHf=> //; apply/idPn=> f_x; case: no_x; exists x. +Qed. + +Definition proj_sat e f := qf_eval e (quantifier_elim (to_rform f)). + +Lemma proj_satP : forall e f, reflect (holds e f) (proj_sat e f). +Proof. +move=> e f; have fP := quantifier_elim_rformP e (to_rform_rformula f). +by apply: (iffP fP); move/to_rformP. +Qed. + +End EvalTerm. + +End ord.
\ No newline at end of file diff --git a/mathcomp/real_closed/polyorder.v b/mathcomp/real_closed/polyorder.v new file mode 100644 index 0000000..be4b7cc --- /dev/null +++ b/mathcomp/real_closed/polyorder.v @@ -0,0 +1,273 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import ssralg poly ssrnum zmodp polydiv interval. + +Import GRing.Theory. +Import Num.Theory. + +Import Pdiv.Idomain. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope ring_scope. + +Section Multiplicity. + +Variable R : idomainType. +Implicit Types x y c : R. +Implicit Types p q r d : {poly R}. + +(* Definition multiplicity (x : R) (p : {poly R}) : nat := *) +(* (odflt ord0 (pick (fun i : 'I_(size p).+1 => ((('X - x%:P) ^+ i %| p)) *) +(* && (~~ (('X - x%:P) ^+ i.+1 %| p))))). *) + +(* Notation "'\mu_' x" := (multiplicity x) *) +(* (at level 8, format "'\mu_' x") : ring_scope. *) + +(* Lemma mu0 : forall x, \mu_x 0 = 0%N. *) +(* Proof. *) +(* by move=> x; rewrite /multiplicity; case: pickP=> //= i; rewrite !dvdp0. *) +(* Qed. *) + +(* Lemma muP : forall p x, p != 0 -> *) +(* (('X - x%:P) ^+ (\mu_x p) %| p) && ~~(('X - x%:P) ^+ (\mu_x p).+1 %| p). *) +(* Proof. *) +(* move=> p x np0; rewrite /multiplicity; case: pickP=> //= hp. *) +(* have {hp} hip: forall i, i <= size p *) +(* -> (('X - x%:P) ^+ i %| p) -> (('X - x%:P) ^+ i.+1 %| p). *) +(* move=> i; rewrite -ltnS=> hi; move/negbT: (hp (Ordinal hi)). *) +(* by rewrite -negb_imply negbK=> /implyP. *) +(* suff: forall i, i <= size p -> ('X - x%:P) ^+ i %| p. *) +(* move=> /(_ _ (leqnn _)) /(size_dvdp np0). *) +(* rewrite -[size _]prednK; first by rewrite size_exp size_XsubC mul1n ltnn. *) +(* by rewrite lt0n size_poly_eq0 expf_eq0 polyXsubC_eq0 lt0n size_poly_eq0 np0. *) +(* elim=> [|i ihi /ltnW hsp]; first by rewrite expr0 dvd1p. *) +(* by rewrite hip // ihi. *) +(* Qed. *) + +(* Lemma cofactor_XsubC : forall p a, p != 0 -> *) +(* exists2 q : {poly R}, (~~ root q a) & p = q * ('X - a%:P) ^+ (\mu_a p). *) +(* Proof. *) +(* move=> p a np0. *) + +Definition multiplicity (x : R) (p : {poly R}) := + if p == 0 then 0%N else sval (multiplicity_XsubC p x). + +Notation "'\mu_' x" := (multiplicity x) + (at level 8, format "'\mu_' x") : ring_scope. + +Lemma mu_spec p a : p != 0 -> + { q : {poly R} & (~~ root q a) + & ( p = q * ('X - a%:P) ^+ (\mu_a p)) }. +Proof. +move=> nz_p; rewrite /multiplicity -if_neg. +by case: (_ p a) => m /=/sig2_eqW[q]; rewrite nz_p; exists q. +Qed. + +Lemma mu0 x : \mu_x 0 = 0%N. +Proof. by rewrite /multiplicity {1}eqxx. Qed. + +Lemma root_mu p x : ('X - x%:P) ^+ (\mu_x p) %| p. +Proof. +case p0: (p == 0); first by rewrite (eqP p0) mu0 expr0 dvd1p. +case: (@mu_spec p x); first by rewrite p0. +by move=> q qn0 hp //=; rewrite {2}hp dvdp_mulIr. +Qed. + +(* Lemma size_exp_XsubC : forall x n, size (('X - x%:P) ^+ n) = n.+1. *) +(* Proof. *) +(* move=> x n; rewrite -[size _]prednK ?size_exp ?size_XsubC ?mul1n //. *) +(* by rewrite ltnNge leqn0 size_poly_eq0 expf_neq0 // polyXsubC_eq0. *) +(* Qed. *) + +Lemma root_muN p x : p != 0 -> + (('X - x%:P)^+(\mu_x p).+1 %| p) = false. +Proof. +move=> pn0; case: (mu_spec x pn0)=> q qn0 hp /=. +rewrite {2}hp exprS dvdp_mul2r; last first. + by rewrite expf_neq0 // polyXsubC_eq0. +apply: negbTE; rewrite -eqp_div_XsubC; apply: contra qn0. +by move/eqP->; rewrite rootM root_XsubC eqxx orbT. +Qed. + +Lemma root_le_mu p x n : p != 0 -> ('X - x%:P)^+n %| p = (n <= \mu_x p)%N. +Proof. +move=> pn0; case: leqP=> hn; last apply/negP=> hp. + apply: (@dvdp_trans _ (('X - x%:P) ^+ (\mu_x p))); last by rewrite root_mu. + by rewrite dvdp_Pexp2l // size_XsubC. +suff : ('X - x%:P) ^+ (\mu_x p).+1 %| p by rewrite root_muN. +by apply: dvdp_trans hp; rewrite dvdp_Pexp2l // size_XsubC. +Qed. + +Lemma muP p x n : p != 0 -> + (('X - x%:P)^+n %| p) && ~~(('X - x%:P)^+n.+1 %| p) = (n == \mu_x p). +Proof. +move=> hp0; rewrite !root_le_mu//; case: (ltngtP n (\mu_x p))=> hn. ++ by rewrite ltnW//=. ++ by rewrite leqNgt hn. ++ by rewrite hn leqnn. +Qed. + +Lemma mu_gt0 p x : p != 0 -> (0 < \mu_x p)%N = root p x. +Proof. by move=> pn0; rewrite -root_le_mu// expr1 root_factor_theorem. Qed. + +Lemma muNroot p x : ~~ root p x -> \mu_x p = 0%N. +Proof. +case p0: (p == 0); first by rewrite (eqP p0) rootC eqxx. +by move=> pnx0; apply/eqP; rewrite -leqn0 leqNgt mu_gt0 ?p0. +Qed. + +Lemma mu_polyC c x : \mu_x (c%:P) = 0%N. +Proof. +case c0: (c == 0); first by rewrite (eqP c0) mu0. +by apply: muNroot; rewrite rootC c0. +Qed. + +Lemma cofactor_XsubC_mu x p n : + ~~ root p x -> \mu_x (p * ('X - x%:P) ^+ n) = n. +Proof. +move=> p0; apply/eqP; rewrite eq_sym -muP//; last first. + apply: contra p0; rewrite mulf_eq0 expf_eq0 polyXsubC_eq0 andbF orbF. + by move/eqP->; rewrite root0. +rewrite dvdp_mulIr /= exprS dvdp_mul2r -?root_factor_theorem //. +by rewrite expf_eq0 polyXsubC_eq0 andbF //. +Qed. + +Lemma mu_mul p q x : p * q != 0 -> + \mu_x (p * q) = (\mu_x p + \mu_x q)%N. +Proof. +move=>hpqn0; apply/eqP; rewrite eq_sym -muP//. +rewrite exprD dvdp_mul ?root_mu//=. +move:hpqn0; rewrite mulf_eq0 negb_or; case/andP=> hp0 hq0. +move: (mu_spec x hp0)=> [qp qp0 hp]. +move: (mu_spec x hq0)=> [qq qq0 hq]. +rewrite {2}hp {2}hq exprS exprD !mulrA [qp * _ * _]mulrAC. +rewrite !dvdp_mul2r ?expf_neq0 ?polyXsubC_eq0 // -eqp_div_XsubC. +move: (mulf_neq0 qp0 qq0); rewrite -hornerM; apply: contra; move/eqP->. +by rewrite hornerM hornerXsubC subrr mulr0. +Qed. + +Lemma mu_XsubC x : \mu_x ('X - x%:P) = 1%N. +Proof. +apply/eqP; rewrite eq_sym -muP; last by rewrite polyXsubC_eq0. +by rewrite expr1 dvdpp/= -{2}[_ - _]expr1 dvdp_Pexp2l // size_XsubC. +Qed. + +Lemma mu_mulC c p x : c != 0 -> \mu_x (c *: p) = \mu_x p. +Proof. +move=> cn0; case p0: (p == 0); first by rewrite (eqP p0) scaler0. +by rewrite -mul_polyC mu_mul ?mu_polyC// mulf_neq0 ?p0 ?polyC_eq0. +Qed. + +Lemma mu_opp p x : \mu_x (-p) = \mu_x p. +Proof. +rewrite -mulN1r -polyC1 -polyC_opp mul_polyC mu_mulC //. +by rewrite -oppr0 (inj_eq (inv_inj (@opprK _))) oner_eq0. +Qed. + +Lemma mu_exp p x n : \mu_x (p ^+ n) = (\mu_x p * n)%N. +Proof. +elim: n p => [|n ihn] p; first by rewrite expr0 mu_polyC muln0. +case p0: (p == 0); first by rewrite (eqP p0) exprS mul0r mu0 mul0n. +by rewrite exprS mu_mul ?ihn ?mulnS// mulf_eq0 expf_eq0 p0 andbF. +Qed. + +Lemma mu_addr p q x : p != 0 -> (\mu_x p < \mu_x q)%N -> + \mu_x (p + q) = \mu_x p. +Proof. +move=> pn0 mupq. +have pqn0 : p + q != 0. + move: mupq; rewrite ltnNge; apply: contra. + by rewrite -[q]opprK subr_eq0; move/eqP->; rewrite opprK mu_opp leqnn. +have qn0: q != 0 by move: mupq; apply: contraL; move/eqP->; rewrite mu0 ltn0. +case: (mu_spec x pn0)=> [qqp qqp0] hp. +case: (mu_spec x qn0)=> [qqq qqq0] hq. +rewrite hp hq -(subnK (ltnW mupq)). +rewrite mu_mul ?mulf_eq0; last first. + rewrite expf_eq0 polyXsubC_eq0 andbF orbF. + by apply: contra qqp0; move/eqP->; rewrite root0. +rewrite mu_exp mu_XsubC mul1n [\mu_x qqp]muNroot // add0n. +rewrite exprD mulrA -mulrDl mu_mul; last first. + by rewrite mulrDl -mulrA -exprD subnK 1?ltnW // -hp -hq. +rewrite muNroot ?add0n ?mu_exp ?mu_XsubC ?mul1n //. +rewrite rootE !hornerE horner_exp hornerXsubC subrr. +by rewrite -subnSK // subnS exprS mul0r mulr0 addr0. +Qed. + +Lemma mu_addl p q x : q != 0 -> (\mu_x p > \mu_x q)%N -> + \mu_x (p + q) = \mu_x q. +Proof. by move=> q0 hmu; rewrite addrC mu_addr. Qed. + +Lemma mu_div p x n : (n <= \mu_x p)%N -> + \mu_x (p %/ ('X - x%:P) ^+ n) = (\mu_x p - n)%N. +Proof. +move=> hn. +case p0: (p == 0); first by rewrite (eqP p0) div0p mu0 sub0n. +case: (@mu_spec p x); rewrite ?p0 // => q hq hp. +rewrite {1}hp -{1}(subnK hn) exprD mulrA. +rewrite Pdiv.IdomainMonic.mulpK; last by apply: monic_exp; exact: monicXsubC. +rewrite mu_mul ?mulf_eq0 ?expf_eq0 ?polyXsubC_eq0 ?andbF ?orbF; last first. + by apply: contra hq; move/eqP->; rewrite root0. +by rewrite mu_exp muNroot // add0n mu_XsubC mul1n. +Qed. + +End Multiplicity. + +Notation "'\mu_' x" := (multiplicity x) + (at level 8, format "'\mu_' x") : ring_scope. + + +Section PolyrealIdomain. + + (*************************************************************************) + (* This should be replaced by a 0-characteristic condition + integrality *) + (* and merged into poly and polydiv *) + (*************************************************************************) + +Variable R : realDomainType. + +Lemma size_deriv (p : {poly R}) : size p^`() = (size p).-1. +Proof. +have [lep1|lt1p] := leqP (size p) 1. + by rewrite {1}[p]size1_polyC // derivC size_poly0 -subn1 (eqnP lep1). +rewrite size_poly_eq // mulrn_eq0 -subn2 -subSn // subn2. +by rewrite lead_coef_eq0 -size_poly_eq0 -(subnKC lt1p). +Qed. + +Lemma derivn_poly0 : forall (p : {poly R}) n, (size p <= n)%N = (p^`(n) == 0). +Proof. +move=> p n; apply/idP/idP. + move=> Hpn; apply/eqP; apply/polyP=>i; rewrite coef_derivn. + rewrite nth_default; first by rewrite mul0rn coef0. + by apply: leq_trans Hpn _; apply leq_addr. +elim: n {-2}n p (leqnn n) => [m | n ihn [| m]] p. +- by rewrite leqn0; move/eqP->; rewrite derivn0 leqn0 -size_poly_eq0. +- by move=> _; apply: ihn; rewrite leq0n. +- rewrite derivSn => hmn hder; case e: (size p) => [|sp] //. + rewrite -(prednK (ltn0Sn sp)) [(_.-1)%N]lock -e -lock -size_deriv ltnS. + exact: ihn. +Qed. + +Lemma mu_deriv : forall x (p : {poly R}), root p x -> + \mu_x (p^`()) = (\mu_x p - 1)%N. +Proof. +move=> x p px0; have [-> | nz_p] := eqVneq p 0; first by rewrite derivC mu0. +have [q nz_qx Dp] := mu_spec x nz_p. +case Dm: (\mu_x p) => [|m]; first by rewrite Dp Dm mulr1 (negPf nz_qx) in px0. +rewrite subn1 Dp Dm !derivCE exprS mul1r mulrnAr -mulrnAl mulrA -mulrDl. +rewrite cofactor_XsubC_mu // rootE !(hornerE, hornerMn) subrr mulr0 add0r. +by rewrite mulrn_eq0. +Qed. + +Lemma mu_deriv_root : forall x (p : {poly R}), p != 0 -> root p x -> + \mu_x p = (\mu_x (p^`()) + 1)%N. +Proof. +by move=> x p p0 rpx; rewrite mu_deriv // subn1 addn1 prednK // mu_gt0. +Qed. + +End PolyrealIdomain. + + + diff --git a/mathcomp/real_closed/polyrcf.v b/mathcomp/real_closed/polyrcf.v new file mode 100644 index 0000000..b49e729 --- /dev/null +++ b/mathcomp/real_closed/polyrcf.v @@ -0,0 +1,1857 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import bigop ssralg poly polydiv ssrnum zmodp. +Require Import polyorder path interval ssrint. + +(****************************************************************************) +(* This files contains basic (and unformatted) theory for polynomials *) +(* over a realclosed fields. From the IVT (contained in the rcfType *) +(* structure), we derive Rolle's Theorem, the Mean Value Theorem, a root *) +(* isolation procedure and the notion of neighborhood. *) +(* *) +(* sgp_minfty p == the sign of p in -oo *) +(* sgp_pinfty p == the sign of p in +oo *) +(* cauchy_bound p == the cauchy bound of p *) +(* (this strictly bounds the norm of roots of p) *) +(* roots p a b == the ordered list of roots of p in `[a, b] *) +(* defaults to [::] when p == 0 *) +(* rootsR p == the ordered list of all roots of p, ([::] if p == 0). *) +(* next_root p x b == the smallest root of p contained in `[x, maxn x b] *) +(* if p has no root on `[x, maxn x b], we pick maxn x b. *) +(* prev_root p x a == the smallest root of p contained in `[minn x a, x] *) +(* if p has no root on `[minn x a, x], we pick minn x a. *) +(* neighpr p a b := `]a, next_root p a b[. *) +(* == an open interval of the form `]a, x[, with x <= b *) +(* in which p has no root. *) +(* neighpl p a b := `]prev_root p a b, b[. *) +(* == an open interval of the form `]x, b[, with a <= x *) +(* in which p has no root. *) +(* sgp_right p a == the sign of p on the right of a. *) +(****************************************************************************) + + +Import GRing.Theory Num.Theory Num.Def. +Import Pdiv.Idomain. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope nat_scope. +Local Open Scope ring_scope. + +Local Notation noroot p := (forall x, ~~ root p x). +Local Notation mid x y := ((x + y) / 2%:R). + +Section more. +Section SeqR. + +Lemma last1_neq0 : forall (R : ringType) (s: seq R) (c:R), c != 0 -> + (last c s != 0) = (last 1 s != 0). +Proof. by move=> R'; elim=> [|t s ihs] c cn0 //; rewrite oner_eq0 cn0. Qed. + +End SeqR. + +Section poly. +Import Pdiv.Ring Pdiv.ComRing. + +Variable R : idomainType. + +Implicit Types p q : {poly R}. + +Lemma lead_coefDr p q : + (size q > size p)%N -> lead_coef (p + q) = lead_coef q. +Proof. by move/lead_coefDl<-; rewrite addrC. Qed. + +Lemma leq1_size_polyC (c : R) : (size c%:P <= 1)%N. +Proof. by rewrite size_polyC; case: (c == 0). Qed. + +Lemma my_size_exp p n : p != 0 -> + (size (p ^+ n)) = ((size p).-1 * n).+1%N. +Proof. +by move=> hp; rewrite -size_exp prednK // lt0n size_poly_eq0 expf_neq0. +Qed. + +Lemma coef_comp_poly p q n : + (p \Po q)`_n = \sum_(i < size p) p`_i * (q ^+ i)`_n. +Proof. +rewrite comp_polyE coef_sum. +by elim/big_ind2: _ => [//|? ? ? ? -> -> //|i]; rewrite coefZ. +Qed. + +Lemma gt_size_poly p n : (size p > n)%N -> p != 0. +Proof. +by move=> h; rewrite -size_poly_eq0 lt0n_neq0 //; apply: leq_ltn_trans h. +Qed. + +Lemma lead_coef_comp_poly p q : (size q > 1)%N -> + lead_coef (p \Po q) = (lead_coef p) * (lead_coef q) ^+ (size p).-1. +Proof. +move=> sq; rewrite !lead_coefE coef_comp_poly size_comp_poly. +case hp: (size p) => [|n]. + move/eqP: hp; rewrite size_poly_eq0 => /eqP ->. + by rewrite big_ord0 coef0 mul0r. +rewrite big_ord_recr /= big1 => [|i _]. + by rewrite add0r -lead_coefE -lead_coef_exp lead_coefE size_exp mulnC. +rewrite [X in _ * X]nth_default ?mulr0 ?(leq_trans (size_exp_leq _ _)) //. +by rewrite mulnC ltn_mul2r -subn1 subn_gt0 sq /=. +Qed. + +End poly. +End more. + +(******************************************************************) +(* Definitions and properties for polynomials in a numDomainType. *) +(******************************************************************) +Section PolyNumDomain. + +Variable R : numDomainType. +Implicit Types (p q : {poly R}). + +Definition sgp_pinfty (p : {poly R}) := sgr (lead_coef p). +Definition sgp_minfty (p : {poly R}) := + sgr ((-1) ^+ (size p).-1 * (lead_coef p)). + +End PolyNumDomain. + +(******************************************************************) +(* Definitions and properties for polynomials in a realFieldType. *) +(******************************************************************) +Section PolyRealField. + +Variable R : realFieldType. +Implicit Types (p q : {poly R}). + +Section SgpInfty. + +Lemma sgp_pinfty_sym p : sgp_pinfty (p \Po -'X) = sgp_minfty p. +Proof. +rewrite /sgp_pinfty /sgp_minfty lead_coef_comp_poly ?size_opp ?size_polyX //. +by rewrite lead_coef_opp lead_coefX mulrC. +Qed. + +Lemma poly_pinfty_gt_lc p : lead_coef p > 0 -> + exists n, forall x, x >= n -> p.[x] >= lead_coef p. +Proof. +elim/poly_ind: p => [| q c IHq]; first by rewrite lead_coef0 ltrr. +have [->|q_neq0] := eqVneq q 0. + by rewrite mul0r add0r lead_coefC => c_gt0; exists 0 => x _; rewrite hornerC. +rewrite lead_coefDl ?size_mul ?polyX_eq0 // ?lead_coefMX; last first. + rewrite size_polyX addn2 size_polyC /= ltnS. + by rewrite (leq_trans (leq_b1 _)) // size_poly_gt0. +move=> lq_gt0; have [y Hy] := IHq lq_gt0. +pose z := (1 + (lead_coef q) ^-1 * `|c|); exists (maxr y z) => x. +have z_gt0 : 0 < z by rewrite ltr_spaddl ?ltr01 ?mulr_ge0 ?invr_ge0 // ltrW. +rewrite !hornerE ler_maxl => /andP[/Hy Hq Hc]. +rewrite (@ler_trans _ (lead_coef q * z + c)) //; last first. + rewrite ler_add2r (@ler_trans _ (q.[x] * z)) // ?ler_pmul2r //. + by rewrite ler_pmul2l // (ltr_le_trans _ Hq). +rewrite mulrDr mulr1 -addrA ler_addl mulVKf ?gtr_eqF //. +by rewrite -[c]opprK subr_ge0 normrN ler_norm. +Qed. + +(* :REMARK: not necessary here ! *) +Lemma poly_lim_infty p m : lead_coef p > 0 -> (size p > 1)%N -> + exists n, forall x, x >= n -> p.[x] >= m. +Proof. +elim/poly_ind: p m => [| q c _] m; first by rewrite lead_coef0 ltrr. +have [-> _|q_neq0] := eqVneq q 0. + by rewrite mul0r add0r size_polyC ltnNge leq_b1. +rewrite lead_coefDl ?size_mul ?polyX_eq0 // ?lead_coefMX; last first. + rewrite size_polyX addn2 size_polyC /= ltnS. + by rewrite (leq_trans (leq_b1 _)) // size_poly_gt0. +move=> lq_gt0; have [y Hy _] := poly_pinfty_gt_lc lq_gt0. +pose z := (1 + (lead_coef q) ^-1 * (`|m| + `|c|)); exists (maxr y z) => x. +have z_gt0 : 0 < z. + by rewrite ltr_spaddl ?ltr01 ?mulr_ge0 ?invr_ge0 ?addr_ge0 // ?ltrW. +rewrite !hornerE ler_maxl => /andP[/Hy Hq Hc]. +rewrite (@ler_trans _ (lead_coef q * z + c)) //; last first. + rewrite ler_add2r (@ler_trans _ (q.[x] * z)) // ?ler_pmul2r //. + by rewrite ler_pmul2l // (ltr_le_trans _ Hq). +rewrite mulrDr mulr1 mulVKf ?gtr_eqF // addrA -addrA ler_paddr //. + by rewrite -[c]opprK subr_ge0 normrN ler_norm. +by rewrite ler_paddl ?ler_norm // ?ltrW. +Qed. + +End SgpInfty. + +Section CauchyBound. + +Definition cauchy_bound (p : {poly R}) := + 1 + `|lead_coef p|^-1 * \sum_(i < size p) `|p`_i|. + +(* Could be a sharp bound, and proof should shrink... *) +Lemma cauchy_boundP : forall (p : {poly R}) x, + p != 0 -> p.[x] = 0 -> `| x | < cauchy_bound p. +Proof. +move=> p x np0 rpx; rewrite ltr_spaddl ?ltr01 //. +case e: (size p) => [|n]; first by move: np0; rewrite -size_poly_eq0 e eqxx. +have lcp : `|lead_coef p| > 0 by move: np0; rewrite -lead_coef_eq0 -normr_gt0. +have lcn0 : `|lead_coef p| != 0 by rewrite normr_eq0 -normr_gt0. +case: (lerP `|x| 1) => cx1. + rewrite (ler_trans cx1) // /cauchy_bound ler_pdivl_mull // mulr1. + by rewrite big_ord_recr /= /lead_coef e ler_addr sumr_ge0. +case es: n e => [|m] e. + suff p0 : p = 0 by rewrite p0 eqxx in np0. + by move: rpx; rewrite (@size1_polyC _ p) ?e ?lerr // hornerC; move->. +move: rpx; rewrite horner_coef e -es big_ord_recr /=; move/eqP; rewrite eq_sym. +rewrite -subr_eq sub0r; move/eqP => h1. +have {h1} h1 : `|p`_n| * `|x| ^+ n <= \sum_(i < n) `|p`_i * x ^+ i|. + rewrite -normrX -normrM -normrN h1. + by rewrite (ler_trans (ler_norm_sum _ _ _)) // lerr. +have xp : `| x | > 0 by rewrite (ltr_trans _ cx1) ?ltr01. +move: h1; rewrite {-1}es exprS mulrA -ler_pdivl_mulr ?exprn_gt0 // big_distrl /=. +rewrite big_ord_recr /= normrM normrX -mulrA es mulfV; last first. + by rewrite expf_eq0 negb_and eq_sym (ltr_eqF xp) orbT. +have pnp : 0 < `|p`_n| by move: lcp; rewrite /lead_coef e es. +rewrite mulr1 -es mulrC -ler_pdivl_mulr //. +rewrite [_ / _]mulrC /cauchy_bound /lead_coef e -es /=. +move=> h1; apply: (ler_trans h1) => //. +rewrite ler_pmul2l ?invr_gt0 ?(ltrW pnp) // big_ord_recr /=. +rewrite es [_ + `|p`_m.+1|]addrC ler_paddl // ?normr_ge0 //. +rewrite big_ord_recr /= ler_add2r; apply: ler_sum => i. +rewrite normrM normrX. +rewrite -mulrA ler_pimulr ?normrE // ler_pdivr_mulr ?exprn_gt0 // mul1r. +by rewrite ler_eexpn2l // 1?ltrW //; case: i=> i hi /=; rewrite ltnW. +(* this could be improved a little bit with int exponents *) +Qed. + +Lemma le_cauchy_bound p : p != 0 -> {in `]-oo, (- cauchy_bound p)], noroot p}. +Proof. +move=> p_neq0 x; rewrite inE /= lerNgt; apply: contra => /rootP. +by move=> /(cauchy_boundP p_neq0) /ltr_normlP []; rewrite ltr_oppl. +Qed. +Hint Resolve le_cauchy_bound. + +Lemma ge_cauchy_bound p : p != 0 -> {in `[cauchy_bound p, +oo[, noroot p}. +Proof. +move=> p_neq0 x; rewrite inE andbT /= lerNgt; apply: contra => /rootP. +by move=> /(cauchy_boundP p_neq0) /ltr_normlP []; rewrite ltr_oppl. +Qed. +Hint Resolve ge_cauchy_bound. + +Lemma cauchy_bound_gt0 p : cauchy_bound p > 0. +Proof. +rewrite ltr_spaddl ?ltr01 ?mulr_ge0 ?invr_ge0 ?normr_ge0 //. +by rewrite sumr_ge0 // => i; rewrite normr_ge0. +Qed. +Hint Resolve cauchy_bound_gt0. + +Lemma cauchy_bound_ge0 p : cauchy_bound p >= 0. +Proof. by rewrite ltrW. Qed. +Hint Resolve cauchy_bound_ge0. + +End CauchyBound. + +End PolyRealField. + +(************************************************************) +(* Definitions and properties for polynomials in a rcfType. *) +(************************************************************) +Section PolyRCF. + +Variable R : rcfType. + +Section Prelim. + +Implicit Types a b c : R. +Implicit Types x y z t : R. +Implicit Types p q r : {poly R}. + +(* we restate poly_ivt in a nicer way. Perhaps the def of PolyRCF should *) +(* be moved in this file, juste above this section *) + +Lemma poly_ivt (p : {poly R}) (a b : R) : + a <= b -> 0 \in `[p.[a], p.[b]] -> { x : R | x \in `[a, b] & root p x }. +Proof. by move=> leab root_p_ab; exact/sig2W/poly_ivt. Qed. + +Lemma polyrN0_itv (i : interval R) (p : {poly R}) : {in i, noroot p} + -> forall y x : R, y \in i -> x \in i -> sgr p.[x] = sgr p.[y]. +Proof. +move=> hi y x hy hx; wlog xy: x y hx hy / x <= y=> [hwlog|]. + by case/orP: (ler_total x y)=> xy; [|symmetry]; apply: hwlog. +have hxyi: {subset `[x, y] <= i}. + move=> z; apply: subitvP=> /=. + by case: i hx hy {hi}=> [[[] ?|] [[] ?|]] /=; do ?[move/itvP->|move=> ?]. +do 2![case: sgrP; first by move/rootP; rewrite (negPf (hi _ _))]=> //. + move=> /ltrW py0 /ltrW p0x; case: (@poly_ivt (- p) x y)=> //. + by rewrite inE !hornerN !oppr_cp0 p0x. + by move=> z hz; rewrite rootN (negPf (hi z _)) // hxyi. +move=> /ltrW p0y /ltrW px0; case: (@poly_ivt p x y); rewrite ?inE ?px0 //. +by move=> z hz; rewrite (negPf (hi z _)) // hxyi. +Qed. + +Lemma poly_div_factor : forall (a:R) (P : {poly R} -> Prop), + (forall k, P k%:P) -> + (forall p n k, p.[a] != 0 -> P p -> + P (p * ('X - a%:P)^+(n.+1) + k%:P)%R) + -> forall p, P p. +Proof. +move=> a P Pk Pq p. +move: {-2}p (leqnn (size p)); elim: (size p)=> {p} [|n ihn] p spn. + move: spn; rewrite leqn0 size_poly_eq0; move/eqP->; rewrite -polyC0. + exact: Pk. +case: (leqP (size p) 1)=> sp1; first by rewrite [p]size1_polyC ?sp1//. +rewrite (Pdiv.IdomainMonic.divp_eq (monicXsubC a) p) [_ %% _]size1_polyC; last first. + rewrite -ltnS. + by rewrite (@leq_trans (size ('X - a%:P))) // ?ltn_modp ?polyXsubC_eq0 ?size_XsubC. +have [n' [q hqa hp]] := multiplicity_XsubC (p %/ ('X - a%:P)) a. +rewrite divpN0 ?size_XsubC ?polyXsubC_eq0 ?sp1 //= in hqa. +rewrite hp -mulrA -exprSr; apply: Pq=> //; apply: ihn. +rewrite (@leq_trans (size (q * ('X - a%:P) ^+ n'))) //. + rewrite size_mul ?expf_eq0 ?polyXsubC_eq0 ?andbF //; last first. + by apply: contra hqa; move/eqP->; rewrite root0. + by rewrite size_exp_XsubC addnS leq_addr. +by rewrite -hp size_divp ?polyXsubC_eq0 ?size_XsubC // leq_subLR. +Qed. + +Lemma nth_root x n : x > 0 -> { y | y > 0 & y ^+ (n.+1) = x }. +Proof. +move=> l0x. +case: (ltrgtP x 1)=> hx; last by exists 1; rewrite ?hx ?lter01// expr1n. + case: (@poly_ivt ('X ^+ n.+1 - x%:P) 0 1); first by rewrite ler01. + rewrite ?(hornerE,horner_exp) ?inE. + by rewrite exprS mul0r sub0r expr1n oppr_cp0 subr_gte0/= !ltrW. + move=> y; case/andP=> [l0y ly1]; rewrite rootE ?(hornerE,horner_exp). + rewrite subr_eq0; move/eqP=> hyx; exists y=> //; rewrite lt0r l0y. + rewrite andbT; apply/eqP=> y0; move: hyx; rewrite y0. + by rewrite exprS mul0r=> x0; move: l0x; rewrite -x0 ltrr. +case: (@poly_ivt ('X ^+ n.+1 - x%:P) 0 x); first by rewrite ltrW. + rewrite ?(hornerE,horner_exp) exprS mul0r sub0r ?inE. + by rewrite oppr_cp0 (ltrW l0x) subr_ge0 ler_eexpr // ltrW. +move=> y; case/andP=> l0y lyx; rewrite rootE ?(hornerE,horner_exp). +rewrite subr_eq0; move/eqP=> hyx; exists y=> //; rewrite lt0r l0y. +rewrite andbT; apply/eqP=> y0; move: hyx; rewrite y0. +by rewrite exprS mul0r=> x0; move: l0x; rewrite -x0 ltrr. +Qed. + +Lemma poly_cont x p e : e > 0 -> exists2 d, + d > 0 & forall y, `|y - x| < d -> `|p.[y] - p.[x]| < e. +Proof. +elim/(@poly_div_factor x): p e. + move=> e ep; exists 1; rewrite ?ltr01// => y hy. + by rewrite !hornerC subrr normr0. +move=> p n k pxn0 Pp e ep. +case: (Pp (`|p.[x]|/2%:R)). + by rewrite pmulr_lgt0 ?invr_gte0//= ?ltr0Sn// normrE. +move=> d' d'p He'. +case: (@nth_root (e / ((3%:R / 2%:R) * `|p.[x]|)) n). + by rewrite ltr_pdivl_mulr ?mul0r ?pmulr_rgt0 ?invr_gt0 ?normrE ?ltr0Sn. +move=> d dp rootd. +exists (minr d d'); first by rewrite ltr_minr dp. +move=> y; rewrite ltr_minr; case/andP=> hxye hxye'. +rewrite !(hornerE, horner_exp) subrr [0 ^+ _]exprS mul0r mulr0 add0r addrK. +rewrite normrM (@ler_lt_trans _ (`|p.[y]| * d ^+ n.+1)) //. + by rewrite ler_wpmul2l ?normrE // normrX ler_expn2r -?topredE /= ?normrE 1?ltrW. +rewrite rootd mulrCA gtr_pmulr //. +rewrite ltr_pdivr_mulr ?mul1r ?pmulr_rgt0 ?invr_gt0 ?ltr0Sn ?normrE //. +rewrite mulrDl mulrDl divff; last by rewrite -mulr2n pnatr_eq0. +rewrite !mul1r mulrC -ltr_subl_addr. +by rewrite (ler_lt_trans _ (He' y _)) // ler_sub_dist. +Qed. + +(* Todo : orderedpoly !! *) +(* Lemma deriv_expz_nat : forall (n : nat) p, (p ^ n)^`() = (p^`() * p ^ (n.-1)) *~ n. *) +(* Proof. *) +(* elim=> [|n ihn] p /=; first by rewrite expr0z derivC mul0zr. *) +(* rewrite exprSz_nat derivM ihn mulzrAr mulrCA -exprSz_nat. *) +(* by case: n {ihn}=> [|n] //; rewrite mul0zr addr0 mul1zr. *) +(* Qed. *) + +(* Definition derivCE := (derivE, deriv_expz_nat). *) + +(* Lemma size_poly_ind : forall K : {poly R} -> Prop, *) +(* K 0 -> *) +(* (forall p sp, size p = sp.+1 -> *) +(* forall q, (size q <= sp)%N -> K q -> K p) *) +(* -> forall p, K p. *) +(* Proof. *) +(* move=> K K0 ihK p. *) +(* move: {-2}p (leqnn (size p)); elim: (size p)=> {p} [|n ihn] p spn. *) +(* by move: spn; rewrite leqn0 size_poly_eq0; move/eqP->. *) +(* case spSn: (size p == n.+1). *) +(* move/eqP:spSn; move/ihK=> ihKp; apply: (ihKp 0)=>//. *) +(* by rewrite size_poly0. *) +(* by move:spn; rewrite leq_eqVlt spSn /= ltnS; by move/ihn. *) +(* Qed. *) + +(* Lemma size_poly_indW : forall K : {poly R} -> Prop, *) +(* K 0 -> *) +(* (forall p sp, size p = sp.+1 -> *) +(* forall q, size q = sp -> K q -> K p) *) +(* -> forall p, K p. *) +(* Proof. *) +(* move=> K K0 ihK p. *) +(* move: {-2}p (leqnn (size p)); elim: (size p)=> {p} [|n ihn] p spn. *) +(* by move: spn; rewrite leqn0 size_poly_eq0; move/eqP->. *) +(* case spSn: (size p == n.+1). *) +(* move/eqP:spSn; move/ihK=> ihKp; case: n ihn spn ihKp=> [|n] ihn spn ihKp. *) +(* by apply: (ihKp 0)=>//; rewrite size_poly0. *) +(* apply: (ihKp 'X^n)=>//; first by rewrite size_polyXn. *) +(* by apply: ihn; rewrite size_polyXn. *) +(* by move:spn; rewrite leq_eqVlt spSn /= ltnS; by move/ihn. *) +(* Qed. *) + +Lemma poly_ltsp_roots : forall p (rs : seq R), + (size rs >= size p)%N -> uniq rs -> all (root p) rs -> p = 0. +Proof. +move=> p rs hrs urs rrs; apply/eqP; apply: contraLR hrs=> np0. +by rewrite -ltnNge; apply: max_poly_roots. +Qed. + +Lemma ivt_sign : forall (p : {poly R}) (a b : R), + a <= b -> sgr p.[a] * sgr p.[b] = -1 -> { x : R | x \in `]a, b[ & root p x}. +Proof. +move=> p a b hab /eqP; rewrite mulrC mulr_sg_eqN1=> /andP [spb0 /eqP spb]. + case: (@poly_ivt (sgr p.[b] *: p) a b)=> //. + by rewrite !hornerZ {1}spb mulNr -!normrEsg inE /= oppr_cp0 !normrE. +move=> c hc; rewrite rootZ ?sgr_eq0 // => rpc; exists c=> //. +(* need for a lemma reditvP *) +rewrite inE /= !ltr_neqAle andbCA -!andbA [_ && (_ <= _)]hc andbT eq_sym -negb_or. +apply/negP=> /orP [] /eqP ec; move: rpc; rewrite -ec /root ?(negPf spb0) //. +by rewrite -sgr_cp0 -[sgr _]opprK -spb eqr_oppLR oppr0 sgr_cp0 (negPf spb0). +Qed. + +Let rolle_weak : forall a b p, a < b -> + p.[a] = 0 -> p.[b] = 0 -> + {c | (c \in `]a, b[) & (p^`().[c] == 0) || (p.[c] == 0)}. +Proof. +move=> a b p lab pa0 pb0; apply/sig2W. +case p0: (p == 0). + rewrite (eqP p0); exists (mid a b); first by rewrite mid_in_itv. + by rewrite derivC horner0 eqxx. +have [n [p' p'a0 hp]] := multiplicity_XsubC p a; rewrite p0 /= in p'a0. +case: n hp pa0 p0 pb0 p'a0=> [ | n -> _ p0 pb0 p'a0]. + by rewrite {1}expr0 mulr1 rootE=> ->; move/eqP->. +have [m [q qb0 hp']] := multiplicity_XsubC p' b. +rewrite (contraNneq _ p'a0) /= in qb0 => [|->]; last exact: root0. +case: m hp' pb0 p0 p'a0 qb0=> [|m]. + rewrite {1}expr0 mulr1=> ->; move/eqP. + rewrite !(hornerE, horner_exp, mulf_eq0). + by rewrite !expf_eq0 !subr_eq0 !(gtr_eqF lab) !andbF !orbF !rootE=> ->. +move=> -> _ p0 p'a0 qb0; case: (sgrP (q.[a] * q.[b])); first 2 last. +- move=> sqasb; case: (@ivt_sign q a b)=> //; first exact: ltrW. + by apply/eqP; rewrite -sgrM sgr_cp0. + move=> c lacb rqc; exists c=> //. + by rewrite !hornerM (eqP rqc) !mul0r eqxx orbT. +- move/eqP; rewrite mulf_eq0 (rootPf qb0) orbF; move/eqP=> qa0. + by move: p'a0; rewrite ?rootM rootE qa0 eqxx. +- move=> hspq; rewrite !derivCE /= !mul1r mulrDl !pmulrn. + set xan := (('X - a%:P) ^+ n); set xbm := (('X - b%:P) ^+ m). + have ->: ('X - a%:P) ^+ n.+1 = ('X - a%:P) * xan by rewrite exprS. + have ->: ('X - b%:P) ^+ m.+1 = ('X - b%:P) * xbm by rewrite exprS. + rewrite -mulrzl -[_ *~ n.+1]mulrzl. + have fac : forall x y z : {poly R}, x * (y * xbm) * (z * xan) + = (y * z * x) * (xbm * xan). + by move=> x y z; rewrite mulrCA !mulrA [_ * y]mulrC mulrA. + rewrite !fac -!mulrDl; set r := _ + _ + _. + case: (@ivt_sign (sgr q.[b] *: r) a b); first exact: ltrW. + rewrite !hornerZ !sgr_smul mulrACA -expr2 sqr_sg (rootPf qb0) mul1r. + rewrite !(subrr, mul0r, mulr0, addr0, add0r, hornerC, hornerXsubC, + hornerD, hornerN, hornerM, hornerMn). + rewrite [_ * _%:R]mulrC -!mulrA !pmulrn !mulrzl !sgrMz -sgrM. + rewrite mulrAC mulrA -mulrA sgrM -opprB mulNr sgrN sgrM. + by rewrite !gtr0_sg ?subr_gt0 ?mulr1 // mulrC. +move=> c lacb; rewrite rootE hornerZ mulf_eq0. +rewrite sgr_cp0 (rootPf qb0) orFb=> rc0. +by exists c=> //; rewrite !hornerM !mulf_eq0 rc0. +Qed. + +Theorem rolle : forall a b p, a < b -> + p.[a] = p.[b] -> {c | c \in `]a, b[ & p^`().[c] = 0}. +Proof. +move=> a b p lab pab. +wlog pb0 : p pab / p.[b] = 0=> [hwlog|]. + case: (hwlog (p - p.[b]%:P)); rewrite ?hornerE ?pab ?subrr //. + by move=> c acb; rewrite derivE derivC subr0=> hc; exists c. +move: pab; rewrite pb0=> pa0. +have: (forall rs : seq R, {subset rs <= `]a, b[} -> + (size p <= size rs)%N -> uniq rs -> all (root p) rs -> p = 0). + by move=> rs hrs; apply: poly_ltsp_roots. +elim: (size p) a b lab pa0 pb0=> [|n ihn] a b lab pa0 pb0 max_roots. + rewrite (@max_roots [::]) //=. + by exists (mid a b); rewrite ?mid_in_itv // derivE horner0. +case: (@rolle_weak a b p); rewrite // ?pa0 ?pb0 //=. +move=> c hc; case: (altP (_ =P 0))=> //= p'c0 pc0; first by exists c. +suff: { d : R | d \in `]a, c[ & (p^`()).[d] = 0 }. + case=> [d hd] p'd0; exists d=> //. + by apply: subitvPr hd; rewrite //= (itvP hc). +apply: ihn=> //; first by rewrite (itvP hc). + exact/eqP. +move=> rs hrs srs urs rrs; apply: (max_roots (c :: rs))=> //=; last exact/andP. + move=> x; rewrite in_cons; case/predU1P=> hx; first by rewrite hx. + have: x \in `]a, c[ by apply: hrs. + by apply: subitvPr; rewrite /= (itvP hc). +by rewrite urs andbT; apply/negP; move/hrs; rewrite bound_in_itv. +Qed. + +Theorem mvt : forall a b p, a < b -> + {c | c \in `]a, b[ & p.[b] - p.[a] = p^`().[c] * (b - a)}. +Proof. +move=> a b p lab. +pose q := (p.[b] - p.[a])%:P * ('X - a%:P) - (b - a)%:P * (p - p.[a]%:P). +case: (@rolle a b q)=> //. + by rewrite /q !hornerE !(subrr,mulr0) mulrC subrr. +move=> c lacb q'x0; exists c=> //. +move: q'x0; rewrite /q !derivE !(mul0r,add0r,subr0,mulr1). +by move/eqP; rewrite !hornerE mulrC subr_eq0; move/eqP. +Qed. + +Lemma deriv_sign : forall a b p, + (forall x, x \in `]a, b[ -> p^`().[x] >= 0) + -> (forall x y, (x \in `]a, b[) && (y \in `]a, b[) + -> x < y -> p.[x] <= p.[y] ). +Proof. +move=> a b p Pab x y; case/andP=> hx hy xy. +rewrite -subr_gte0; case: (@mvt x y p)=> //. +move=> c hc ->; rewrite pmulr_lge0 ?subr_gt0 ?Pab //. +by apply: subitvP hc; rewrite //= ?(itvP hx) ?(itvP hy). +Qed. + +End Prelim. + +Section MonotonictyAndRoots. + +Section NoRoot. + +Variable (p : {poly R}). + +Variables (a b : R). + +Hypothesis der_pos : forall x, x \in `]a, b[ -> (p^`()).[x] > 0. + +Lemma derp0r : 0 <= p.[a] -> forall x, x \in `]a, b] -> p.[x] > 0. +Proof. +move=> pa0 x; case/itv_dec=> ax xb; case: (mvt p ax) => c acx. +move/(canRL (@subrK _ _))->; rewrite ltr_paddr //. +by rewrite pmulr_rgt0 ?subr_gt0 // der_pos //; apply: subitvPr acx. +Qed. + +Lemma derppr : 0 < p.[a] -> forall x, x \in `[a, b] -> p.[x] > 0. +Proof. +move=> pa0 x hx; case exa: (x == a); first by rewrite (eqP exa). +case: (@mvt a x p); first by rewrite ltr_def exa (itvP hx). +move=> c hc; move/eqP; rewrite subr_eq; move/eqP->; rewrite ltr_spsaddr //. +rewrite pmulr_rgt0 ?subr_gt0 //; first by rewrite ltr_def exa (itvP hx). +by rewrite der_pos // (subitvPr _ hc) //= ?(itvP hx). +Qed. + +Lemma derp0l : p.[b] <= 0 -> forall x, x \in `[a, b[ -> p.[x] < 0. +Proof. +move=> pb0 x hx; rewrite -oppr_gte0 /=. +case: (@mvt x b p)=> //; first by rewrite (itvP hx). +move=> c hc; move/(canRL (@addKr _ _))->; rewrite ltr_spaddr ?oppr_ge0 //. +rewrite pmulr_rgt0 // ?subr_gt0 ?(itvP hx) //. +by rewrite der_pos // (subitvPl _ hc) //= (itvP hx). +Qed. + +Lemma derpnl : p.[b] < 0 -> forall x, x \in `[a, b] -> p.[x] < 0. +Proof. +move=> pbn x hx; case xb: (b == x) pbn; first by rewrite -(eqP xb). +case: (@mvt x b p); first by rewrite ltr_def xb ?(itvP hx). +move=> y hy; move/eqP; rewrite subr_eq; move/eqP->. +rewrite !ltrNge; apply: contra=> hpx; rewrite ler_paddr // ltrW //. +rewrite pmulr_rgt0 ?subr_gt0 ?(itvP hy) //. +by rewrite der_pos // (subitvPl _ hy) //= (itvP hx). +Qed. + +End NoRoot. + +Section NoRoot_sg. + +Variable (p : {poly R}). + +Variables (a b c : R). + +Hypothesis derp_neq0 : {in `]a, b[, noroot p^`()}. +Hypothesis acb : c \in `]a, b[. + +Local Notation sp'c := (sgr p^`().[c]). +Local Notation q := (sp'c *: p). + +Fact derq_pos x : x \in `]a, b[ -> (q^`()).[x] > 0. +Proof. +move=> hx; rewrite derivZ hornerZ -sgr_cp0. +rewrite sgrM sgr_id mulr_sg_eq1 derp_neq0 //=. +by apply/eqP; apply: (@polyrN0_itv `]a, b[). +Qed. + +Fact sgp x : sgr p.[x] = sp'c * sgr q.[x]. +Proof. +by rewrite hornerZ sgr_smul mulrA -expr2 sqr_sg derp_neq0 ?mul1r. +Qed. + +Fact hsgp x : 0 < q.[x] -> sgr p.[x] = sp'c. +Proof. by rewrite -sgr_cp0 sgp => /eqP->; rewrite mulr1. Qed. + +Fact hsgpN x : q.[x] < 0 -> sgr p.[x] = - sp'c. +Proof. by rewrite -sgr_cp0 sgp => /eqP->; rewrite mulrN1. Qed. + +Lemma ders0r : p.[a] = 0 -> forall x, x \in `]a, b] -> sgr p.[x] = sp'c. +Proof. +move=> pa0 x hx; rewrite hsgp // (@derp0r _ a b) //; first exact: derq_pos. +by rewrite hornerZ pa0 mulr0. +Qed. + +Lemma derspr : sgr p.[a] = sp'c -> forall x, x \in `[a, b] -> sgr p.[x] = sp'c. +Proof. +move=> spa x hx; rewrite hsgp // (@derppr _ a b) //; first exact: derq_pos. +by rewrite -sgr_cp0 hornerZ sgrM sgr_id spa -expr2 sqr_sg derp_neq0. +Qed. + +Lemma ders0l : p.[b] = 0 -> forall x, x \in `[a, b[ -> sgr p.[x] = -sp'c. +Proof. +move=> pa0 x hx; rewrite hsgpN // (@derp0l _ a b) //; first exact: derq_pos. +by rewrite hornerZ pa0 mulr0. +Qed. + +Lemma derspl : sgr p.[b] = -sp'c -> forall x, x \in `[a, b] -> sgr p.[x] = -sp'c. +Proof. +move=> spb x hx; rewrite hsgpN // (@derpnl _ a b) //; first exact: derq_pos. +by rewrite -sgr_cp0 hornerZ sgr_smul spb mulrN -expr2 sqr_sg derp_neq0. +Qed. + +End NoRoot_sg. + +Variable (p : {poly R}). + +Variables (a b : R). + +Section der_root. + +Hypothesis der_pos : forall x, x \in `]a, b[ -> (p^`()).[x] > 0. + +Lemma derp_root : a <= b -> 0 \in `]p.[a], p.[b][ -> + { r : R | + [/\ forall x, x \in `[a, r[ -> p.[x] < 0, + p.[r] = 0, + r \in `]a, b[ & + forall x, x \in `]r, b] -> p.[x] > 0] }. +Proof. +move=> leab hpab. +have /eqP hs : sgr p.[a] * sgr p.[b] == -1. + by rewrite -sgrM sgr_cp0 pmulr_llt0 ?(itvP hpab). +case: (ivt_sign leab hs) => r arb pr0; exists r; split => //; last 2 first. +- by move/eqP: pr0. +- move=> x rxb; have hd : forall t, t \in `]r, b[ -> 0 < (p^`()).[t]. + by move=> t ht; rewrite der_pos // ?(subitvPl _ ht) //= ?(itvP arb). + by rewrite (derp0r hd) ?(eqP pr0). +- move=> x rxb; have hd : forall t, t \in `]a, r[ -> 0 < (p^`()).[t]. + by move=> t ht; rewrite der_pos // ?(subitvPr _ ht) //= ?(itvP arb). + by rewrite (derp0l hd) ?(eqP pr0). +Qed. + +End der_root. + +(* Section der_root_sg. *) + +(* Hypothesis der_pos : forall x, x \in `]a, b[ -> (p^`()).[x] != 0. *) + +(* Lemma derp_root : a <= b -> sgr p.[a] != sgr p.[b] -> *) +(* { r : R | *) +(* [/\ forall x, x \in `[a, r[ -> p.[x] < 0, *) +(* p.[r] = 0, *) +(* r \in `]a, b[ & *) +(* forall x, x \in `]r, b] -> p.[x] > 0] }. *) +(* Proof. *) +(* move=> leab hpab. *) +(* have hs : sgr p.[a] * sgr p.[b] == -1. *) +(* by rewrite -sgrM sgr_cp0 mulr_lt0_gt0 ?(itvP hpab). *) +(* case: (ivt_sign ivt leab hs) => r arb pr0; exists r; split => //; last 2 first. *) +(* - by move/eqP: pr0. *) +(* - move=> x rxb; have hd : forall t, t \in `]r, b[ -> 0 < (p^`()).[t]. *) +(* by move=> t ht; rewrite der_pos // ?(subitvPl _ ht) //= ?(itvP arb). *) +(* by rewrite (derp0r hd) ?(eqP pr0). *) +(* - move=> x rxb; have hd : forall t, t \in `]a, r[ -> 0 < (p^`()).[t]. *) +(* by move=> t ht; rewrite der_pos // ?(subitvPr _ ht) //= ?(itvP arb). *) +(* by rewrite (derp0l hd) ?(eqP pr0). *) +(* Qed. *) + +(* End der_root. *) + +End MonotonictyAndRoots. + +Section RootsOn. + +Variable T : predType R. + +Definition roots_on (p : {poly R}) (i : T) (s : seq R) := + forall x, (x \in i) && (root p x) = (x \in s). + +Lemma roots_onP : forall p i s, roots_on p i s -> {in i, root p =1 mem s}. +Proof. by move=> p i s hp x hx; move: (hp x); rewrite hx. Qed. + +Lemma roots_on_in : forall p i s, + roots_on p i s -> forall x, x \in s -> x \in i. +Proof. by move=> p i s hp x; rewrite -hp; case/andP. Qed. + +Lemma roots_on_root : forall p i s, + roots_on p i s -> forall x, x \in s -> root p x. +Proof. by move=> p i s hp x; rewrite -hp; case/andP. Qed. + +Lemma root_roots_on : forall p i s, + roots_on p i s -> forall x, x \in i -> root p x -> x \in s. +Proof. by move=> p i s hp x; rewrite -hp=> ->. Qed. + +Lemma roots_on_opp : forall p i s, + roots_on (- p) i s -> roots_on p i s. +Proof. by move=> p i s hp x; rewrite -hp rootN. Qed. + +Lemma roots_on_nil : forall p i, roots_on p i [::] -> {in i, noroot p}. +Proof. +by move=> p i hp x hx; move: (hp x); rewrite in_nil hx /=; move->. +Qed. + +Lemma roots_on_same : forall s' p i s, + s =i s' -> (roots_on p i s <-> roots_on p i s'). +Proof. by move=> s' p i s hss'; split=> hs x; rewrite (hss', (I, hss')). Qed. + +End RootsOn. + + +(* (* Symmetry of center a *) *) +(* Definition symr (a x : R) := a - x. *) + +(* Lemma symr_inv : forall a, involutive (symr a). *) +(* Proof. by move=> a y; rewrite /symr opprD addrA subrr opprK add0r. Qed. *) + +(* Lemma symr_inj : forall a, injective (symr a). *) +(* Proof. by move=> a; apply: inv_inj; apply: symr_inv. Qed. *) + +(* Lemma ltr_sym : forall a x y, (symr a x < symr a y) = (y < x). *) +(* Proof. by move=> a x y; rewrite lter_add2r lter_oppr opprK. Qed. *) + +(* Lemma symr_add_itv : forall a b x, *) +(* (a < symr (a + b) x < b) = (a < x < b). *) +(* Proof. *) +(* move=> a b x; rewrite andbC. *) +(* by rewrite lter_subrA lter_add2r -lter_addlA lter_add2l. *) +(* Qed. *) + +Lemma roots_on_comp : forall p a b s, + roots_on (p \Po (-'X)) `](-b), (-a)[ + (map (-%R) s) <-> roots_on p `]a, b[ s. +Proof. +move=> p a b /= s; split=> hs x; rewrite ?root_comp ?hornerE. + move: (hs (-x)); rewrite mem_map; last exact: (inv_inj (@opprK _)). + by rewrite root_comp ?hornerE oppr_itv !opprK. +rewrite -[x]opprK oppr_itv /= mem_map; last exact: (inv_inj (@opprK _)). +by move: (hs (-x)); rewrite !opprK. +Qed. + +Lemma min_roots_on : forall p a b x s, + all (> x) s -> roots_on p `]a, b[ (x :: s) + -> [/\ x \in `]a, b[, (roots_on p `]a, x[ [::]), + (root p x) & (roots_on p `]x, b[ s)]. +Proof. +move=> p a b x s lxs hxs. +have hx: x \in `]a, b[ by rewrite (roots_on_in hxs) ?mem_head. +rewrite hx (roots_on_root hxs) ?mem_head //. +split=> // y; move: (hxs y); rewrite ?in_nil ?in_cons /=. + case hy: (y \in `]a, x[)=> //=. + rewrite (subitvPr _ hy) //= ?(itvP hx) //= => ->. + rewrite ltr_eqF ?(itvP hy) //=; apply/negP. + by move/allP: lxs=> lxs /lxs; rewrite ltrNge ?(itvP hy). +move/allP:lxs=>lxs; case eyx: (y == _)=> /=. + case/andP=> hy _; rewrite (eqP eyx). + rewrite boundl_in_itv /=; symmetry. + by apply/negP; move/lxs; rewrite ltrr. +case py0: root; rewrite !(andbT, andbF) //. +case ys: (y \in s); first by move/lxs:ys; rewrite ?inE => ->; case/andP. +move/negP; move/negP=> nhy; apply: negbTE; apply: contra nhy. +by apply: subitvPl; rewrite //= ?(itvP hx). +Qed. + +Lemma max_roots_on : forall p a b x s, + all (< x) s -> roots_on p `]a, b[ (x :: s) + -> [/\ x \in `]a, b[, (roots_on p `]x, b[ [::]), + (root p x) & (roots_on p `]a, x[ s)]. +Proof. +move=> p a b x s; move/allP=> lsx; move/roots_on_comp=> /=. +move/min_roots_on; case. + apply/allP=> y; rewrite -[y]opprK mem_map. + by move/lsx; rewrite ltr_oppr opprK. + exact: (inv_inj (@opprK _)). +rewrite oppr_itv root_comp !hornerE !opprK=> -> rxb -> rax. +by split=> //; apply/roots_on_comp. +Qed. + +Lemma roots_on_cons : forall p a b r s, + sorted <%R (r :: s) -> roots_on p `]a, b[ (r :: s) -> roots_on p `]r, b[ s. +Proof. +move=> p a b r s /= hrs hr. +have:= (order_path_min (@ltr_trans _) hrs)=> allrs. +by case: (min_roots_on allrs hr). +Qed. +(* move=> p a b r s hp hr x; apply/andP/idP. *) +(* have:= (order_path_min (@ltr_trans _) hp) => /=; case/andP=> ar1 _. *) +(* case; move/ooitvP=> rxb rpx; move: (hr x); rewrite in_cons rpx andbT. *) +(* by rewrite rxb andbT (ltr_trans ar1) 1?eq_sym ?ltr_eqF ?rxb. *) +(* move=> spx. *) +(* have xrsp: x \in r :: s by rewrite in_cons spx orbT. *) +(* rewrite (roots_on_root hr) //. *) +(* rewrite (roots_on_in hr xrsp); move: hp => /=; case/andP=> _. *) +(* by move/(order_path_min (@ltr_trans _)); move/allP; move/(_ _ spx)->. *) +(* Qed. *) + +Lemma roots_on_rcons : forall p a b r s, + sorted <%R (rcons s r) -> roots_on p `]a, b[ (rcons s r) + -> roots_on p `]a, r[ s. +Proof. +move=> p a b r s; rewrite -{1}[s]revK -!rev_cons rev_sorted /=. +move=> hrs hr. +have := (order_path_min (rev_trans (@ltr_trans _)) hrs)=> allrrs. +have allrs: (all (< r) s). + by apply/allP=> x hx; move/allP:allrrs; apply; rewrite mem_rev. +move/(@roots_on_same _ _ _ _ (r::s)):hr=>hr. +case: (max_roots_on allrs (hr _))=> //. +by move=> x; rewrite mem_rcons. +Qed. + + +(* move=> p a b r s; rewrite -{1}[s]revK -!rev_cons rev_sorted. *) +(* rewrite [r :: _]lock /=; unlock; move=> hp hr x; apply/andP/idP. *) +(* have:= (order_path_min (rev_trans (@ltr_trans _)) hp) => /=. *) +(* case/andP=> ar1 _; case; move/ooitvP=> axr rpx. *) +(* move: (hr x); rewrite mem_rcons in_cons rpx andbT axr andTb. *) +(* by rewrite ((rev_trans (@ltr_trans _) ar1)) ?ltr_eqF ?axr. *) +(* move=> spx. *) +(* have xrsp: x \in rcons s r by rewrite mem_rcons in_cons spx orbT. *) +(* rewrite (roots_on_root hr) //. *) +(* rewrite (roots_on_in hr xrsp); move: hp => /=; case/andP=> _. *) +(* move/(order_path_min (rev_trans (@ltr_trans _))); move/allP. *) +(* by move/(_ x)=> -> //; rewrite mem_rev. *) +(* Qed. *) + +Lemma no_roots_on : forall (p : {poly R}) a b, + {in `]a, b[, noroot p} -> roots_on p `]a, b[ [::]. +Proof. +move=> p a b hr x; rewrite in_nil; case hx: (x \in _) => //=. +by apply: negPf; apply: hr hx. +Qed. + +Lemma monotonic_rootN : forall (p : {poly R}) (a b : R), + {in `]a, b[, noroot p^`()} -> + ((roots_on p `]a, b[ [::]) + ({r : R | roots_on p `]a, b[ [:: r]}))%type. +Proof. +move=> p a b hp'; case: (ltrP a b); last first => leab. + by left => x; rewrite in_nil itv_gte. +wlog {hp'} hp'sg: p / forall x, x \in `]a, b[ -> sgr (p^`()).[x] = 1. + move=> hwlog. have := (polyrN0_itv hp'). + move: (mid_in_itvoo leab)=> hm /(_ _ _ hm). + case: (sgrP _.[mid a b])=> hpm. + - by move=> norm; move: (hp' _ hm); rewrite rootE hpm eqxx. + - by move/(hwlog p). + - move=> hp'N; case: (hwlog (-p))=> [x|h|[r hr]]. + * by rewrite derivE hornerN sgrN=> /hp'N->; rewrite opprK. + * by left; apply: roots_on_opp. + * by right; exists r; apply: roots_on_opp. +have hp'pos: forall x, x \in `]a, b[ -> (p^`()).[x] > 0. + by move=> x; move/hp'sg; move/eqP; rewrite sgr_cp0. +case: (lerP 0 p.[a]) => ha. +- left; apply: no_roots_on => x axb. + by rewrite rootE gtr_eqF // (@derp0r _ a b) // (subitvPr _ axb) /=. +- case: (lerP p.[b] 0) => hb. + + left => x; rewrite in_nil; apply: negbTE; case axb: (x \in `]a, b[) => //=. + by rewrite rootE ltr_eqF // (@derp0l _ a b) // (subitvPl _ axb) /=. + + case: (derp_root hp'pos (ltrW leab) _). + by rewrite ?inE; apply/andP. + move=> r [h1 h2 h3] h4; right. + exists r => x; rewrite in_cons in_nil (itv_splitU2 h3). + case exr: (x == r); rewrite ?(andbT, orbT, andbF, orbF) /=. + by rewrite rootE (eqP exr) h2 eqxx. + case px0: root; rewrite (andbT, andbF) //. + move/eqP: px0=> px0; apply/negP; case/orP=> hx. + by move: (h1 x); rewrite (subitvPl _ hx) //= px0 ltrr; move/implyP. + by move: (h4 x); rewrite (subitvPr _ hx) //= px0 ltrr; move/implyP. +Qed. + +(* Inductive polN0 : Type := PolN0 : forall p : {poly R}, p != 0 -> polN0. *) + +(* Coercion pol_of_polN0 i := let: PolN0 p _ := i in p. *) + +(* Canonical Structure polN0_subType := [subType for pol_of_polN0]. *) +(* Definition polN0_eqMixin := Eval hnf in [eqMixin of polN0 by <:]. *) +(* Canonical Structure polN0_eqType := *) +(* Eval hnf in EqType polN0 polN0_eqMixin. *) +(* Definition polN0_choiceMixin := [choiceMixin of polN0 by <:]. *) +(* Canonical Structure polN0_choiceType := *) +(* Eval hnf in ChoiceType polN0 polN0_choiceMixin. *) + +(* Todo : Lemmas about operations of intervall : + itversection, restriction and splitting *) +Lemma cat_roots_on : forall (p : {poly R}) a b x, + x \in `]a, b[ -> ~~ (root p x) + -> forall s s', + sorted <%R s -> sorted <%R s' + -> roots_on p `]a, x[ s -> roots_on p `]x, b[ s' + -> roots_on p `]a, b[ (s ++ s'). +Proof. +move=> p a b x hx /= npx0 s. +elim: s a hx => [|y s ihs] a hx s' //= ss ss'. + move/roots_on_nil=> hax hs' y. + rewrite -hs'; case py0: root; rewrite ?(andbT, andbF) //. + rewrite (itv_splitU2 hx); case: (y \in `]x, b[); rewrite ?orbF ?orbT //=. + apply/negP; case/orP; first by move/hax; rewrite py0. + by move/eqP=> exy; rewrite -exy py0 in npx0. +move/min_roots_on; rewrite order_path_min //; last exact: ltr_trans. +case=> // hy hay py0 hs hs' z. +rewrite in_cons; case ezy: (z == y)=> /=. + by rewrite (eqP ezy) py0 andbT (subitvPr _ hy) //= ?(itvP hx). +rewrite -(ihs y) //; last exact: path_sorted ss; last first. + by rewrite inE /= (itvP hx) (itvP hy). +case pz0: root; rewrite ?(andbT, andbF) //. +rewrite (@itv_splitU2 _ y); last by rewrite (subitvPr _ hy) //= (itvP hx). +rewrite ezy /=; case: (z \in `]y, b[); rewrite ?orbF ?orbT //. +by apply/negP=> hz; move: (hay z); rewrite hz pz0 in_nil. +Qed. + +CoInductive roots_spec (p : {poly R}) (i : pred R) (s : seq R) : + {poly R} -> bool -> seq R -> Type := +| Roots0 of p = 0 :> {poly R} & s = [::] : roots_spec p i s 0 true [::] +| Roots of p != 0 & roots_on p i s + & sorted <%R s : roots_spec p i s p false s. + +(* still much too long *) +Lemma itv_roots : forall (p :{poly R}) (a b : R), + {s : seq R & roots_spec p (topred `]a, b[) s p (p == 0) s}. +Proof. +move=> p a b; case p0: (_ == 0). + by rewrite (eqP p0); exists [::]; constructor. +elim: (size p) {-2}p (leqnn (size p)) p0 a b => {p} [| n ihn] p sp p0 a b. + by exists [::]; move: p0; rewrite -size_poly_eq0 -leqn0 sp. +move/negbT: (p0)=> np0. +case p'0 : (p^`() == 0). + move: p'0; rewrite -derivn1 -derivn_poly0; move/size1_polyC => pC. + exists [::]; constructor=> // x; rewrite in_nil pC rootC; apply: negPf. + by rewrite negb_and -polyC_eq0 -pC p0 orbT. +move/negbT: (p'0) => np'0. +have sizep' : (size p^`() <= n)%N. + rewrite -ltnS; apply: leq_trans sp; rewrite size_deriv prednK // lt0n. + by rewrite size_poly_eq0 p0. +case: (ihn _ sizep' p'0 a b) => sp' ih {ihn}. +case ltab : (a < b); last first. + exists [::]; constructor=> // x; rewrite in_nil. + case axb : (x \in _) => //=. + by case/andP: axb => ax xb; move: ltab; rewrite (ltr_trans ax xb). +elim: sp' a b ltab ih => [|r1 sp' hsp'] a b ltab hp'. + case: hp' np'0; rewrite ?eqxx // => np'0 hroots' _ _. + move/roots_on_nil : hroots' => hroots'. + case: (monotonic_rootN hroots') => [h| [r rh]]. + by exists [::]; constructor. + by exists [:: r]; constructor=> //=; rewrite andbT. +case: hp' np'0; rewrite ?eqxx // => np'0 hroots' /= hpath' _. +case: (min_roots_on _ hroots'). + by rewrite order_path_min //; apply: ltr_trans. +move=> hr1 har1 p'r10 hr1b. +case: (hsp' r1 b); first by rewrite (itvP hr1). + by constructor=> //; rewrite (path_sorted hpath'). +move=> s spec_s. +case: spec_s np0; rewrite ?eqxx //. +move=> np0 hroot hsort _. +move: (roots_on_nil har1). +case pr1 : (root p r1); case/monotonic_rootN => hrootsl; last 2 first. +- exists s; constructor=> //. + by rewrite -[s]cat0s; apply: (cat_roots_on hr1)=> //; rewrite pr1. +- case:hrootsl=> r hr; exists (r::s); constructor=> //=. + by rewrite -cat1s; apply: (cat_roots_on hr1)=> //; rewrite pr1. + rewrite path_min_sorted // => y; rewrite -hroot; case/andP=> hy _. + rewrite (@ltr_trans _ r1) ?(itvP hy) //. + by rewrite (itvP (roots_on_in hr (mem_head _ _))). +- exists (r1::s); constructor=> //=; last first. + rewrite path_min_sorted=> // y; rewrite -hroot. + by case/andP; move/itvP->. + move=> x; rewrite in_cons; case exr1: (x == r1)=> /=. + by rewrite (eqP exr1) pr1 andbT. + rewrite -hroot; case px: root; rewrite ?(andbT, andbF) //. + rewrite (itv_splitU2 hr1) exr1 /=. + case: (_ \in `]r1, _[); rewrite ?(orbT, orbF) //. + by apply/negP=> hx; move: (hrootsl x); rewrite hx px in_nil. +- case: hrootsl => r0 hrootsl. + move/min_roots_on:hrootsl; case=> // hr0 har0 pr0 hr0r1. + exists [:: r0, r1 & s]; constructor=> //=; last first. + rewrite (itvP hr0) /= path_min_sorted // => y. + by rewrite -hroot; case/andP; move/itvP->. + move=> y; rewrite !in_cons (itv_splitU2 hr1) (itv_splitU2 hr0). + case eyr0: (y == r0); rewrite ?(orbT, orbF, orTb, orFb). + by rewrite (eqP eyr0) pr0. + case eyr1: (y == r1); rewrite ?(orbT, orbF, orTb, orFb). + by rewrite (eqP eyr1) pr1. + rewrite -hroot; case py0: root; rewrite ?(andbF, andbT) //. + case: (_ \in `]r1, _[); rewrite ?(orbT, orbF) //. + apply/negP; case/orP=> hy; first by move: (har0 y); rewrite hy py0 in_nil. + by move: (hr0r1 y); rewrite hy py0 in_nil. +Qed. + +Definition roots (p : {poly R}) a b := projT1 (itv_roots p a b). + +Lemma rootsP : forall p a b, + roots_spec p (topred `]a, b[) (roots p a b) p (p == 0) (roots p a b). +Proof. by move=> p a b; rewrite /roots; case hp: itv_roots. Qed. + +Lemma roots0 : forall a b, roots 0 a b = [::]. +Proof. by move=> a b; case: rootsP=> //=; rewrite eqxx. Qed. + +Lemma roots_on_roots : forall p a b, p != 0 -> + roots_on p `]a, b[ (roots p a b). +Proof. by move=> a b p; case:rootsP. Qed. +Hint Resolve roots_on_roots. + +Lemma sorted_roots : forall a b p, sorted <%R (roots p a b). +Proof. by move=> p a b; case: rootsP. Qed. +Hint Resolve sorted_roots. + +Lemma path_roots : forall p a b, path <%R a (roots p a b). +Proof. +move=> p a b; case: rootsP=> //= p0 hp sp; rewrite path_min_sorted //. +by move=> y; rewrite -hp; case/andP; move/itvP->. +Qed. +Hint Resolve path_roots. + +Lemma root_is_roots : + forall (p : {poly R}) (a b : R), p != 0 -> + forall x, x \in `]a, b[ -> root p x = (x \in roots p a b). +Proof. +by move=> p a b; case: rootsP=> // p0 hs ps _ y hy /=; rewrite -hs hy. +Qed. + +Lemma root_in_roots : forall (p : {poly R}) a b, p != 0 -> + forall x, x \in `]a, b[ -> root p x -> x \in (roots p a b). +Proof. by move=> p a b p0 x axb rpx; rewrite -root_is_roots. Qed. + +Lemma root_roots : forall p a b x, x \in roots p a b -> root p x. +Proof. by move=> p a b x; case: rootsP=> // p0 <- _; case/andP. Qed. + +Lemma roots_nil : forall p a b, p != 0 -> + roots p a b = [::] -> {in `]a, b[, noroot p}. +Proof. +move=> p a b; case: rootsP=> // p0 hs ps _ s0 x axb. +by move: (hs x); rewrite s0 in_nil !axb /= => ->. +Qed. + +Lemma roots_in p a b x : x \in roots p a b -> x \in `]a, b[. +Proof. by case: rootsP=> //= np0 ron_p *; exact: (roots_on_in ron_p). Qed. + +Lemma rootsEba : forall p a b, b <= a -> roots p a b = [::]. +Proof. +move=> p a b; case: rootsP=> // p0; case: (roots _ _ _)=> [|x s] hs ps ba //; +by move: (hs x); rewrite itv_gte //= mem_head. +Qed. + +Lemma roots_on_uniq : forall p a b s1 s2, + sorted <%R s1 -> sorted <%R s2 -> + roots_on p `]a, b[ s1 -> roots_on p `]a, b[ s2 -> s1 = s2. +Proof. +move=> p a b s1. +elim: s1 p a b => [| r1 s1 ih] p a b [| r2 s2] ps1 ps2 rs1 rs2 //. +- have rpr2 : root p r2 by apply: (roots_on_root rs2); rewrite mem_head. + have abr2 : r2 \in `]a, b[ by apply: (roots_on_in rs2); rewrite mem_head. + by have:= (rs1 r2); rewrite rpr2 !abr2 in_nil. +- have rpr1 : root p r1 by apply: (roots_on_root rs1); rewrite mem_head. + have abr1 : r1 \in `]a, b[ by apply: (roots_on_in rs1); rewrite mem_head. + by have:= (rs2 r1); rewrite rpr1 !abr1 in_nil. +- have er1r2 : r1 = r2. + move: (rs1 r2); rewrite (roots_on_root rs2) ?mem_head //. + rewrite !(roots_on_in rs2) ?mem_head //= in_cons. + move/(@sym_eq _ true); case/orP => hr2; first by rewrite (eqP hr2). + move: ps1=> /=; move/(order_path_min (@ltr_trans R)); move/allP. + move/(_ r2 hr2) => h1. + move: (rs2 r1); rewrite (roots_on_root rs1) ?mem_head //. + rewrite !(roots_on_in rs1) ?mem_head //= in_cons. + move/(@sym_eq _ true); case/orP => hr1; first by rewrite (eqP hr1). + move: ps2=> /=; move/(order_path_min (@ltr_trans R)); move/allP. + by move/(_ r1 hr1); rewrite ltrNge ltrW. +congr (_ :: _) => //; rewrite er1r2 in ps1 rs1. +have h3 := (roots_on_cons ps1 rs1). +have h4 := (roots_on_cons ps2 rs2). +move: ps1 ps2=> /=; move/path_sorted=> hs1; move/path_sorted=> hs2. +exact: (ih p _ b _ hs1 hs2 h3 h4). +Qed. + +Lemma roots_eq : forall (p q : {poly R}) (a b : R), + p != 0 -> q != 0 -> + ({in `]a, b[, root p =1 root q} + <-> roots p a b = roots q a b). +Proof. +move=> p q a b p0 q0. +case hab : (a < b); last first. + split; first by rewrite !rootsEba // lerNgt hab. + move=> _ x. rewrite !inE; case/andP=> ax xb. + by move: hab; rewrite (@ltr_trans _ x) /=. +split=> hpq. + apply: (@roots_on_uniq p a b); rewrite ?path_roots ?p0 ?q0 //. + by apply: roots_on_roots. + rewrite /roots_on => x; case hx: (_ \in _). + by rewrite -hx hpq //; apply: roots_on_roots. + by rewrite /= -(andFb (q.[x] == 0)) -hx; apply: roots_on_roots. +move=> x axb /=. +by rewrite (@root_is_roots q a b) // (@root_is_roots p a b) // hpq. +Qed. + +Lemma roots_opp : forall p, roots (- p) =2 roots p. +Proof. +move=> p a b; case p0 : (p == 0); first by rewrite (eqP p0) oppr0. +by apply/roots_eq=> [||x]; rewrite ?oppr_eq0 ?p0 ?rootN. +Qed. + +Lemma no_root_roots : forall (p : {poly R}) a b, + {in `]a, b[ , noroot p} -> roots p a b = [::]. +Proof. +move=> p a b hr; case: rootsP=> // p0 hs ps. +apply: (@roots_on_uniq p a b)=> // x; rewrite in_nil. +by apply/negP; case/andP; move/hr; move/negPf->. +Qed. + +Lemma head_roots_on_ge : forall p a b s, a < b -> + roots_on p `]a, b[ s -> a < head b s. +Proof. +move=> p a b [|x s] ab //; move/(_ x). +by rewrite in_cons eqxx; case/andP; case/andP. +Qed. + +Lemma head_roots_ge : forall p a b, a < b -> a < head b (roots p a b). +Proof. +by move=> p a b; case: rootsP=> // *; apply: head_roots_on_ge. +Qed. + +Lemma last_roots_on_le : forall p a b s, a < b -> + roots_on p `]a, b[ s -> last a s < b. +Proof. +move=> p a b [|x s] ab rs //. +by rewrite (itvP (roots_on_in rs _)) //= mem_last. +Qed. + +Lemma last_roots_le : forall p a b, a < b -> last a (roots p a b) < b. +Proof. +by move=> p a b; case: rootsP=> // *; apply: last_roots_on_le. +Qed. + +Lemma roots_uniq : forall p a b s, p != 0 -> + roots_on p `]a, b[ s -> sorted <%R s -> s = roots p a b. +Proof. +move=> p a b s; case: rootsP=> // p0 hs' ps' _ hs ss. +exact: (@roots_on_uniq p a b)=> //. +Qed. + +Lemma roots_cons : forall p a b x s, + (roots p a b == x :: s) + = [&& p != 0, x \in `]a, b[, + (roots p a x == [::]), + (root p x) & (roots p x b == s)]. +Proof. +move=> p a b x s; case: rootsP=> // p0 hs' ps' /=. +apply/idP/idP. + move/eqP=> es'; move: ps' hs'; rewrite es' /= => sxs. + case/min_roots_on; first by apply: order_path_min=> //; apply: ltr_trans. + move=> -> rax px0 rxb. + rewrite px0 (@roots_uniq p a x [::]) // (@roots_uniq p x b s) ?eqxx //=. + by move/path_sorted:sxs. +case: rootsP p0=> // p0 rax sax _. +case/and3P=> hx hax; rewrite (eqP hax) in rax sax. +case: rootsP p0=> // p0 rxb sxb _. +case/andP=> px0 hxb; rewrite (eqP hxb) in rxb sxb. +rewrite [_ :: _](@roots_uniq p a b) //; last first. + rewrite /= path_min_sorted // => y. + by rewrite -(eqP hxb); move/roots_in; move/itvP->. +move=> y; rewrite (itv_splitU2 hx) !andb_orl in_cons. +case hy: (y == x); first by rewrite (eqP hy) px0 orbT. +by rewrite andFb orFb rax rxb in_nil. +Qed. + +Lemma roots_rcons : forall p a b x s, + (roots p a b == rcons s x) + = [&& p != 0, x \in `]a , b[, + (roots p x b == [::]), + (root p x) & (roots p a x == s)]. +Proof. +move=> p a b x s; case: rootsP; first by case: s. +move=> // p0 hs' ps' /=. +apply/idP/idP. + move/eqP=> es'; move: ps' hs'; rewrite es' /= => sxs. + have hsx: rcons s x =i x :: rev s. + by move=> y; rewrite mem_rcons !in_cons mem_rev. + move/(roots_on_same _ _ hsx). + case/max_roots_on. + move: sxs; rewrite -[rcons _ _]revK rev_sorted rev_rcons. + by apply: order_path_min=> u v w /=; move/(ltr_trans _); apply. + move=> -> rax px0 rxb. + move/(@roots_on_same _ s): rxb; move/(_ (mem_rev _))=> rxb. + rewrite px0 (@roots_uniq p x b [::]) // (@roots_uniq p a x s) ?eqxx //=. + move: sxs; rewrite -[rcons _ _]revK rev_sorted rev_rcons. + by move/path_sorted; rewrite -rev_sorted revK. +case: rootsP p0=> // p0 rax sax _. +case/and3P=> hx hax; rewrite (eqP hax) in rax sax. +case: rootsP p0=> // p0 rxb sxb _. +case/andP=> px0 hxb; rewrite (eqP hxb) in rxb sxb. +rewrite [rcons _ _](@roots_uniq p a b) //; last first. + rewrite -[rcons _ _]revK rev_sorted rev_rcons /= path_min_sorted. + by rewrite -rev_sorted revK. + move=> y; rewrite mem_rev; rewrite -(eqP hxb). + by move/roots_in; move/itvP->. +move=> y; rewrite (itv_splitU2 hx) mem_rcons in_cons !andb_orl. +case hy: (y == x); first by rewrite (eqP hy) px0 orbT. +by rewrite rxb rax in_nil !orbF. +Qed. + +Section NeighborHood. + +Implicit Types a b : R. + +Implicit Types p : {poly R}. + +Definition next_root (p : {poly R}) x b := if p == 0 then x else head (maxr b x) (roots p x b). + +Lemma next_root0 : forall a b, next_root 0 a b = a. +Proof. by move=> *; rewrite /next_root eqxx. Qed. + +CoInductive next_root_spec (p : {poly R}) x b : bool -> R -> Type := +| NextRootSpec0 of p = 0 : next_root_spec p x b true x +| NextRootSpecRoot y of p != 0 & p.[y] = 0 & y \in `]x, b[ + & {in `]x, y[, forall z, ~~(root p z)} : next_root_spec p x b true y +| NextRootSpecNoRoot c of p != 0 & c = maxr b x + & {in `]x, b[, forall z, ~~(root p z)} : next_root_spec p x b (p.[c] == 0) c. + +Lemma next_rootP : forall (p : {poly R}) a b, next_root_spec p a b (p.[next_root p a b] == 0) (next_root p a b). +Proof. +move=> p a b; rewrite /next_root /=. +case hs: roots=> [|x s] /=. + case: (altP (p =P 0))=> p0. + by rewrite {2}p0 hornerC eqxx; constructor; rewrite p0. + by constructor=> // y hy; apply: (roots_nil p0 hs). +move/eqP: hs; rewrite roots_cons; case/and5P=> p0 hx; move/eqP=> rap rpx rx. +rewrite (negPf p0) (rootPt rpx); constructor=> //; first by move/eqP: rpx. +by move=> y hy /=; move/(roots_nil p0): (rap); apply. +Qed. + +Lemma next_root_in : forall p a b, next_root p a b \in `[a, maxr b a]. +Proof. +move=> p a b; case: next_rootP=> [p0|y np0 py0 hy _|c np0 hc _]. +* by rewrite bound_in_itv /= ler_maxr lerr orbT. +* by apply: subitvP hy=> /=; rewrite ler_maxr !lerr. +* by rewrite hc bound_in_itv /= ler_maxr lerr orbT. +Qed. + +Lemma next_root_gt : forall p a b, a < b -> p != 0 -> next_root p a b > a. +Proof. +move=> p a b ab np0; case: next_rootP=> [p0|y _ py0 hy _|c _ -> _]. +* by rewrite p0 eqxx in np0. +* by rewrite (itvP hy). +* by rewrite maxr_l // ltrW. +Qed. + +Lemma next_noroot : forall p a b, {in `]a, (next_root p a b)[, noroot p}. +Proof. +move=> p a b z; case: next_rootP; first by rewrite itv_xx. + by move=> y np0 py0 hy hp hz; rewrite (negPf (hp _ _)). +move=> c p0 -> hp hz; rewrite (negPf (hp _ _)) //. +by case: maxrP hz; last by rewrite itv_xx. +Qed. + +Lemma is_next_root : forall p a b x, next_root_spec p a b (root p x) x -> x = next_root p a b. +Proof. +move=> p a b x []; first by move->; rewrite /next_root eqxx. + move=> y; case: next_rootP; first by move->; rewrite eqxx. + move=> y' np0 py'0 hy' hp' _ py0 hy hp. + wlog: y y' hy' hy hp' hp py0 py'0 / y <= y'. + by case/orP: (ler_total y y')=> lyy' hw; [|symmetry]; apply: hw. + case: ltrgtP=> // hyy' _; move: (hp' y). + by rewrite rootE py0 eqxx inE /= (itvP hy) hyy'; move/(_ isT). + move=> c p0 ->; case: maxrP=> hab; last by rewrite itv_gte //= ltrW. + by move=> hpz _ py0 hy; move/hpz:hy; rewrite rootE py0 eqxx. +case: next_rootP=> //; first by move->; rewrite eqxx. + by move=> y np0 py0 hy _ c _ _; move/(_ _ hy); rewrite rootE py0 eqxx. +by move=> c _ -> _ c' _ ->. +Qed. + +Definition prev_root (p : {poly R}) a x := if p == 0 then x else last (minr a x) (roots p a x). + +Lemma prev_root0 : forall a b, prev_root 0 a b = b. +Proof. by move=> *; rewrite /prev_root eqxx. Qed. + +CoInductive prev_root_spec (p : {poly R}) a x : bool -> R -> Type := +| PrevRootSpec0 of p = 0 : prev_root_spec p a x true x +| PrevRootSpecRoot y of p != 0 & p.[y] = 0 & y \in`]a, x[ + & {in `]y, x[, forall z, ~~(root p z)} : prev_root_spec p a x true y +| PrevRootSpecNoRoot c of p != 0 & c = minr a x + & {in `]a, x[, forall z, ~~(root p z)} : prev_root_spec p a x (p.[c] == 0) c. + +Lemma prev_rootP : forall (p : {poly R}) a b, prev_root_spec p a b (p.[prev_root p a b] == 0) (prev_root p a b). +Proof. +move=> p a b; rewrite /prev_root /=. +move hs: (roots _ _ _)=> s. +case: (lastP s) hs=> {s} [|s x] hs /=. + case: (altP (p =P 0))=> p0. + by rewrite {2}p0 hornerC eqxx; constructor; rewrite p0. + by constructor=> // y hy; apply: (roots_nil p0 hs). +rewrite last_rcons; move/eqP: hs. +rewrite roots_rcons; case/and5P=> p0 hx; move/eqP=> rap rpx rx. +rewrite (negPf p0) (rootPt rpx); constructor=> //; first by move/eqP: rpx. +by move=> y hy /=; move/(roots_nil p0): (rap); apply. +Qed. + +Lemma prev_root_in : forall p a b, prev_root p a b \in `[minr a b, b]. +Proof. +move=> p a b; case: prev_rootP=> [p0|y np0 py0 hy _|c np0 hc _]. +* by rewrite bound_in_itv /= ler_minl lerr orbT. +* by apply: subitvP hy=> /=; rewrite ler_minl !lerr. +* by rewrite hc bound_in_itv /= ler_minl lerr orbT. +Qed. + +Lemma prev_noroot : forall p a b, {in `](prev_root p a b), b[, noroot p}. +Proof. +move=> p a b z; case: prev_rootP; first by rewrite itv_xx. + by move=> y np0 py0 hy hp hz; rewrite (negPf (hp _ _)). +move=> c np0 ->; case: minrP=> hab; last by rewrite itv_xx. +by move=> hp hz; rewrite (negPf (hp _ _)). +Qed. + +Lemma prev_root_lt : forall p a b, a < b -> p != 0 -> prev_root p a b < b. +Proof. +move=> p a b ab np0; case: prev_rootP=> [p0|y _ py0 hy _|c _ -> _]. +* by rewrite p0 eqxx in np0. +* by rewrite (itvP hy). +* by rewrite minr_l // ltrW. +Qed. + +Lemma is_prev_root : forall p a b x, prev_root_spec p a b (root p x) x -> x = prev_root p a b. +Proof. +move=> p a b x []; first by move->; rewrite /prev_root eqxx. + move=> y; case: prev_rootP; first by move->; rewrite eqxx. + move=> y' np0 py'0 hy' hp' _ py0 hy hp. + wlog: y y' hy' hy hp' hp py0 py'0 / y <= y'. + by case/orP: (ler_total y y')=> lyy' hw; [|symmetry]; apply: hw. + case: ltrgtP=> // hyy' _; move/implyP: (hp y'). + by rewrite rootE py'0 eqxx inE /= (itvP hy') hyy'. + by move=> c _ _ hpz _ py0 hy; move/hpz:hy; rewrite rootE py0 eqxx. +case: prev_rootP=> //; first by move->; rewrite eqxx. + move=> y ? py0 hy _ c _ ->; case: minrP hy=> hab; last first. + by rewrite itv_gte //= ltrW. + by move=> hy; move/(_ _ hy); rewrite rootE py0 eqxx. +by move=> c _ -> _ c' _ ->. +Qed. + +Definition neighpr p a b := `]a, (next_root p a b)[. + +Definition neighpl p a b := `](prev_root p a b), b[. + +Lemma neighpl_root : forall p a x, {in neighpl p a x, noroot p}. +Proof. exact: prev_noroot. Qed. + +Lemma sgr_neighplN : forall p a x, ~~root p x -> + {in neighpl p a x, forall y, (sgr p.[y] = sgr p.[x])}. +Proof. +rewrite /neighpl=> p a x nrpx /= y hy. +apply: (@polyrN0_itv `[y, x]); do ?by rewrite bound_in_itv /= (itvP hy). +move=> z; rewrite (@itv_splitU _ x false) ?itv_xx /=; last first. +(* Todo : Lemma itv_splitP *) + by rewrite bound_in_itv /= (itvP hy). +rewrite orbC => /predU1P[-> // | hz]. +rewrite (@prev_noroot _ a x) //. +by apply: subitvPl hz; rewrite /= (itvP hy). +Qed. + +Lemma sgr_neighpl_same : forall p a x, + {in neighpl p a x &, forall y z, (sgr p.[y] = sgr p.[z])}. +Proof. +by rewrite /neighpl=> p x b y z *; apply: (polyrN0_itv (@prev_noroot p x b)). +Qed. + +Lemma neighpr_root : forall p x b, {in neighpr p x b, noroot p}. +Proof. exact: next_noroot. Qed. + +Lemma sgr_neighprN : forall p x b, p.[x] != 0 -> + {in neighpr p x b, forall y, (sgr p.[y] = sgr p.[x])}. +Proof. +rewrite /neighpr=> p x b nrpx /= y hy; symmetry. +apply: (@polyrN0_itv `[x, y]); do ?by rewrite bound_in_itv /= (itvP hy). +move=> z; rewrite (@itv_splitU _ x true) ?itv_xx /=; last first. +(* Todo : Lemma itv_splitP *) + by rewrite bound_in_itv /= (itvP hy). +case/orP=> [|hz]; first by rewrite inE /=; move/eqP->. +rewrite (@next_noroot _ x b) //. +by apply: subitvPr hz; rewrite /= (itvP hy). +Qed. + +Lemma sgr_neighpr_same : forall p x b, + {in neighpr p x b &, forall y z, (sgr p.[y] = sgr p.[z])}. +Proof. +by rewrite /neighpl=> p x b y z *; apply: (polyrN0_itv (@next_noroot p x b)). +Qed. + +Lemma uniq_roots : forall a b p, uniq (roots p a b). +Proof. +move=> a b p; case p0: (p == 0); first by rewrite (eqP p0) roots0. +by apply: (@sorted_uniq _ <%R); [exact: ltr_trans | exact: ltrr|]. +Qed. + +Hint Resolve uniq_roots. + +Lemma in_roots : forall p a b, forall x : R, + (x \in roots p a b) = [&& root p x, x \in `]a, b[ & p != 0]. +Proof. +move=> p a b x; case: rootsP=> //=; first by rewrite in_nil !andbF. +by move=> p0 hr sr; rewrite andbT -hr andbC. +Qed. + +(* Todo : move to polyorder => need char 0 *) +Lemma gdcop_eq0 : forall p q, (gdcop p q == 0) = (q == 0) && (p != 0). +Proof. +move=> p q; case: (eqVneq q 0) => [-> | q0]. + rewrite gdcop0 /= eqxx /=. + by case: (eqVneq p 0) => [-> | pn0]; rewrite ?(negPf pn0) eqxx ?oner_eq0. +rewrite /gdcop; move: {-1}(size q) (leqnn (size q))=> k hk. +case: (eqVneq p 0) => [-> | p0]. + rewrite eqxx andbF; apply: negPf. + elim: k q q0 {hk} => [|k ihk] q q0 /=; first by rewrite eqxx oner_eq0. + case: ifP=> _ //. + apply: ihk; rewrite gcdp0 divpp ?q0 // polyC_eq0; exact: lc_expn_scalp_neq0. +rewrite p0 (negPf q0) /=; apply: negPf. +elim: k q q0 hk => [|k ihk] /= q q0 hk. + by move: hk q0; rewrite leqn0 size_poly_eq0; move->. +case: ifP=> cpq; first by rewrite (negPf q0). +apply: ihk. + rewrite divpN0; last by rewrite gcdp_eq0 negb_and q0. + by rewrite dvdp_leq // dvdp_gcdl. +rewrite -ltnS; apply: leq_trans hk; move: (dvdp_gcdl q p); rewrite dvdp_eq. +move/eqP=> eqq; move/(f_equal (fun x : {poly R} => size x)): (eqq). +rewrite size_scale; last exact: lc_expn_scalp_neq0. +have gcdn0 : gcdp q p != 0 by rewrite gcdp_eq0 negb_and q0. +have qqn0 : q %/ gcdp q p != 0. + apply: contraTneq q0; rewrite negbK => e. + move: (scaler_eq0 (lead_coef (gcdp q p) ^+ scalp q (gcdp q p)) q). + by rewrite (negPf (lc_expn_scalp_neq0 _ _)) /=; move<-; rewrite eqq e mul0r. +move->; rewrite size_mul //; case sgcd: (size (gcdp q p)) => [|n]. + by move/eqP: sgcd gcdn0; rewrite size_poly_eq0; move->. +case: n sgcd => [|n]; first by move/eqP; rewrite size_poly_eq1 gcdp_eqp1 cpq. +by rewrite addnS /= -{1}[size (_ %/ _)]addn0 ltn_add2l. +Qed. + +Lemma roots_mul : forall a b, a < b -> forall p q, + p != 0 -> q != 0 -> perm_eq (roots (p*q) a b) + (roots p a b ++ roots ((gdcop p q)) a b). +Proof. +move=> a b hab p q np0 nq0. +apply: uniq_perm_eq; first exact: uniq_roots. + rewrite cat_uniq ?uniq_roots andbT /=; apply/hasPn=> x /=. + move/root_roots; rewrite root_gdco //; case/andP=> _. + by rewrite in_roots !negb_and=> ->. +move=> x; rewrite mem_cat !in_roots root_gdco //. +rewrite rootM mulf_eq0 gdcop_eq0 negb_and. +case: (x \in `]_, _[); last by rewrite !andbF. +by rewrite negb_or !np0 !nq0 !andbT /=; do 2?case: root=> //=. +Qed. + +Lemma roots_mul_coprime : forall a b, a < b -> forall p q, + p != 0 -> q != 0 -> coprimep p q -> + perm_eq (roots (p * q) a b) + (roots p a b ++ roots q a b). +Proof. +move=> a b hab p q np0 nq0 cpq. +rewrite (perm_eq_trans (roots_mul hab np0 nq0)) //. +suff ->: roots (gdcop p q) a b = roots q a b by apply: perm_eq_refl. +case: gdcopP=> r rq hrp; move/(_ q (dvdpp _)). +rewrite coprimep_sym; move/(_ cpq)=> qr. +have erq : r %= q by rewrite /eqp rq qr. +(* Todo : relate eqp with roots *) +apply/roots_eq=> // [|x hx]; last exact: eqp_root. +by rewrite -size_poly_eq0 (eqp_size erq) size_poly_eq0. +Qed. + + +Lemma next_rootM : forall a b (p q : {poly R}), + next_root (p * q) a b = minr (next_root p a b) (next_root q a b). +Proof. +move=> a b p q; symmetry; apply: is_next_root. +wlog: p q / next_root p a b <= next_root q a b. + case: minrP=> hpq; first by move/(_ _ _ hpq); case: minrP hpq. + by move/(_ _ _ (ltrW hpq)); rewrite mulrC minrC; case: minrP hpq. +case: minrP=> //; case: next_rootP=> [|y np0 py0 hy|c np0 ->] hp hpq _. +* by rewrite hp mul0r root0; constructor. +* rewrite rootM; move/rootP:(py0)->; constructor=> //. + - by rewrite mulf_neq0 //; case: next_rootP hpq; rewrite // (itvP hy). + - by rewrite hornerM py0 mul0r. + - move=> z hz /=; rewrite rootM negb_or ?hp //. + by rewrite (@next_noroot _ a b) //; apply: subitvPr hz. +* case: (altP (q =P 0))=> q0. + move: hpq; rewrite q0 mulr0 root0 next_root0 ler_maxl lerr andbT. + by move=> hba; rewrite maxr_r //; constructor. + constructor=> //; first by rewrite mulf_neq0. + move=> z hz /=; rewrite rootM negb_or ?hp //. + rewrite (@next_noroot _ a b) //; apply: subitvPr hz=> /=. + by move: hpq; rewrite ler_maxl; case/andP. +Qed. + +Lemma neighpr_mul : forall a b p q, + (neighpr (p * q) a b) =i [predI (neighpr p a b) & (neighpr q a b)]. +Proof. +move=> a b p q x; rewrite inE /= !inE /= next_rootM. +by case: (a < x); rewrite // ltr_minr. +Qed. + +Lemma prev_rootM : forall a b (p q : {poly R}), + prev_root (p * q) a b = maxr (prev_root p a b) (prev_root q a b). +Proof. +move=> a b p q; symmetry; apply: is_prev_root. +wlog: p q / prev_root p a b >= prev_root q a b. + case: maxrP=> hpq; first by move/(_ _ _ hpq); case: maxrP hpq. + by move/(_ _ _ (ltrW hpq)); rewrite mulrC maxrC; case: maxrP hpq. +case: maxrP=> //; case: (@prev_rootP p)=> [|y np0 py0 hy|c np0 ->] hp hpq _. +* by rewrite hp mul0r root0; constructor. +* rewrite rootM; move/rootP:(py0)->; constructor=> //. + - by rewrite mulf_neq0 //; case: prev_rootP hpq; rewrite // (itvP hy). + - by rewrite hornerM py0 mul0r. + - move=> z hz /=; rewrite rootM negb_or ?hp //. + by rewrite (@prev_noroot _ a b) //; apply: subitvPl hz. +* case: (altP (q =P 0))=> q0. + move: hpq; rewrite q0 mulr0 root0 prev_root0 ler_minr lerr andbT. + by move=> hba; rewrite minr_r //; constructor. + constructor=> //; first by rewrite mulf_neq0. + move=> z hz /=; rewrite rootM negb_or ?hp //. + rewrite (@prev_noroot _ a b) //; apply: subitvPl hz=> /=. + by move: hpq; rewrite ler_minr; case/andP. +Qed. + +Lemma neighpl_mul : forall a b p q, + (neighpl (p * q) a b) =i [predI (neighpl p a b) & (neighpl q a b)]. +Proof. +move=> a b p q x; rewrite !inE /= prev_rootM. +by case: (x < b); rewrite // ltr_maxl !(andbT, andbF). +Qed. + +Lemma neighpr_wit : forall p x b, x < b -> p != 0 -> {y | y \in neighpr p x b}. +Proof. +move=> p x b xb; exists (mid x (next_root p x b)). +by rewrite mid_in_itv //= next_root_gt. +Qed. + +Lemma neighpl_wit : forall p a x, a < x -> p != 0 -> {y | y \in neighpl p a x}. +Proof. +move=> p a x xb; exists (mid (prev_root p a x) x). +by rewrite mid_in_itv //= prev_root_lt. +Qed. + +End NeighborHood. + +Section SignRight. + +Definition sgp_right (p : {poly R}) x := + let fix aux (p : {poly R}) n := + if n is n'.+1 + then if ~~ root p x + then sgr p.[x] + else aux p^`() n' + else 0 + in aux p (size p). + +Lemma sgp_right0 : forall x, sgp_right 0 x = 0. +Proof. by move=> x; rewrite /sgp_right size_poly0. Qed. + +Lemma sgr_neighpr : forall b p x, + {in neighpr p x b, forall y, (sgr p.[y] = sgp_right p x)}. +Proof. +move=> b p x. +elim: (size p) {-2}p (leqnn (size p))=> [|n ihn] {p} p. + rewrite leqn0 size_poly_eq0 /neighpr; move/eqP=> -> /=. + by move=>y; rewrite next_root0 itv_xx. +rewrite leq_eqVlt ltnS; case/orP; last exact: ihn. +move/eqP=> sp; rewrite /sgp_right sp /=. +case px0: root=> /=; last first. + move=> y; rewrite/neighpr => hy /=; symmetry. + apply: (@polyrN0_itv `[x, y]); do ?by rewrite bound_in_itv /= (itvP hy). + move=> z; rewrite (@itv_splitU _ x true) ?bound_in_itv /= ?(itvP hy) //. + rewrite itv_xx /=; case/predU1P=> hz; first by rewrite hz px0. + rewrite (@next_noroot p x b) //. + by apply: subitvPr hz=> /=; rewrite (itvP hy). +have <-: size p^`() = n by rewrite size_deriv sp. +rewrite -/(sgp_right p^`() x). +move=> y; rewrite /neighpr=> hy /=. +case: (@neighpr_wit (p * p^`()) x b)=> [||m hm]. +* case: next_rootP hy; first by rewrite itv_xx. + by move=> ? ? ?; move/itvP->. + by move=> c p0 -> _; case: maxrP=> _; rewrite ?itv_xx //; move/itvP->. +* rewrite mulf_neq0 //. + by move/eqP:sp; apply: contraTneq=> ->; rewrite size_poly0. + (* Todo : a lemma for this *) + move: (size_deriv p); rewrite sp /=; move/eqP; apply: contraTneq=> ->. + rewrite size_poly0; apply: contraTneq px0=> hn; rewrite -hn in sp. + by move/eqP: sp; case/size_poly1P=> c nc0 ->; rewrite rootC. +* move: hm; rewrite neighpr_mul /neighpr inE /=; case/andP=> hmp hmp'. + rewrite (polyrN0_itv _ hmp) //; last exact: next_noroot. + rewrite (@ders0r p x m (mid x m)) ?(eqP px0) ?mid_in_itv ?bound_in_itv //; + rewrite /= ?(itvP hmp) //; last first. + move=> u hu /=; rewrite (@next_noroot _ x b) //. + by apply: subitvPr hu; rewrite /= (itvP hmp'). + rewrite ihn ?size_deriv ?sp /neighpr //. + by rewrite (subitvP _ (@mid_in_itv _ true true _ _ _)) //= ?lerr (itvP hmp'). +Qed. + +Lemma sgr_neighpl : forall a p x, + {in neighpl p a x, forall y, + (sgr p.[y] = (-1) ^+ (odd (\mu_x p)) * sgp_right p x) + }. +Proof. +move=> a p x. +elim: (size p) {-2}p (leqnn (size p))=> [|n ihn] {p} p. + rewrite leqn0 size_poly_eq0 /neighpl; move/eqP=> -> /=. + by move=>y; rewrite prev_root0 itv_xx. +rewrite leq_eqVlt ltnS; case/orP; last exact: ihn. +move/eqP=> sp; rewrite /sgp_right sp /=. +case px0: root=> /=; last first. + move=> y; rewrite/neighpl => hy /=; symmetry. + move: (negbT px0); rewrite -mu_gt0; last first. + by apply: contraFN px0; move/eqP->; rewrite rootC. + rewrite -leqNgt leqn0; move/eqP=> -> /=; rewrite expr0 mul1r. + symmetry; apply: (@polyrN0_itv `[y, x]); + do ?by rewrite bound_in_itv /= (itvP hy). + move=> z; rewrite (@itv_splitU _ x false) ?bound_in_itv /= ?(itvP hy) //. + rewrite itv_xx /= orbC; case/predU1P=> hz; first by rewrite hz px0. + rewrite (@prev_noroot p a x) //. + by apply: subitvPl hz=> /=; rewrite (itvP hy). +have <-: size p^`() = n by rewrite size_deriv sp. +rewrite -/(sgp_right p^`() x). +move=> y; rewrite /neighpl=> hy /=. +case: (@neighpl_wit (p * p^`()) a x)=> [||m hm]. +* case: prev_rootP hy; first by rewrite itv_xx. + by move=> ? ? ?; move/itvP->. + by move=> c p0 -> _; case: minrP=> _; rewrite ?itv_xx //; move/itvP->. +* rewrite mulf_neq0 //. + by move/eqP:sp; apply: contraTneq=> ->; rewrite size_poly0. + (* Todo : a lemma for this *) + move: (size_deriv p); rewrite sp /=; move/eqP; apply: contraTneq=> ->. + rewrite size_poly0; apply: contraTneq px0=> hn; rewrite -hn in sp. + by move/eqP: sp; case/size_poly1P=> c nc0 ->; rewrite rootC. +* move: hm; rewrite neighpl_mul /neighpl inE /=; case/andP=> hmp hmp'. + rewrite (polyrN0_itv _ hmp) //; last exact: prev_noroot. + rewrite (@ders0l p m x (mid m x)) ?(eqP px0) ?mid_in_itv ?bound_in_itv //; + rewrite /= ?(itvP hmp) //; last first. + move=> u hu /=; rewrite (@prev_noroot _ a x) //. + by apply: subitvPl hu; rewrite /= (itvP hmp'). + rewrite ihn ?size_deriv ?sp /neighpl //; last first. + by rewrite (subitvP _ (@mid_in_itv _ true true _ _ _)) //= ?lerr (itvP hmp'). + rewrite mu_deriv // odd_sub ?mu_gt0 //=; last by rewrite -size_poly_eq0 sp. + by rewrite signr_addb /= mulrN1 mulNr opprK. +Qed. + +Lemma sgp_right_deriv : forall (p : {poly R}) x, root p x -> + sgp_right p x = sgp_right (p^`()) x. +Proof. +move=> p; elim: (size p) {-2}p (erefl (size p)) => {p} [p|sp hp p hsp x]. + by move/eqP; rewrite size_poly_eq0; move/eqP=> -> x _; rewrite derivC. +by rewrite /sgp_right size_deriv hsp /= => ->. +Qed. + +Lemma sgp_rightNroot : forall (p : {poly R}) x, + ~~ root p x -> sgp_right p x = sgr p.[x]. +Proof. +move=> p x nrpx; rewrite /sgp_right; case hsp: (size _)=> [|sp]. + by move/eqP:hsp; rewrite size_poly_eq0; move/eqP->; rewrite hornerC sgr0. +by rewrite nrpx. +Qed. + +Lemma sgp_right_mul : forall p q x, + sgp_right (p * q) x = sgp_right p x * sgp_right q x. +Proof. +move=> p q x. +case: (altP (q =P 0))=> q0; first by rewrite q0 /sgp_right !(size_poly0,mulr0). +case: (altP (p =P 0))=> p0; first by rewrite p0 /sgp_right !(size_poly0,mul0r). +case: (@neighpr_wit (p * q) x (1 + x))=> [||m hpq]; do ?by rewrite mulf_neq0. + by rewrite ltr_spaddl ?ltr01. +rewrite -(@sgr_neighpr (1 + x) _ _ m) //. +move: hpq; rewrite neighpr_mul inE /=; case/andP=> hp hq. +by rewrite hornerM sgrM !(@sgr_neighpr (1 + x) _ x) /neighpr. +Qed. + +Lemma sgp_right_scale c p x : sgp_right (c *: p) x = sgr c * sgp_right p x. +Proof. +case c0: (c == 0); first by rewrite (eqP c0) scale0r sgr0 mul0r sgp_right0. +by rewrite -mul_polyC sgp_right_mul sgp_rightNroot ?hornerC ?rootC ?c0. +Qed. + +Lemma sgp_right_square : forall p x, p != 0 -> sgp_right p x * sgp_right p x = 1. +Proof. +move=> p x np0; case: (@neighpr_wit p x (1 + x))=> [||m hpq] //. + by rewrite ltr_spaddl ?ltr01. +rewrite -(@sgr_neighpr (1 + x) _ _ m) //. +by rewrite -expr2 sqr_sg (@next_noroot _ x (1 + x)). +Qed. + +Lemma sgp_right_rec p x : sgp_right p x = + if p == 0 then 0 else + if ~~ root p x then sgr p.[x] + else sgp_right p^`() x. +Proof. +rewrite /sgp_right; case hs: size=> [|s]; rewrite -size_poly_eq0 hs //=. +by rewrite size_deriv hs. +Qed. + +Lemma sgp_right_addp0 : forall (p q : {poly R}) x, q != 0 -> + (\mu_x p > \mu_x q)%N -> sgp_right (p + q) x = sgp_right q x. +Proof. +move=> p q x nq0; move hm: (\mu_x q)=> m. +elim: m p q nq0 hm => [|mq ihmq] p q nq0 hmq; case hmp: (\mu_x p)=> // [mp]; + do[ + rewrite ltnS=> hm; + rewrite sgp_right_rec {1}addrC; + rewrite GRing.Theory.addr_eq0]. (* Todo : fix this ! *) + case: (altP (_ =P _))=> hqp. + move: (nq0); rewrite {1}hqp oppr_eq0=> np0. + rewrite sgp_right_rec (negPf nq0) -mu_gt0 // hmq /=. + apply/eqP; rewrite eq_sym hqp hornerN sgrN. + by rewrite oppr_eq0 sgr_eq0 -[_ == _]mu_gt0 ?hmp. + rewrite rootE hornerD. + have ->: p.[x] = 0. + apply/eqP; rewrite -[_ == _]mu_gt0 ?hmp //. + by move/eqP: hmp; apply: contraTneq=> ->; rewrite mu0. + symmetry; rewrite sgp_right_rec (negPf nq0) add0r. + by rewrite -/(root _ _) -mu_gt0 // hmq. +case: (altP (_ =P _))=> hqp. + by move: hm; rewrite -ltnS -hmq -hmp hqp mu_opp ltnn. +have px0: p.[x] = 0. + apply/rootP; rewrite -mu_gt0 ?hmp //. + by move/eqP: hmp; apply: contraTneq=> ->; rewrite mu0. +have qx0: q.[x] = 0 by apply/rootP; rewrite -mu_gt0 ?hmq //. +rewrite rootE hornerD px0 qx0 add0r eqxx /=; symmetry. +rewrite sgp_right_rec rootE (negPf nq0) qx0 eqxx /=. +rewrite derivD ihmq // ?mu_deriv ?rootE ?px0 ?qx0 ?hmp ?hmq ?subn1 //. +apply: contra nq0; rewrite -size_poly_eq0 size_deriv. +case hsq: size=> [|sq] /=. + by move/eqP: hsq; rewrite size_poly_eq0. +move/eqP=> sq0; move/eqP: hsq qx0; rewrite sq0; case/size_poly1P=> c c0 ->. +by rewrite hornerC; move/eqP; rewrite (negPf c0). +Qed. + +End SignRight. + +(* redistribute some of what follows with in the file *) +Section PolyRCFPdiv. +Import Pdiv.Ring Pdiv.ComRing. + +Lemma sgp_rightc (x c : R) : sgp_right c%:P x = sgr c. +Proof. +rewrite /sgp_right size_polyC. +case cn0: (_ == 0)=> /=; first by rewrite (eqP cn0) sgr0. +by rewrite rootC hornerC cn0. +Qed. + +Lemma sgp_right_eq0 (x : R) p : (sgp_right p x == 0) = (p == 0). +Proof. +case: (altP (p =P 0))=> p0; first by rewrite p0 sgp_rightc sgr0 eqxx. +rewrite /sgp_right. +elim: (size p) {-2}p (erefl (size p)) p0=> {p} [|sp ihsp] p esp p0. + by move/eqP:esp; rewrite size_poly_eq0 (negPf p0). +rewrite esp /=; case px0: root=> //=; rewrite ?sgr_cp0 ?px0//. +have hsp: sp = size p^`() by rewrite size_deriv esp. +rewrite hsp ihsp // -size_poly_eq0 -hsp; apply/negP; move/eqP=> sp0. +move: px0; rewrite root_factor_theorem. +by move=> /rdvdp_leq // /(_ p0); rewrite size_XsubC esp sp0. +Qed. + +(* :TODO: backport to polydiv *) +Lemma lc_expn_rscalp_neq0 (p q : {poly R}): lead_coef q ^+ rscalp p q != 0. +Proof. +case: (eqVneq q 0) => [->|nzq]; last by rewrite expf_neq0 ?lead_coef_eq0. +by rewrite /rscalp unlock /= eqxx /= expr0 oner_neq0. +Qed. +Notation lcn_neq0 := lc_expn_rscalp_neq0. + +Lemma sgp_right_mod : forall p q x, (\mu_x p < \mu_x q)%N -> + sgp_right (rmodp p q) x = (sgr (lead_coef q)) ^+ (rscalp p q) * sgp_right p x. +Proof. +move=> p q x mupq; case p0: (p == 0). + by rewrite (eqP p0) rmod0p !sgp_right0 mulr0. +have qn0 : q != 0. + by apply/negP; move/eqP=> q0; rewrite q0 mu0 ltn0 in mupq. +move/eqP: (rdivp_eq q p). +rewrite eq_sym (can2_eq (addKr _ ) (addNKr _)); move/eqP->. +case qpq0: ((rdivp p q) == 0). + by rewrite (eqP qpq0) mul0r oppr0 add0r sgp_right_scale // sgrX. +rewrite sgp_right_addp0 ?sgp_right_scale ?sgrX //. + by rewrite scaler_eq0 negb_or p0 lcn_neq0. +rewrite mu_mulC ?lcn_neq0 // mu_opp mu_mul ?mulf_neq0 ?qpq0 //. +by rewrite ltn_addl. +Qed. + +Lemma rootsC (a b c : R) : roots c%:P a b = [::]. +Proof. +case: (altP (c =P 0))=> hc; first by rewrite hc roots0. +by apply: no_root_roots=> x hx; rewrite rootC. +Qed. + +Lemma rootsZ a b c p : c != 0 -> roots (c *: p) a b = roots p a b. +Proof. +have [->|p_neq0 c_neq0] := eqVneq p 0; first by rewrite scaler0. +by apply/roots_eq => [||x axb]; rewrite ?scaler_eq0 ?(negPf c_neq0) ?rootZ. +Qed. + +Lemma root_bigrgcd (x : R) (ps : seq {poly R}) : + root (\big[(@rgcdp _)/0]_(p <- ps) p) x = all (root^~ x) ps. +Proof. +elim: ps; first by rewrite big_nil root0. +move=> p ps ihp; rewrite big_cons /=. +by rewrite (eqp_root (eqp_rgcd_gcd _ _)) root_gcd ihp. +Qed. + +Definition rootsR p := roots p (- cauchy_bound p) (cauchy_bound p). + +Lemma roots_on_rootsR p : p != 0 -> roots_on p `]-oo, +oo[ (rootsR p). +Proof. +rewrite /rootsR => p_neq0 x /=; rewrite -roots_on_roots // andbC. +by have [/(cauchy_boundP p_neq0) /=|//] := altP rootP; rewrite ltr_norml. +Qed. + +Lemma rootsR0 : rootsR 0 = [::]. Proof. exact: roots0. Qed. + +Lemma rootsRC c : rootsR c%:P = [::]. Proof. exact: rootsC. Qed. + +Lemma rootsRP p a b : + {in `]-oo, a], noroot p} -> {in `[b , +oo[, noroot p} -> + roots p a b = rootsR p. +Proof. +move=> rpa rpb. +have [->|p_neq0] := eqVneq p 0; first by rewrite rootsR0 roots0. +apply: (eq_sorted_irr (@ltr_trans _)); rewrite ?sorted_roots // => x. +rewrite -roots_on_rootsR -?roots_on_roots //=. +have [rpx|] := boolP (root _ _); rewrite ?(andbT, andbF) //. +apply: contraLR rpx; rewrite inE negb_and -!lerNgt. +by move=> /orP[/rpa //|xb]; rewrite rpb // inE andbT. +Qed. + +Lemma sgp_pinftyP x (p : {poly R}) : {in `[x , +oo[, noroot p} -> + {in `[x, +oo[, forall y, sgr p.[y] = sgp_pinfty p}. +Proof. +rewrite /sgp_pinfty; wlog lp_gt0 : x p / lead_coef p > 0 => [hwlog|rpx y Hy]. + have [|/(hwlog x p) //|/eqP] := ltrgtP (lead_coef p) 0; last first. + by rewrite lead_coef_eq0 => /eqP -> ? ? ?; rewrite lead_coef0 horner0. + rewrite -[p]opprK lead_coef_opp oppr_cp0 => /(hwlog x _) Hp HNp y Hy. + by rewrite hornerN !sgrN Hp => // z /HNp; rewrite rootN. +have [z Hz] := poly_pinfty_gt_lc lp_gt0. +have {Hz} Hz u : u \in `[z, +oo[ -> Num.sg p.[u] = 1. + by rewrite inE andbT => /Hz pu_ge1; rewrite gtr0_sg // (ltr_le_trans lp_gt0). +rewrite (@polyrN0_itv _ _ rpx (maxr y z)) ?inE ?ler_maxr ?(itvP Hy) //. +by rewrite Hz ?gtr0_sg // inE ler_maxr lerr orbT. +Qed. + +Lemma sgp_minftyP x (p : {poly R}) : {in `]-oo, x], noroot p} -> + {in `]-oo, x], forall y, sgr p.[y] = sgp_minfty p}. +Proof. +move=> rpx y Hy; rewrite -sgp_pinfty_sym. +have -> : p.[y] = (p \Po -'X).[-y] by rewrite horner_comp !hornerE opprK. +apply: (@sgp_pinftyP (- x)); last by rewrite inE ler_opp2 (itvP Hy). +by move=> z Hz /=; rewrite root_comp !hornerE rpx // inE ler_oppl (itvP Hz). +Qed. + +Lemma odd_poly_root (p : {poly R}) : ~~ odd (size p) -> {x | root p x}. +Proof. +move=> size_p_even. +have [->|p_neq0] := altP (p =P 0); first by exists 0; rewrite root0. +pose b := cauchy_bound p. +have [] := @ivt_sign p (-b) b; last by move=> x _; exists x. + by rewrite ge0_cp // ?cauchy_bound_ge0. +rewrite (sgp_minftyP (le_cauchy_bound p_neq0)) ?bound_in_itv //. +rewrite (sgp_pinftyP (ge_cauchy_bound p_neq0)) ?bound_in_itv //. +move: size_p_even; rewrite polySpred //= negbK /sgp_minfty -signr_odd => ->. +by rewrite expr1 mulN1r sgrN mulNr -expr2 sqr_sg lead_coef_eq0 p_neq0. +Qed. +End PolyRCFPdiv. + +End PolyRCF. diff --git a/mathcomp/real_closed/qe_rcf.v b/mathcomp/real_closed/qe_rcf.v new file mode 100644 index 0000000..1791aca --- /dev/null +++ b/mathcomp/real_closed/qe_rcf.v @@ -0,0 +1,1008 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import finfun path matrix. +Require Import bigop ssralg poly polydiv ssrnum zmodp div ssrint. +Require Import polyorder polyrcf interval polyXY. +Require Import qe_rcf_th ordered_qelim mxtens. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing.Theory Num.Theory. + +Local Open Scope nat_scope. +Local Open Scope ring_scope. + +Import ord. + +Section QF. + +Variable R : Type. + +Inductive term : Type := +| Var of nat +| Const of R +| NatConst of nat +| Add of term & term +| Opp of term +| NatMul of term & nat +| Mul of term & term +| Exp of term & nat. + +Inductive formula : Type := +| Bool of bool +| Equal of term & term +| Lt of term & term +| Le of term & term +| And of formula & formula +| Or of formula & formula +| Implies of formula & formula +| Not of formula. + +Coercion rterm_to_term := fix loop (t : term) : GRing.term R := + match t with + | Var x => GRing.Var _ x + | Const x => GRing.Const x + | NatConst n => GRing.NatConst _ n + | Add u v => GRing.Add (loop u) (loop v) + | Opp u => GRing.Opp (loop u) + | NatMul u n => GRing.NatMul (loop u) n + | Mul u v => GRing.Mul (loop u) (loop v) + | Exp u n => GRing.Exp (loop u) n + end. + +Coercion qfr_to_formula := fix loop (f : formula) : ord.formula R := + match f with + | Bool b => ord.Bool b + | Equal x y => ord.Equal x y + | Lt x y => ord.Lt x y + | Le x y => ord.Le x y + | And f g => ord.And (loop f) (loop g) + | Or f g => ord.Or (loop f) (loop g) + | Implies f g => ord.Implies (loop f) (loop g) + | Not f => ord.Not (loop f) + end. + +Definition to_rterm := fix loop (t : GRing.term R) : term := + match t with + | GRing.Var x => Var x + | GRing.Const x => Const x + | GRing.NatConst n => NatConst n + | GRing.Add u v => Add (loop u) (loop v) + | GRing.Opp u => Opp (loop u) + | GRing.NatMul u n => NatMul (loop u) n + | GRing.Mul u v => Mul (loop u) (loop v) + | GRing.Exp u n => Exp (loop u) n + | _ => NatConst 0 + end. + +End QF. + +Bind Scope qf_scope with term. +Bind Scope qf_scope with formula. +Arguments Scope Add [_ qf_scope qf_scope]. +Arguments Scope Opp [_ qf_scope]. +Arguments Scope NatMul [_ qf_scope nat_scope]. +Arguments Scope Mul [_ qf_scope qf_scope]. +Arguments Scope Mul [_ qf_scope qf_scope]. +Arguments Scope Exp [_ qf_scope nat_scope]. +Arguments Scope Equal [_ qf_scope qf_scope]. +Arguments Scope And [_ qf_scope qf_scope]. +Arguments Scope Or [_ qf_scope qf_scope]. +Arguments Scope Implies [_ qf_scope qf_scope]. +Arguments Scope Not [_ qf_scope]. + +Implicit Arguments Bool [R]. +Prenex Implicits Const Add Opp NatMul Mul Exp Bool Unit And Or Implies Not Lt. +Prenex Implicits to_rterm. + +Notation True := (Bool true). +Notation False := (Bool false). + +Delimit Scope qf_scope with qfT. +Notation "''X_' i" := (Var _ i) : qf_scope. +Notation "n %:R" := (NatConst _ n) : qf_scope. +Notation "x %:T" := (Const x) : qf_scope. +Notation "0" := 0%:R%qfT : qf_scope. +Notation "1" := 1%:R%qfT : qf_scope. +Infix "+" := Add : qf_scope. +Notation "- t" := (Opp t) : qf_scope. +Notation "t - u" := (Add t (- u)) : qf_scope. +Infix "*" := Mul : qf_scope. +Infix "*+" := NatMul : qf_scope. +Infix "^+" := Exp : qf_scope. +Notation "t ^- n" := (t^-1 ^+ n)%qfT : qf_scope. +Infix "==" := Equal : qf_scope. +Infix "<%" := Lt : qf_scope. +Infix "<=%" := Le : qf_scope. +Infix "/\" := And : qf_scope. +Infix "\/" := Or : qf_scope. +Infix "==>" := Implies : qf_scope. +Notation "~ f" := (Not f) : qf_scope. +Notation "x != y" := (Not (x == y)) : qf_scope. + +Section evaluation. + +Variable R : realDomainType. + +Fixpoint eval (e : seq R) (t : term R) {struct t} : R := + match t with + | ('X_i)%qfT => e`_i + | (x%:T)%qfT => x + | (n%:R)%qfT => n%:R + | (t1 + t2)%qfT => eval e t1 + eval e t2 + | (- t1)%qfT => - eval e t1 + | (t1 *+ n)%qfT => eval e t1 *+ n + | (t1 * t2)%qfT => eval e t1 * eval e t2 + | (t1 ^+ n)%qfT => eval e t1 ^+ n + end. + +Lemma evalE (e : seq R) (t : term R) : eval e t = GRing.eval e t. +Proof. by elim: t=> /=; do ?[move->|move=>?]. Qed. + +Definition qf_eval e := fix loop (f : formula R) : bool := + match f with + | Bool b => b + | t1 == t2 => (eval e t1 == eval e t2)%bool + | t1 <% t2 => (eval e t1 < eval e t2)%bool + | t1 <=% t2 => (eval e t1 <= eval e t2)%bool + | f1 /\ f2 => loop f1 && loop f2 + | f1 \/ f2 => loop f1 || loop f2 + | f1 ==> f2 => (loop f1 ==> loop f2)%bool + | ~ f1 => ~~ loop f1 + end%qfT. + +Lemma qf_evalE (e : seq R) (f : formula R) : qf_eval e f = ord.qf_eval e f. +Proof. by elim: f=> /=; do ?[rewrite evalE|move->|move=>?]. Qed. + +Lemma to_rtermE (t : GRing.term R) : + GRing.rterm t -> to_rterm t = t :> GRing.term _. +Proof. +elim: t=> //=; do ? + [ by move=> u hu v hv /andP[ru rv]; rewrite hu ?hv + | by move=> u hu *; rewrite hu]. +Qed. + +End evaluation. + +Import Pdiv.Ring. + +Definition bind_def T1 T2 T3 (f : (T1 -> T2) -> T3) (k : T1 -> T2) := f k. +Notation "'bind' x <- y ; z" := + (bind_def y (fun x => z)) (at level 99, x at level 0, y at level 0, + format "'[hv' 'bind' x <- y ; '/' z ']'"). + +Section ProjDef. + +Variable F : realFieldType. + +Notation fF := (formula F). +Notation tF := (term F). +Definition polyF := seq tF. + +Lemma qf_formF (f : fF) : qf_form f. +Proof. by elim: f=> // *; apply/andP; split. Qed. + +Lemma rtermF (t : tF) : GRing.rterm t. +Proof. by elim: t=> //=; do ?[move->|move=>?]. Qed. + +Lemma rformulaF (f : fF) : rformula f. +Proof. by elim: f=> /=; do ?[rewrite rtermF|move->|move=>?]. Qed. + +Section If. + +Implicit Types (pf tf ef : formula F). + +Definition If pf tf ef := (pf /\ tf \/ ~ pf /\ ef)%qfT. + +End If. + +Notation "'If' c1 'Then' c2 'Else' c3" := (If c1 c2 c3) + (at level 200, right associativity, format +"'[hv ' 'If' c1 '/' '[' 'Then' c2 ']' '/' '[' 'Else' c3 ']' ']'"). + +Notation cps T := ((T -> fF) -> fF). + +Section Pick. + +Variables (I : finType) (pred_f then_f : I -> fF) (else_f : fF). + +Definition Pick := + \big[Or/False]_(p : {ffun pred I}) + ((\big[And/True]_i (if p i then pred_f i else ~ pred_f i)) + /\ (if pick p is Some i then then_f i else else_f))%qfT. + +Lemma eval_Pick e (qev := qf_eval e) : + let P i := qev (pred_f i) in + qev Pick = (if pick P is Some i then qev (then_f i) else qev else_f). +Proof. +move=> P; rewrite ((big_morph qev) false orb) //= big_orE /=. +apply/existsP/idP=> [[p] | true_at_P]. + rewrite ((big_morph qev) true andb) //= big_andE /=. + case/andP=> /forallP eq_p_P. + rewrite (@eq_pick _ _ P) => [|i]; first by case: pick. + by move/(_ i): eq_p_P => /=; case: (p i) => //=; move/negbTE. +exists [ffun i => P i] => /=; apply/andP; split. + rewrite ((big_morph qev) true andb) //= big_andE /=. + by apply/forallP=> i; rewrite /= ffunE; case Pi: (P i) => //=; apply: negbT. +rewrite (@eq_pick _ _ P) => [|i]; first by case: pick true_at_P. +by rewrite ffunE. +Qed. + +End Pick. + +Fixpoint eval_poly (e : seq F) pf := + if pf is c :: qf then (eval_poly e qf) * 'X + (eval e c)%:P else 0. + +Lemma eval_polyP e p : eval_poly e p = Poly (map (eval e) p). +Proof. by elim: p=> // a p /= ->; rewrite cons_poly_def. Qed. + +Fixpoint Size (p : polyF) : cps nat := fun k => + if p is c :: q then + bind n <- Size q; + if n is m.+1 then k m.+2 + else If c == 0 Then k 0%N Else k 1%N + else k 0%N. + +Definition Isnull (p : polyF) : cps bool := fun k => + bind n <- Size p; k (n == 0%N). + +Definition LtSize (p q : polyF) : cps bool := fun k => + bind n <- Size p; bind m <- Size q; k (n < m)%N. + +Fixpoint LeadCoef p : cps tF := fun k => + if p is c :: q then + bind l <- LeadCoef q; If l == 0 Then k c Else k l + else k (Const 0). + +Fixpoint AmulXn (a : tF) (n : nat) : polyF:= + if n is n'.+1 then (Const 0) :: (AmulXn a n') else [::a]. + +Fixpoint AddPoly (p q : polyF) := + if p is a::p' then + if q is b::q' then (a + b)%qfT :: (AddPoly p' q') + else p + else q. +Local Infix "++" := AddPoly : qf_scope. + +Definition ScalPoly (c : tF) (p : polyF) : polyF := map (Mul c) p. +Local Infix "*:" := ScalPoly : qf_scope. + +Fixpoint MulPoly (p q : polyF) := if p is a :: p' + then (a *: q ++ (0 :: (MulPoly p' q)))%qfT else [::]. +Local Infix "**" := MulPoly (at level 40) : qf_scope. + +Lemma map_poly0 (R R' : ringType) (f : R -> R') : map_poly f 0 = 0. +Proof. by rewrite map_polyE polyseq0. Qed. + +Definition ExpPoly p n := iterop n MulPoly p [::1%qfT]. +Local Infix "^^+" := ExpPoly (at level 29) : qf_scope. + +Definition OppPoly := ScalPoly (@Const F (-1)). +Local Notation "-- p" := (OppPoly p) (at level 35) : qf_scope. +Local Notation "p -- q" := (p ++ (-- q))%qfT (at level 50) : qf_scope. + +Definition NatMulPoly n := ScalPoly (NatConst F n). +Local Infix "+**" := NatMulPoly (at level 40) : qf_scope. + +Fixpoint Horner (p : polyF) (x : tF) : tF := + if p is a :: p then (Horner p x * x + a)%qfT else 0%qfT. + +Fixpoint Deriv (p : polyF) : polyF := + if p is a :: q then (q ++ (0 :: Deriv q))%qfT else [::]. + +Fixpoint Rediv_rec_loop (q : polyF) sq cq + (c : nat) (qq r : polyF) (n : nat) {struct n} : + cps (nat * polyF * polyF) := fun k => + bind sr <- Size r; + if (sr < sq)%N then k (c, qq, r) else + bind lr <- LeadCoef r; + let m := AmulXn lr (sr - sq) in + let qq1 := (qq ** [::cq] ++ m)%qfT in + let r1 := (r ** [::cq] -- m ** q)%qfT in + if n is n1.+1 then Rediv_rec_loop q sq cq c.+1 qq1 r1 n1 k + else k (c.+1, qq1, r1). + + Definition Rediv (p : polyF) (q : polyF) : cps (nat * polyF * polyF) := + fun k => + bind b <- Isnull q; + if b then k (0%N, [::Const 0], p) + else bind sq <- Size q; + bind sp <- Size p; + bind lq <- LeadCoef q; + Rediv_rec_loop q sq lq 0 [::Const 0] p sp k. + +Definition Rmod (p : polyF) (q : polyF) (k : polyF -> fF) : fF := + Rediv p q (fun d => k d.2)%PAIR. +Definition Rdiv (p : polyF) (q : polyF) (k : polyF -> fF) : fF := + Rediv p q (fun d => k d.1.2)%PAIR. +Definition Rscal (p : polyF) (q : polyF) (k : nat -> fF) : fF := + Rediv p q (fun d => k d.1.1)%PAIR. +Definition Rdvd (p : polyF) (q : polyF) (k : bool -> fF) : fF := + bind r <- Rmod p q; bind r_null <- Isnull r; k r_null. + +Fixpoint rgcdp_loop n (pp qq : {poly F}) {struct n} := + if rmodp pp qq == 0 then qq + else if n is n1.+1 then rgcdp_loop n1 qq (rmodp pp qq) + else rmodp pp qq. + +Fixpoint Rgcd_loop n pp qq k {struct n} := + bind r <- Rmod pp qq; bind b <- Isnull r; + if b then (k qq) + else if n is n1.+1 then Rgcd_loop n1 qq r k else k r. + +Definition Rgcd (p : polyF) (q : polyF) : cps polyF := fun k => + let aux p1 q1 k := (bind b <- Isnull p1; + if b then k q1 else bind n <- Size p1; Rgcd_loop n p1 q1 k) in + bind b <- LtSize p q; + if b then aux q p k else aux p q k. + +Fixpoint BigRgcd (ps : seq polyF) : cps (seq tF) := fun k => + if ps is p :: pr then bind r <- BigRgcd pr; Rgcd p r k else k [::Const 0]. + +Fixpoint Changes (s : seq tF) : cps nat := fun k => + if s is a :: q then + bind v <- Changes q; + If (Lt (a * head 0 q) 0)%qfT Then k (1 + v)%N Else k v + else k 0%N. + +Fixpoint SeqPInfty (ps : seq polyF) : cps (seq tF) := fun k => + if ps is p :: ps then + bind lp <- LeadCoef p; + bind lps <- SeqPInfty ps; + k (lp :: lps) + else k [::]. + +Fixpoint SeqMInfty (ps : seq polyF) : cps (seq tF) := fun k => + if ps is p :: ps then + bind lp <- LeadCoef p; + bind sp <- Size p; + bind lps <- SeqMInfty ps; + k ((-1)%:T ^+ (~~ odd sp) * lp :: lps)%qfT + else k [::]. + +Definition ChangesPoly ps : cps int := fun k => + bind mps <- SeqMInfty ps; + bind pps <- SeqPInfty ps; + bind vm <- Changes mps; bind vp <- Changes pps; k (vm%:Z - vp%:Z). + +Definition NextMod (p q : polyF) : cps polyF := fun k => + bind lq <- LeadCoef q; + bind spq <- Rscal p q; + bind rpq <- Rmod p q; k (- lq ^+ spq *: rpq)%qfT. + +Fixpoint ModsAux (p q : polyF) n : cps (seq polyF) := fun k => + if n is m.+1 + then + bind p_eq0 <- Isnull p; + if p_eq0 then k [::] + else + bind npq <- NextMod p q; + bind ps <- ModsAux q npq m; + k (p :: ps) + else k [::]. + +Definition Mods (p q : polyF) : cps (seq polyF) := fun k => + bind sp <- Size p; bind sq <- Size q; + ModsAux p q (maxn sp sq.+1) k. + +Definition PolyComb (sq : seq polyF) (sc : seq int) := + reducebig [::1%qfT] (iota 0 (size sq)) + (fun i => BigBody i MulPoly true (nth [::] sq i ^^+ comb_exp sc`_i)%qfT). + +Definition Pcq sq i := (nth [::] (map (PolyComb sq) (sg_tab (size sq))) i). + +Definition TaqR (p : polyF) (q : polyF) : cps int := fun k => + bind r <- Mods p (Deriv p ** q)%qfT; ChangesPoly r k. + +Definition TaqsR (p : polyF) (sq : seq polyF) (i : nat) : cps tF := + fun k => bind n <- TaqR p (Pcq sq i); k ((n%:~R) %:T)%qfT. + +Fixpoint ProdPoly T (s : seq T) (f : T -> cps polyF) : cps polyF := fun k => + if s is a :: s then + bind fa <- f a; + bind fs <- ProdPoly s f; + k (fa ** fs)%qfT + else k [::1%qfT]. + +Definition BoundingPoly (sq : seq polyF) : polyF := + Deriv (reducebig [::1%qfT] sq (fun i => BigBody i MulPoly true i)). + +Definition Coefs (n i : nat) : tF := + Const (match n with + | 0 => (i == 0%N)%:R + | 1 => [:: 2%:R^-1; 2%:R^-1; 0]`_i + | n => coefs _ n i + end). + +Definition CcountWeak (p : polyF) (sq : seq polyF) : cps tF := fun k => + let fix aux s (i : nat) k := if i is i'.+1 + then bind x <- TaqsR p sq i'; + aux (x * (Coefs (size sq) i') + s)%qfT i' k + else k s in + aux 0%qfT (3 ^ size sq)%N k. + +Definition CcountGt0 (sp sq : seq polyF) : fF := + bind p <- BigRgcd sp; bind p0 <- Isnull p; + if ~~ p0 then + bind c <- CcountWeak p sq; + Lt 0%qfT c + else + let bq := BoundingPoly sq in + bind cw <- CcountWeak bq sq; + ((reducebig True sq (fun q => + BigBody q And true (LeadCoef q (fun lq => Lt 0 lq)))) + \/ ((reducebig True sq (fun q => + BigBody q And true + (bind sq <- Size q; + bind lq <- LeadCoef q; + Lt 0 ((Opp 1) ^+ (sq).-1 * lq) + ))) \/ Lt 0 cw))%qfT. + + +Fixpoint abstrX (i : nat) (t : tF) : polyF := + (match t with + | 'X_n => if n == i then [::0; 1] else [::t] + | - x => -- abstrX i x + | x + y => abstrX i x ++ abstrX i y + | x * y => abstrX i x ** abstrX i y + | x *+ n => n +** abstrX i x + | x ^+ n => abstrX i x ^^+ n + | _ => [::t] + end)%qfT. + +Definition wproj (n : nat) (s : seq (GRing.term F) * seq (GRing.term F)) : + formula F := + let sp := map (abstrX n \o to_rterm) s.1%PAIR in + let sq := map (abstrX n \o to_rterm) s.2%PAIR in + CcountGt0 sp sq. + +Definition rcf_sat := proj_sat wproj. + +End ProjDef. + +Section ProjCorrect. + +Variable F : rcfType. +Implicit Types (e : seq F). + +Notation fF := (formula F). +Notation tF := (term F). +Notation polyF := (polyF F). + +Notation "'If' c1 'Then' c2 'Else' c3" := (If c1 c2 c3) + (at level 200, right associativity, format +"'[hv ' 'If' c1 '/' '[' 'Then' c2 ']' '/' '[' 'Else' c3 ']' ']'"). + +Notation cps T := ((T -> fF) -> fF). + +Local Infix "**" := MulPoly (at level 40) : qf_scope. +Local Infix "+**" := NatMulPoly (at level 40) : qf_scope. +Local Notation "-- p" := (OppPoly p) (at level 35) : qf_scope. +Local Notation "p -- q" := (p ++ (-- q))%qfT (at level 50) : qf_scope. +Local Infix "^^+" := ExpPoly (at level 29) : qf_scope. +Local Infix "**" := MulPoly (at level 40) : qf_scope. +Local Infix "*:" := ScalPoly : qf_scope. +Local Infix "++" := AddPoly : qf_scope. + +Lemma eval_If e pf tf ef (ev := qf_eval e) : + ev (If pf Then tf Else ef) = (if ev pf then ev tf else ev ef). +Proof. by unlock (If _ Then _ Else _)=> /=; case: ifP => _; rewrite ?orbF. Qed. + +Lemma eval_Size k p e : + qf_eval e (Size p k) = qf_eval e (k (size (eval_poly e p))). +Proof. +elim: p e k=> [|c p ihp] e k; first by rewrite size_poly0. +rewrite ihp /= size_MXaddC -size_poly_eq0; case: size=> //. +by rewrite eval_If /=; case: (_ == _). +Qed. + +Lemma eval_Isnull k p e : qf_eval e (Isnull p k) + = qf_eval e (k (eval_poly e p == 0)). +Proof. by rewrite eval_Size size_poly_eq0. Qed. + +Lemma eval_LeadCoef e p k k' : + (forall x, qf_eval e (k x) = (k' (eval e x))) -> + qf_eval e (LeadCoef p k) = k' (lead_coef (eval_poly e p)). +Proof. +move=> Pk; elim: p k k' Pk=> [|a p ihp] k k' Pk //=. + by rewrite lead_coef0 Pk. +rewrite (ihp _ (fun l => if l == 0 then qf_eval e (k a) else (k' l))); last first. + by move=> x; rewrite eval_If /= !Pk. +rewrite lead_coef_eq0; have [->|p_neq0] := altP (_ =P 0). + by rewrite mul0r add0r lead_coefC. +rewrite lead_coefDl ?lead_coefMX ?size_mulX // ltnS size_polyC. +by rewrite (leq_trans (leq_b1 _)) // size_poly_gt0. +Qed. + +Implicit Arguments eval_LeadCoef [e p k]. +Prenex Implicits eval_LeadCoef. + +Lemma eval_AmulXn a n e : eval_poly e (AmulXn a n) = (eval e a)%:P * 'X^n. +Proof. +elim: n=> [|n] /=; first by rewrite expr0 mulr1 mul0r add0r. +by move->; rewrite addr0 -mulrA -exprSr. +Qed. + +Lemma eval_AddPoly p q e : + eval_poly e (p ++ q)%qfT = (eval_poly e p) + (eval_poly e q). +Proof. +elim: p q => [|a p Hp] q /=; first by rewrite add0r. +case: q => [|b q] /=; first by rewrite addr0. +by rewrite Hp mulrDl rmorphD /= !addrA [X in _ = X + _]addrAC. +Qed. + +Lemma eval_ScalPoly e t p : + eval_poly e (ScalPoly t p) = (eval e t) *: (eval_poly e p). +Proof. +elim: p=> [|a p ihp] /=; first by rewrite scaler0. +by rewrite ihp scalerDr scalerAl -!mul_polyC rmorphM. +Qed. + +Lemma eval_MulPoly e p q : + eval_poly e (p ** q)%qfT = (eval_poly e p) * (eval_poly e q). +Proof. +elim: p q=> [|a p Hp] q /=; first by rewrite mul0r. +rewrite eval_AddPoly /= eval_ScalPoly Hp. +by rewrite addr0 mulrDl addrC mulrAC mul_polyC. +Qed. + +Lemma eval_ExpPoly e p n : eval_poly e (p ^^+ n)%qfT = (eval_poly e p) ^+ n. +Proof. +case: n=> [|n]; first by rewrite /= expr0 mul0r add0r. +rewrite /ExpPoly iteropS exprSr; elim: n=> [|n ihn] //=. + by rewrite expr0 mul1r. +by rewrite eval_MulPoly ihn exprS mulrA. +Qed. + +Lemma eval_NatMulPoly p n e : + eval_poly e (n +** p)%qfT = (eval_poly e p) *+ n. +Proof. +elim: p; rewrite //= ?mul0rn // => c p ->. +rewrite mulrnDl mulr_natl polyC_muln; congr (_+_). +by rewrite -mulr_natl mulrAC -mulrA mulr_natl mulrC. +Qed. + +Lemma eval_OppPoly p e : eval_poly e (-- p)%qfT = - eval_poly e p. +Proof. +elim: p; rewrite //= ?oppr0 // => t ts ->. +by rewrite !mulNr !opprD polyC_opp mul1r. +Qed. + +Lemma eval_Horner e p x : eval e (Horner p x) = (eval_poly e p).[eval e x]. +Proof. by elim: p => /= [|a p ihp]; rewrite !(horner0, hornerE) // ihp. Qed. + +Lemma eval_ConstPoly e c : eval_poly e [::c] = (eval e c)%:P. +Proof. by rewrite /= mul0r add0r. Qed. + +Lemma eval_Deriv e p : eval_poly e (Deriv p) = (eval_poly e p)^`(). +Proof. +elim: p=> [|a p ihp] /=; first by rewrite deriv0. +by rewrite eval_AddPoly /= addr0 ihp !derivE. +Qed. + +Definition eval_OpPoly := + (eval_MulPoly, eval_AmulXn, eval_AddPoly, eval_OppPoly, eval_NatMulPoly, + eval_ConstPoly, eval_Horner, eval_ExpPoly, eval_Deriv, eval_ScalPoly). + +Lemma eval_Changes e s k : qf_eval e (Changes s k) + = qf_eval e (k (changes (map (eval e) s))). +Proof. +elim: s k=> //= a q ihq k; rewrite ihq eval_If /= -nth0. +by case: q {ihq}=> /= [|b q]; [rewrite /= mulr0 ltrr add0n | case: ltrP]. +Qed. + +Lemma eval_SeqPInfty e ps k k' : + (forall xs, qf_eval e (k xs) = k' (map (eval e) xs)) -> + qf_eval e (SeqPInfty ps k) + = k' (map lead_coef (map (eval_poly e) ps)). +Proof. +elim: ps k k' => [|p ps ihps] k k' Pk /=; first by rewrite Pk. +rewrite (eval_LeadCoef (fun lp => + k' (lp :: [seq lead_coef i |i <- [seq eval_poly e i | i <- ps]]))) => // lp. +rewrite (ihps _ (fun ps => k' (eval e lp :: ps))) => //= lps. +by rewrite Pk. +Qed. + +Implicit Arguments eval_SeqPInfty [e ps k]. +Prenex Implicits eval_SeqPInfty. + +Lemma eval_SeqMInfty e ps k k' : + (forall xs, qf_eval e (k xs) = k' (map (eval e) xs)) -> + qf_eval e (SeqMInfty ps k) + = k' (map (fun p : {poly F} => (-1) ^+ (~~ odd (size p)) * lead_coef p) + (map (eval_poly e) ps)). +Proof. +elim: ps k k' => [|p ps ihps] k k' Pk /=; first by rewrite Pk. +rewrite (eval_LeadCoef (fun lp => + k' ((-1) ^+ (~~ odd (size (eval_poly e p))) * lp + :: [seq (-1) ^+ (~~ odd (size p)) * lead_coef p + | p : {poly _} <- [seq eval_poly e i | i <- ps]]))) => // lp. +rewrite eval_Size /= (ihps _ (fun ps => + k' (((-1) ^+ (~~ odd (size (eval_poly e p))) * eval e lp) :: ps))) => //= lps. +by rewrite Pk. +Qed. + +Implicit Arguments eval_SeqMInfty [e ps k]. +Prenex Implicits eval_SeqMInfty. + +Lemma eval_ChangesPoly e ps k : qf_eval e (ChangesPoly ps k) = + qf_eval e (k (changes_poly (map (eval_poly e) ps))). +Proof. +rewrite (eval_SeqMInfty (fun mps => + qf_eval e (k ((changes mps)%:Z - + (changes_pinfty [seq eval_poly e i | i <- ps])%:Z)))) => // mps. +rewrite (eval_SeqPInfty (fun pps => + qf_eval e (k ((changes (map (eval e) mps))%:Z - (changes pps)%:Z)))) => // pps. +by rewrite !eval_Changes. +Qed. + +Fixpoint redivp_rec_loop (q : {poly F}) sq cq + (k : nat) (qq r : {poly F})(n : nat) {struct n} := + if (size r < sq)%N then (k, qq, r) else + let m := (lead_coef r) *: 'X^(size r - sq) in + let qq1 := qq * cq%:P + m in + let r1 := r * cq%:P - m * q in + if n is n1.+1 then redivp_rec_loop q sq cq k.+1 qq1 r1 n1 else (k.+1, qq1, r1). + +Lemma redivp_rec_loopP q c qq r n : redivp_rec q c qq r n + = redivp_rec_loop q (size q) (lead_coef q) c qq r n. +Proof. by elim: n c qq r => [| n Pn] c qq r //=; rewrite Pn. Qed. + +Lemma eval_Rediv_rec_loop e q sq cq c qq r n k k' + (d := redivp_rec_loop (eval_poly e q) sq (eval e cq) + c (eval_poly e qq) (eval_poly e r) n) : + (forall c qq r, qf_eval e (k (c, qq, r)) + = k' (c, eval_poly e qq, eval_poly e r)) -> + qf_eval e (Rediv_rec_loop q sq cq c qq r n k) = k' d. +Proof. +move=> Pk; elim: n c qq r k Pk @d=> [|n ihn] c qq r k Pk /=. + rewrite eval_Size /=; have [//=|gtq] := ltnP. + rewrite (eval_LeadCoef (fun lr => + let m := lr *: 'X^(size (eval_poly e r) - sq) in + let qq1 := (eval_poly e qq) * (eval e cq)%:P + m in + let r1 := (eval_poly e r) * (eval e cq)%:P - m * (eval_poly e q) in + k' (c.+1, qq1, r1))) //. + by move=> x /=; rewrite Pk /= !eval_OpPoly /= !mul_polyC. +rewrite eval_Size /=; have [//=|gtq] := ltnP. +rewrite (eval_LeadCoef (fun lr => + let m := lr *: 'X^(size (eval_poly e r) - sq) in + let qq1 := (eval_poly e qq) * (eval e cq)%:P + m in + let r1 := (eval_poly e r) * (eval e cq)%:P - m * (eval_poly e q) in + k' (redivp_rec_loop (eval_poly e q) sq (eval e cq) c.+1 qq1 r1 n))) //=. +by move=> x; rewrite ihn // !eval_OpPoly /= !mul_polyC. +Qed. + +Implicit Arguments eval_Rediv_rec_loop [e q sq cq c qq r n k]. +Prenex Implicits eval_Rediv_rec_loop. + +Lemma eval_Rediv e p q k k' (d := (redivp (eval_poly e p) (eval_poly e q))) : + (forall c qq r, qf_eval e (k (c, qq, r)) = k' (c, eval_poly e qq, eval_poly e r)) -> + qf_eval e (Rediv p q k) = k' d. +Proof. +move=> Pk; rewrite eval_Isnull /d unlock. +have [_|p_neq0] /= := boolP (_ == _); first by rewrite Pk /= mul0r add0r. +rewrite !eval_Size; set p' := eval_poly e p; set q' := eval_poly e q. +rewrite (eval_LeadCoef (fun lq => + k' (redivp_rec_loop q' (size q') lq 0 0 p' (size p')))) /=; last first. + by move=> x; rewrite (eval_Rediv_rec_loop k') //= mul0r add0r. +by rewrite redivp_rec_loopP. +Qed. + +Implicit Arguments eval_Rediv [e p q k]. +Prenex Implicits eval_Rediv. + +Lemma eval_NextMod e p q k k' : + (forall p, qf_eval e (k p) = k' (eval_poly e p)) -> + qf_eval e (NextMod p q k) = + k' (next_mod (eval_poly e p) (eval_poly e q)). +Proof. +move=> Pk; set p' := eval_poly e p; set q' := eval_poly e q. +rewrite (eval_LeadCoef (fun lq => + k' (- lq ^+ rscalp p' q' *: rmodp p' q'))) => // lq. +rewrite (eval_Rediv (fun spq => + k' (- eval e lq ^+ spq.1.1%PAIR *: rmodp p' q'))) => //= spq _ _. +rewrite (eval_Rediv (fun mpq => + k' (- eval e lq ^+ spq *: mpq.2%PAIR))) => //= _ _ mpq. +by rewrite Pk !eval_OpPoly. +Qed. + +Implicit Arguments eval_NextMod [e p q k]. +Prenex Implicits eval_NextMod. + +Lemma eval_Rgcd_loop e n p q k k' : + (forall p, qf_eval e (k p) = k' (eval_poly e p)) + -> qf_eval e (Rgcd_loop n p q k) = + k' (rgcdp_loop n (eval_poly e p) (eval_poly e q)). +Proof. +elim: n p q k k'=> [|n ihn] p q k k' Pk /=. + rewrite (eval_Rediv (fun r => + if r.2%PAIR == 0 then k' (eval_poly e q) else k' r.2%PAIR)) /=. + by case: eqP. + by move=> _ _ r; rewrite eval_Isnull; case: eqP. +pose q' := eval_poly e q. +rewrite (eval_Rediv (fun r => + if r.2%PAIR == 0 then k' q' else k' (rgcdp_loop n q' r.2%PAIR))) /=. + by case: eqP. +move=> _ _ r; rewrite eval_Isnull; case: eqP; first by rewrite Pk. +by rewrite (ihn _ _ _ k'). +Qed. + +Lemma eval_Rgcd e p q k k' : + (forall p, qf_eval e (k p) = k' (eval_poly e p)) -> + qf_eval e (Rgcd p q k) = + k' (rgcdp (eval_poly e p) (eval_poly e q)). +Proof. +move=> Pk; rewrite /Rgcd /LtSize !eval_Size /rgcdp. +case: ltnP=> _; rewrite !eval_Isnull; case: eqP=> // _; +by rewrite eval_Size; apply: eval_Rgcd_loop. +Qed. + + +Lemma eval_BigRgcd e ps k k' : + (forall p, qf_eval e (k p) = k' (eval_poly e p)) -> + qf_eval e (BigRgcd ps k) = + k' (\big[@rgcdp _/0%:P]_(i <- ps) (eval_poly e i)). +Proof. +elim: ps k k'=> [|p sp ihsp] k k' Pk /=. + by rewrite big_nil Pk /= mul0r add0r. +rewrite big_cons (ihsp _ (fun r => k' (rgcdp (eval_poly e p) r))) //. +by move=> r; apply: eval_Rgcd. +Qed. + +Implicit Arguments eval_Rgcd [e p q k]. +Prenex Implicits eval_Rgcd. + + +Fixpoint mods_aux (p q : {poly F}) (n : nat) : seq {poly F} := + if n is m.+1 + then if p == 0 then [::] + else p :: (mods_aux q (next_mod p q) m) + else [::]. + +Lemma eval_ModsAux e p q n k k' : + (forall sp, qf_eval e (k sp) = k' (map (eval_poly e) sp)) -> + qf_eval e (ModsAux p q n k) = + k' (mods_aux (eval_poly e p) (eval_poly e q) n). +Proof. +elim: n p q k k'=> [|n ihn] p q k k' Pk; first by rewrite /= Pk. +rewrite /= eval_Isnull; have [|ep_neq0] := altP (_ =P _); first by rewrite Pk. +set q' := eval_poly e q; set p' := eval_poly e p. +rewrite (eval_NextMod (fun npq => k' (p' :: mods_aux q' npq n))) => // npq. +by rewrite (ihn _ _ _ (fun ps => k' (p' :: ps))) => // ps; rewrite Pk. +Qed. + +Implicit Arguments eval_ModsAux [e p q n k]. +Prenex Implicits eval_ModsAux. + +Lemma eval_Mods e p q k k' : + (forall sp, qf_eval e (k sp) = k' (map (eval_poly e) sp)) -> + qf_eval e (Mods p q k) = k' (mods (eval_poly e p) (eval_poly e q)). +Proof. by move=> Pk; rewrite !eval_Size; apply: eval_ModsAux. Qed. + +Implicit Arguments eval_Mods [e p q k]. +Prenex Implicits eval_Mods. + +Lemma eval_TaqR e p q k : + qf_eval e (TaqR p q k) = + qf_eval e (k (taqR (eval_poly e p) (eval_poly e q))). +Proof. +rewrite (eval_Mods (fun r => qf_eval e (k (changes_poly r)))). + by rewrite !eval_OpPoly. +by move=> sp; rewrite !eval_ChangesPoly. +Qed. + +Lemma eval_PolyComb e sq sc : + eval_poly e (PolyComb sq sc) = poly_comb (map (eval_poly e) sq) sc. +Proof. +rewrite /PolyComb /poly_comb size_map. +rewrite -BigOp.bigopE -val_enum_ord -filter_index_enum !big_map. +apply: (big_ind2 (fun u v => eval_poly e u = v)). ++ by rewrite /= mul0r add0r. ++ by move=> x x' y y'; rewrite eval_MulPoly=> -> ->. +by move=> i _; rewrite eval_ExpPoly /= (nth_map [::]). +Qed. + +Definition pcq (sq : seq {poly F}) i := + (map (poly_comb sq) (sg_tab (size sq)))`_i. + +Lemma eval_Pcq e sq i : + eval_poly e (Pcq sq i) = pcq (map (eval_poly e) sq) i. +Proof. +rewrite /Pcq /pcq size_map; move: (sg_tab _)=> s. +have [ge_is|lt_is] := leqP (size s) i. + by rewrite !nth_default ?size_map // /=. +rewrite -(nth_map _ 0) ?size_map //; congr _`_i; rewrite -map_comp. +by apply: eq_map=> x /=; rewrite eval_PolyComb. +Qed. + +Lemma eval_TaqsR e p sq i k k' : + (forall x, qf_eval e (k x) = k' (eval e x)) -> + qf_eval e (TaqsR p sq i k) = + k' (taqsR (eval_poly e p) (map (eval_poly e) sq) i). +Proof. by move=> Pk; rewrite /TaqsR /taqsR eval_TaqR Pk /= eval_Pcq. Qed. + +Implicit Arguments eval_TaqsR [e p sq i k]. +Prenex Implicits eval_TaqsR. + +Fact invmx_ctmat1 : invmx (map_mx (intr : int -> F) ctmat1) = + \matrix_(i, j) + (nth [::] [:: [:: 2%:R^-1; - 2%:R^-1; 0]; + [:: 2%:R^-1; 2%:R^-1; -1]; + [:: 0; 0; 1]] i)`_j :> 'M[F]_3. +Proof. +rewrite -[lhs in lhs = _]mul1r; apply: (canLR (mulrK _)). + exact: ctmat1_unit. +symmetry; rewrite /ctmat1. +apply/matrixP => i j; rewrite !(big_ord_recl, big_ord0, mxE) /=. +have halfP (K : numFieldType) : 2%:R^-1 + 2%:R^-1 = 1 :> K. + by rewrite -mulr2n -[_ *+ 2]mulr_natl mulfV // pnatr_eq0. +move: i; do ?[case => //=]; move: j; do ?[case => //=] => _ _; +rewrite !(mulr1, mul1r, mulrN1, mulN1r, mulr0, mul0r, opprK); +by rewrite !(addr0, add0r, oppr0, subrr, addrA, halfP). +Qed. + +Lemma eval_Coefs e n i : eval e (Coefs F n i) = coefs F n i. +Proof. +case: n => [|[|n]] //=; rewrite /coefs /=. + case: i => [|i]; last first. + by rewrite nth_default // size_map size_enum_ord expn0. + rewrite (nth_map 0) ?size_enum_ord //. + set O := _`_0; rewrite (_ : O = ord0). + by rewrite ?castmxE ?cast_ord_id map_mx1 invmx1 mxE. + by apply: val_inj => /=; rewrite nth_enum_ord. +have [lt_i3|le_3i] := ltnP i 3; last first. + by rewrite !nth_default // size_map size_enum_ord. +rewrite /ctmat /= ?ntensmx1 invmx_ctmat1 /=. +rewrite (nth_map 0) ?size_enum_ord // castmxE /=. +rewrite !mxE !cast_ord_id //= nth_enum_ord //=. +by move: i lt_i3; do 3?case. +Qed. + +Lemma eval_CcountWeak e p sq k k' : + (forall x, qf_eval e (k x) = k' (eval e x)) -> + qf_eval e (CcountWeak p sq k) = + k' (ccount_weak (eval_poly e p) (map (eval_poly e) sq)). +Proof. +move=> Pk; rewrite /CcountWeak /ccount_weak. +set Aux := (fix Aux s i k := match i with 0 => _ | _ => _ end). +set aux := (fix aux s i := match i with 0 => _ | _ => _ end). +rewrite size_map -[0]/(eval e 0%qfT); move: 0%qfT=> x. +elim: (_ ^ _)%N k k' Pk x=> /= [|n ihn] k k' Pk x. + by rewrite Pk. +rewrite (eval_TaqsR + (fun y => k' (aux (y * (coefs F (size sq) n) + eval e x) n))). + by rewrite size_map. +by move=> y; rewrite (ihn _ k') // -(eval_Coefs e). +Qed. + +Implicit Arguments eval_CcountWeak [e p sq k]. +Prenex Implicits eval_CcountWeak. + +Lemma eval_ProdPoly e T s f k f' k' : + (forall x k k', (forall p, (qf_eval e (k p) = k' (eval_poly e p))) -> + qf_eval e (f x k) = k' (f' x)) -> + (forall p, qf_eval e (k p) = k' (eval_poly e p)) -> + qf_eval e (@ProdPoly _ T s f k) = k' (\prod_(x <- s) f' x). +Proof. +move=> Pf; elim: s k k'=> [|a s ihs] k k' Pk /=. + by rewrite big_nil Pk /= !(mul0r, add0r). +rewrite (Pf _ _ (fun fa => k' (fa * \prod_(x <- s) f' x))). + by rewrite big_cons. +move=> fa; rewrite (ihs _ (fun fs => k' (eval_poly e fa * fs))) //. +by move=> fs; rewrite Pk eval_OpPoly. +Qed. + +Implicit Arguments eval_ProdPoly [e T s f k]. +Prenex Implicits eval_ProdPoly. + +Lemma eval_BoundingPoly e sq : + eval_poly e (BoundingPoly sq) = bounding_poly (map (eval_poly e) sq). +Proof. +rewrite eval_Deriv -BigOp.bigopE; congr _^`(); rewrite big_map. +by apply: big_morph => [p q | ]/=; rewrite ?eval_MulPoly // mul0r add0r. +Qed. + +Lemma eval_CcountGt0 e sp sq : qf_eval e (CcountGt0 sp sq) = + ccount_gt0 (map (eval_poly e) sp) (map (eval_poly e) sq). +Proof. +pose sq' := map (eval_poly e) sq; rewrite /ccount_gt0. +rewrite (@eval_BigRgcd _ _ _ (fun p => if p != 0 + then 0 < ccount_weak p sq' + else let bq := bounding_poly sq' in + [|| \big[andb/true]_(q <- sq') (0 < lead_coef q), + \big[andb/true]_(q <- sq') (0 < (-1) ^+ (size q).-1 * lead_coef q) + | 0 < ccount_weak bq sq'])). + by rewrite !big_map. +move=> p; rewrite eval_Isnull; case: eqP=> _ /=; last first. + by rewrite (eval_CcountWeak (> 0)). +rewrite (eval_CcountWeak (fun n => + [|| \big[andb/true]_(q <- sq') (0 < lead_coef q), + \big[andb/true]_(q <- sq') (0 < (-1) ^+ (size q).-1 * lead_coef q) + | 0 < n ])). + by rewrite eval_BoundingPoly. +move=> n /=; rewrite -!BigOp.bigopE !big_map; congr [|| _, _| _]. + apply: (big_ind2 (fun u v => qf_eval e u = v))=> //=. + by move=> u v u' v' -> ->. + by move=> i _; rewrite (eval_LeadCoef (> 0)). +apply: (big_ind2 (fun u v => qf_eval e u = v))=> //=. + by move=> u v u' v' -> ->. +by move=> i _; rewrite eval_Size (eval_LeadCoef (fun lq => + (0 < (-1) ^+ (size (eval_poly e i)).-1 * lq))). +Qed. + +Lemma abstrXP e i t x : + (eval_poly e (abstrX i t)).[x] = eval (set_nth 0 e i x) t. +Proof. +elim: t. +- move=> n /=; case ni: (_ == _); + rewrite //= ?(mul0r,add0r,addr0,polyC1,mul1r,hornerX,hornerC); + by rewrite // nth_set_nth /= ni. +- by move=> r; rewrite /= mul0r add0r hornerC. +- by move=> r; rewrite /= mul0r add0r hornerC. +- by move=> t tP s sP; rewrite /= eval_AddPoly hornerD tP ?sP. +- by move=> t tP; rewrite /= eval_OppPoly hornerN tP. +- by move=> t tP n; rewrite /= eval_NatMulPoly hornerMn tP. +- by move=> t tP s sP; rewrite /= eval_MulPoly hornerM tP ?sP. +- by move=> t tP n; rewrite /= eval_ExpPoly horner_exp tP. +Qed. + +Lemma wf_QE_wproj i bc (bc_i := @wproj F i bc) : + dnf_rterm (w_to_oclause bc) -> qf_form bc_i && rformula bc_i. +Proof. +case: bc @bc_i=> sp sq /=; rewrite /dnf_rterm /= /wproj andbT=> /andP[rsp rsq]. +by rewrite qf_formF rformulaF. +Qed. + +Lemma valid_QE_wproj i bc (bc' := w_to_oclause bc) + (ex_i_bc := ('exists 'X_i, odnf_to_oform [:: bc'])%oT) e : + dnf_rterm bc' -> reflect (holds e ex_i_bc) (ord.qf_eval e (wproj i bc)). +Proof. +case: bc @bc' @ex_i_bc=> sp sq /=; rewrite /dnf_rterm /wproj /= andbT. +move=> /andP[rsp rsq]; rewrite -qf_evalE. +rewrite eval_CcountGt0 /=; apply: (equivP (ccount_gt0P _ _)). +set P1 := (fun x => _); set P2 := (fun x => _). +suff: forall x, P1 x <-> P2 x. + by move=> hP; split=> [] [x Px]; exists x; rewrite (hP, =^~ hP). +move=> x; rewrite /P1 /P2 {P1 P2} !big_map !(big_seq_cond xpredT) /=. +rewrite (eq_bigr (fun t => GRing.eval (set_nth 0 e i x) t == 0)); last first. + by move=> t /andP[t_in_sp _]; rewrite abstrXP evalE to_rtermE ?(allP rsp). +rewrite [X in _ && X](eq_bigr (fun t => 0 < GRing.eval (set_nth 0 e i x) t)); + last by move=> t /andP[tsq _]; rewrite abstrXP evalE to_rtermE ?(allP rsq). +rewrite -!big_seq_cond !(rwP (qf_evalP _ _)); first last. ++ elim: sp rsp => //= p sp ihsp /andP[rp rsp]; first by rewrite ihsp. ++ elim: sq rsq => //= q sq ihsq /andP[rq rsq]; first by rewrite ihsq. +rewrite !(rwP andP) (rwP orP) orbF !andbT /=. +have unfoldr P s : foldr (fun t => ord.And (P t)) ord.True s = + \big[ord.And/ord.True]_(t <- s) P t by rewrite unlock /reducebig. +rewrite !unfoldr; set e' := set_nth _ _ _ _. +by rewrite !(@big_morph _ _ (ord.qf_eval _) true andb). +Qed. + +Lemma rcf_satP e f : reflect (holds e f) (rcf_sat e f). +Proof. exact: (proj_satP wf_QE_wproj valid_QE_wproj). Qed. + +End ProjCorrect. + +(* Section Example. *) +(* no chances it computes *) + +(* Require Import rat. *) + +(* Eval vm_compute in (54%:R / 289%:R + 2%:R^-1 :rat). *) + +(* Local Open Scope qf_scope. *) + +(* Notation polyF := (polyF [realFieldType of rat]). *) +(* Definition p : polyF := [::'X_2; 'X_1; 'X_0]. *) +(* Definition q : polyF := [:: 0; 1]. *) +(* Definition sq := [::q]. *) + +(* Eval vm_compute in MulPoly p q. *) + +(* Eval vm_compute in Rediv ([:: 1] : polyF) [::1]. *) + +(* Definition fpq := Eval vm_compute in (CcountWeak p [::q]). *) + +(* End Example. *) diff --git a/mathcomp/real_closed/qe_rcf_th.v b/mathcomp/real_closed/qe_rcf_th.v new file mode 100644 index 0000000..f1e5a61 --- /dev/null +++ b/mathcomp/real_closed/qe_rcf_th.v @@ -0,0 +1,1293 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice path fintype. +Require Import div bigop ssralg poly polydiv ssrnum perm zmodp ssrint. +Require Import polyorder polyrcf interval matrix mxtens. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing.Theory Num.Theory Num.Def Pdiv.Ring Pdiv.ComRing. + +Local Open Scope nat_scope. +Local Open Scope ring_scope. + +Section extra. + +Variable R : rcfType. +Implicit Types (p q : {poly R}). + + +(* Proof. *) +(* move=> sq; rewrite comp_polyE; case hp: (size p) => [|n]. *) +(* move/eqP: hp; rewrite size_poly_eq0 => /eqP ->. *) +(* by rewrite !big_ord0 mulr1 lead_coef0. *) +(* rewrite big_ord_recr /= addrC lead_coefDl. *) +(* by rewrite lead_coefZ lead_coef_exp // !lead_coefE hp. *) +(* rewrite (leq_ltn_trans (size_sum _ _ _)) // size_scale; last first. *) +(* rewrite -[n]/(n.+1.-1) -hp -lead_coefE ?lead_coef_eq0 //. *) +(* by rewrite -size_poly_eq0 hp. *) +(* rewrite polySpred ?ltnS ?expf_eq0; last first. *) +(* by rewrite andbC -size_poly_eq0 gtn_eqF // ltnW. *) +(* apply/bigmax_leqP => i _; rewrite size_exp. *) +(* have [->|/size_scale->] := eqVneq p`_i 0; first by rewrite scale0r size_poly0. *) +(* by rewrite (leq_trans (size_exp_leq _ _)) // ltn_mul2l -subn1 subn_gt0 sq /=. *) +(* Qed. *) + + +Lemma mul2n n : (2 * n = n + n)%N. Proof. by rewrite mulSn mul1n. Qed. +Lemma mul3n n : (3 * n = n + (n + n))%N. Proof. by rewrite !mulSn addn0. Qed. +Lemma exp3n n : (3 ^ n)%N = (3 ^ n).-1.+1. +Proof. by elim: n => // n IHn; rewrite expnS IHn. Qed. + +Definition exp3S n : (3 ^ n.+1 = 3 ^ n + (3 ^ n + 3 ^ n))%N + := etrans (expnS 3 n) (mul3n (3 ^ n)). + +Lemma tens_I3_mx (cR : comRingType) m n (M : 'M[cR]_(m,n)) : + 1%:M *t M = castmx (esym (mul3n _ ), esym (mul3n _ )) + (block_mx M 0 + 0 (block_mx M 0 + 0 M : 'M_(m+m,n+n)%N)). +Proof. +rewrite [1%:M : 'M_(1+2)%N]scalar_mx_block. +rewrite [1%:M : 'M_(1+1)%N]scalar_mx_block. +rewrite !tens_block_mx. +apply/eqP; rewrite -(can2_eq (castmxKV _ _) (castmxK _ _)); apply/eqP. +rewrite castmx_comp !tens_scalar_mx !tens0mx !scale1r. +rewrite (castmx_block (mul1n _) (mul1n _) (mul2n _) (mul2n _)). +rewrite !castmx_comp /= !castmx_id. +rewrite (castmx_block (mul1n _) (mul1n _) (mul1n _) (mul1n _)). +by rewrite !castmx_comp /= !castmx_id !castmx_const /=. +Qed. + +Lemma mul_1tensmx (cR : comRingType) (m n p: nat) + (e3n : (n + (n + n) = 3 * n)%N) + (A B C : 'M[cR]_(m, n)) (M : 'M[cR]_(n, p)) : + castmx (erefl _, e3n) + (row_mx A (row_mx B C)) *m (1%:M *t M) + = castmx (erefl _, esym (mul3n _)) + (row_mx (A *m M) (row_mx (B *m M) (C *m M))). +Proof. +apply/eqP; rewrite -(can2_eq (castmxKV _ _) (castmxK _ _)); apply/eqP. +rewrite tens_I3_mx mulmx_cast castmx_mul !castmx_comp /= !castmx_id /=. +by rewrite !mul_row_block /= !mulmx0 !addr0 !add0r. +Qed. + +(* :TODO: backport to polydiv *) +Lemma coprimep_rdiv_gcd p q : (p != 0) || (q != 0) -> + coprimep (rdivp p (gcdp p q)) (rdivp q (gcdp p q)). +Proof. +move=> hpq. +have gpq0: gcdp p q != 0 by rewrite gcdp_eq0 negb_and. +rewrite -gcdp_eqp1 -(@eqp_mul2r _ (gcdp p q)) // mul1r. +have: gcdp p q %| p by rewrite dvdp_gcdl. +have: gcdp p q %| q by rewrite dvdp_gcdr. +rewrite !dvdpE !rdvdp_eq eq_sym; move/eqP=> hq; rewrite eq_sym; move/eqP=> hp. +rewrite (eqp_ltrans (mulp_gcdl _ _ _)) hq hp. +have lcn0 k : (lead_coef (gcdp p q)) ^+ k != 0. + by rewrite expf_neq0 ?lead_coef_eq0. +by apply: eqp_gcd; rewrite ?eqp_scale. +Qed. + +(* :TODO: generalize to non idomainTypes and backport to polydiv *) +Lemma rgcdp_eq0 p q : rgcdp p q == 0 = (p == 0) && (q == 0). +Proof. by rewrite -eqp0 (eqp_ltrans (eqp_rgcd_gcd _ _)) eqp0 gcdp_eq0. Qed. + +(* :TODO: : move in polyorder *) +Lemma mu_eq0 : forall p x, p != 0 -> (\mu_x p == 0%N) = (~~ root p x). +Proof. by move=> p x p0; rewrite -mu_gt0 // -leqNgt leqn0. Qed. + +Notation lcn_neq0 := lc_expn_rscalp_neq0. + +(* :TODO: : move to polyorder *) +Lemma mu_mod p q x : (\mu_x p < \mu_x q)%N -> + \mu_x (rmodp p q) = \mu_x p. +Proof. +move=> mupq; have [->|p0] := eqVneq p 0; first by rewrite rmod0p. +have qn0 : q != 0 by apply: contraTneq mupq => ->; rewrite mu0 ltn0. +have /(canLR (addKr _)) <- := (rdivp_eq q p). +have [->|divpq_eq0] := eqVneq (rdivp p q) 0. + by rewrite mul0r oppr0 add0r mu_mulC ?lcn_neq0. +rewrite mu_addl ?mu_mulC ?scaler_eq0 ?negb_or ?mulf_neq0 ?lcn_neq0 //. +by rewrite mu_opp mu_mul ?ltn_addl // ?mulf_neq0. +Qed. + +(* :TODO: : move to polyorder *) +Lemma mu_add p q x : p + q != 0 -> + (minn (\mu_x p) (\mu_x q) <= \mu_x (p + q)%R)%N . +Proof. +have [->|p0] := eqVneq p 0; first by rewrite mu0 min0n add0r. +have [->|q0] := eqVneq q 0; first by rewrite mu0 minn0 addr0. +have [Hpq|Hpq|Hpq] := (ltngtP (\mu_x p) (\mu_x q)). ++ by rewrite mu_addr ?geq_minl. ++ by rewrite mu_addl ?geq_minr. +have [//|p' nrp'x hp] := (@mu_spec _ p x). +have [//|q' nrq'x hq] := (@mu_spec _ q x). +rewrite Hpq minnn hp {1 3}hq Hpq -mulrDl => pq0. +by rewrite mu_mul // mu_exp mu_XsubC mul1n leq_addl. +Qed. + +(* :TODO: : move to polydiv *) +Lemma mu_mod_leq : forall p q x, ~~ (q %| p) -> + (\mu_x q <= \mu_x p)%N -> + (\mu_x q <= \mu_x (rmodp p q)%R)%N. +Proof. +move=> p q x; rewrite dvdpE /rdvdp=> rn0 mupq. +case q0: (q == 0); first by rewrite (eqP q0) mu0 leq0n. +move/eqP: (rdivp_eq q p). +rewrite eq_sym (can2_eq (addKr _ ) (addNKr _)); move/eqP=> hr. +rewrite hr; case qpq0: (rdivp p q == 0). + by rewrite (eqP qpq0) mul0r oppr0 add0r mu_mulC // lcn_neq0. +rewrite (leq_trans _ (mu_add _ _)) // -?hr //. +rewrite leq_min mu_opp mu_mul ?mulf_neq0 ?qpq0 ?q0 // leq_addl. +by rewrite mu_mulC // lcn_neq0. +Qed. + +(* Lemma sgp_right0 : forall (x : R), sgp_right 0 x = 0. *) +(* Proof. by move=> x; rewrite /sgp_right size_poly0. Qed. *) + +End extra. + +Section ctmat. + +Variable R : numFieldType. + +Definition ctmat1 := \matrix_(i < 3, j < 3) + (nth [::] [:: [:: 1%:Z ; 1 ; 1 ] + ; [:: -1 ; 1 ; 1 ] + ; [:: 0 ; 0 ; 1 ] ] i)`_j. + +Lemma det_ctmat1 : \det ctmat1 = 2. +Proof. +(* Developpement direct ? *) +by do ?[rewrite (expand_det_row _ ord0) //=; + rewrite ?(big_ord_recl,big_ord0) //= ?mxE //=; + rewrite /cofactor /= ?(addn0, add0n, expr0, exprS); + rewrite ?(mul1r,mulr1,mulN1r,mul0r,mul1r,addr0) /=; + do ?rewrite [row' _ _]mx11_scalar det_scalar1 !mxE /=]. +Qed. + +Notation zmxR := ((map_mx ((intmul 1) : int -> R)) _ _). + +Lemma ctmat1_unit : zmxR ctmat1 \in unitmx. +Proof. +rewrite /mem /in_mem /= /unitmx det_map_mx //. +by rewrite det_ctmat1 unitfE intr_eq0. +Qed. + +Definition ctmat n := (ctmat1 ^t n). + +Lemma ctmat_unit : forall n, zmxR (ctmat n) \in unitmx. +Proof. +case=> [|n] /=; first by rewrite map_mx1 ?unitmx1//; apply: zinjR_morph. +elim:n=> [|n ihn] /=; first by apply: ctmat1_unit. +rewrite map_mxT //. +apply: tensmx_unit=> //; last exact: ctmat1_unit. +by elim: n {ihn}=> // n ihn; rewrite muln_eq0. +Qed. + +Lemma ctmat1_blocks : ctmat1 = (block_mx + 1 (row_mx 1 1) + (col_mx (-1) 0) (block_mx 1 1 0 1 : 'M_(1+1)%N)). +Proof. +apply/matrixP=> i j; rewrite !mxE. +by do 4?[case: splitP => ?; rewrite !mxE ?ord1=> ->]. +Qed. + +Lemma tvec_sub n : (3 * (3 ^ n).-1.+1 = 3 ^ (n.+1) )%N. +Proof. by rewrite -exp3n expnS. Qed. + +Lemma tens_ctmat1_mx n (M : 'M_n) : + ctmat1 *t M = castmx (esym (mul3n _ ), esym (mul3n _ )) + (block_mx M (row_mx M M) + (col_mx (-M) 0) (block_mx M M + 0 M : 'M_(n+n)%N)). +Proof. +rewrite ctmat1_blocks !tens_block_mx !tens_row_mx !tens_col_mx. +rewrite [-1]mx11_scalar !mxE /= !tens_scalar_mx !tens0mx scaleNr !scale1r. +apply/eqP; rewrite -(can2_eq (castmxKV _ _) (castmxK _ _)); apply/eqP. +rewrite !castmx_comp !esymK /=. +rewrite !(castmx_block (mul1n _) (mul1n _) (mul2n _) (mul2n _)). +rewrite !castmx_comp !castmx_id /=. +rewrite !(castmx_row (mul1n _) (mul1n _)). +rewrite !(castmx_block (mul1n _) (mul1n _) (mul1n _) (mul1n _)) /=. +rewrite !(castmx_col (mul1n _) (mul1n _)) !castmx_comp !castmx_id /=. +by rewrite !castmx_const. +Qed. + +Definition coefs n i := + [seq (castmx (erefl _, exp3n _) (invmx (zmxR (ctmat n)))) i ord0 + | i <- enum 'I__]`_i. + +End ctmat. + +Section QeRcfTh. + +Variable R : rcfType. +Implicit Types a b : R. +Implicit Types p q : {poly R}. + +Notation zmxR := ((map_mx ((intmul 1) : int -> R)) _ _). +Notation midf a b := ((a + b) / 2%:~R). + +(* Constraints and Tarski queries *) + +Local Notation sgp_is q s := (fun x => (sgr q.[x] == s)). + +Definition constraints (z : seq R) (sq : seq {poly R}) (sigma : seq int) := + (\sum_(x <- z) \prod_(i < size sq) (sgz (sq`_i).[x] == sigma`_i))%N. + +Definition taq (z : seq R) (q : {poly R}) : int := \sum_(x <- z) (sgz q.[x]). + +Lemma taq_constraint1 z q : + taq z q = (constraints z [::q] [::1])%:~R - (constraints z [::q] [::-1])%:~R. +Proof. +rewrite /constraints /taq !sumMz -sumrB /=; apply: congr_big=> // x _. +by rewrite !big_ord_recl big_ord0 !muln1 /=; case: sgzP. +Qed. + +Lemma taq_constraint0 z q : + taq z 1 = (constraints z [::q] [:: 0])%:~R + + (constraints z [::q] [:: 1])%:~R + + (constraints z [::q] [::-1])%:~R. +Proof. +rewrite /constraints /taq !sumMz //= -!big_split /=; apply: congr_big=> // x _. +by rewrite hornerC sgz1 !big_ord_recl big_ord0 !muln1 /=; case: sgzP. +Qed. + +Lemma taq_no_constraint z : taq z 1 = (constraints z [::] [::])%:~R. +Proof. +rewrite /constraints /taq !sumMz; apply: congr_big=> // x _. +by rewrite hornerC sgz1 big_ord0. +Qed. + +Lemma taq_constraint2 z q : + taq z (q ^+ 2) = (constraints z [::q] [:: 1])%:~R + + (constraints z [::q] [::-1])%:~R. +Proof. +rewrite /constraints /taq !sumMz -big_split /=; apply: congr_big=> // x _. +rewrite !big_ord_recl big_ord0 !muln1 /= horner_exp sgzX. +by case: (sgzP q.[x])=> _. +Qed. + +Fixpoint sg_tab n : seq (seq int) := + if n is m.+1 + then flatten (map (fun x => map (fun l => x :: l) (sg_tab m)) [::1;-1;0]) + else [::[::]]. + +Lemma sg_tab_nil n : (sg_tab n == [::]) = false. +Proof. by elim: n => //= n; case: sg_tab. Qed. + +Lemma size_sg_tab n : size (sg_tab n) = (3 ^ n)%N. +Proof. +by elim: n => [|n] // ihn; rewrite !size_cat !size_map ihn addn0 exp3S. +Qed. + +Lemma size_sg_tab_neq0 n : size (sg_tab n) != 0%N. +Proof. by rewrite size_sg_tab exp3n. Qed. + + +Definition comb_exp (R : realDomainType) (s : R) := + match sgz s with Posz 1 => 1%N | Negz 0 => 2 | _ => 0%N end. + +Definition poly_comb (sq : seq {poly R}) (sc : seq int) : {poly R} := + \prod_(i < size sq) ((sq`_i) ^+ (comb_exp sc`_i)). + +(* Eval compute in sg_tab 4. *) + +Definition cvec z sq := let sg_tab := sg_tab (size sq) in + \row_(i < 3 ^ size sq) ((constraints z sq (nth [::] sg_tab i))%:~R : int). +Definition tvec z sq := let sg_tab := sg_tab (size sq) in + \row_(i < 3 ^ size sq) (taq z (map (poly_comb sq) sg_tab)`_i). + + +Lemma tvec_cvec1 z q : tvec z [::q] = (cvec z [::q]) *m ctmat1. +Proof. +apply/rowP => j. +rewrite /tvec !mxE /poly_comb /= !big_ord_recl !big_ord0 //=. +rewrite !(expr0,expr1,mulr1) /=. +case: j=> [] [|[|[|j]]] hj //. +* by rewrite !mxE /= mulr0 add0r mulr1 mulrN1 addr0 taq_constraint1. +* by rewrite !mxE /= mulr0 !mulr1 add0r addr0 taq_constraint2. +* by rewrite !mxE /= addrA (@taq_constraint0 _ q) !mulr1 addr0 -addrA addrC. +Qed. + +Lemma cvec_rec z q sq : + cvec z (q :: sq) = castmx (erefl _, esym (exp3S _)) + (row_mx (cvec (filter (sgp_is q 1) z) (sq)) + (row_mx (cvec (filter (sgp_is q (-1)) z) (sq)) + (cvec (filter (sgp_is q 0) z) (sq)))). +Proof. +apply/eqP; rewrite -(can2_eq (castmxKV _ _) (castmxK _ _)); apply/eqP. +apply/rowP=> [] i; rewrite !(mxE, castmxE, esymK, cast_ord_id) /=. +symmetry; case: splitP=> j hj /=; rewrite !mxE hj. + case hst: sg_tab (sg_tab_nil (size sq))=> [|l st] // _. + have sst: (size st).+1 = (3 ^ size sq)%N. + transitivity (size (sg_tab (size sq))); first by rewrite hst //. + by rewrite size_sg_tab. + rewrite /constraints big_filter big_mkcond !sumMz; apply: congr_big=> // x _. + rewrite nth_cat size_map ![size (_::_)]/= sst ltn_ord. + rewrite (nth_map [::]) /= ?sst ?ltn_ord // big_ord_recl /=. + by rewrite sgr_cp0 sgz_cp0; case: (_ < _); first by rewrite mul1n. +case: splitP=> k hk; rewrite !mxE /= hk. + case hst: sg_tab (sg_tab_nil (size sq))=> [|l st] // _. + have sst: (size st).+1 = (3 ^ size sq)%N. + transitivity (size (sg_tab (size sq))); first by rewrite hst //. + by rewrite size_sg_tab. + rewrite /constraints big_filter big_mkcond !sumMz; apply: congr_big=> // x _. + rewrite nth_cat nth_cat !size_map ![size (_ :: _)]/= sst ltnNge leq_addr. + rewrite (@nth_map _ [::] _ _ [eta cons (-1)] _ (l::st)) /= ?sst addKn ltn_ord //. + rewrite big_ord_recl /= sgr_cp0 sgz_cp0. + by case: (_ < _); first by rewrite mul1n. +case hst: sg_tab (sg_tab_nil (size sq))=> [|l st] // _. +have sst: (size st).+1 = (3 ^ size sq)%N. + transitivity (size (sg_tab (size sq))); first by rewrite hst //. + by rewrite size_sg_tab. +rewrite /constraints big_filter big_mkcond !sumMz; apply: congr_big=> // x _. +rewrite nth_cat nth_cat nth_cat !size_map ![size (_ :: _)]/= sst. +rewrite (@nth_map _ [::] _ _ [eta cons 0] _ (l::st)) /=; last first. + by rewrite !addKn sst ltn_ord. +rewrite ltnNge leq_addr /= !addKn ltnNge leq_addr /= ltn_ord. +rewrite big_ord_recl /= sgr_cp0 sgz_cp0. +by case: (_ == _); first by rewrite mul1n. +Qed. + + +Lemma poly_comb_cons q sq s ss : + poly_comb (q :: sq) (s :: ss) = (q ^ (comb_exp s)) * poly_comb sq ss. +Proof. by rewrite /poly_comb /= big_ord_recl /=. Qed. + +Lemma comb_expE (rR : realDomainType): + (comb_exp (1 : rR) = 1%N) * (comb_exp (-1 : rR) = 2%N) * (comb_exp (0 : rR) = 0%N). +Proof. by rewrite /comb_exp sgzN sgz1 sgz0. Qed. + +Lemma tvec_rec z q sq : + tvec z (q :: sq) = + castmx (erefl _, esym (exp3S _)) ( + (row_mx (tvec (filter (sgp_is q 1) z) (sq)) + (row_mx (tvec (filter (sgp_is q (-1)) z) (sq)) + (tvec (filter (sgp_is q 0) z) (sq)))) *m + (castmx (mul3n _, mul3n _) (ctmat1 *t 1%:M))). +Proof. +rewrite tens_ctmat1_mx !castmx_comp !castmx_id /=. +rewrite !(mul_row_block, mul_row_col, mul_mx_row) !(mulmx1, mulmx0, mulmxN, addr0) /=. +apply/eqP; rewrite -(can2_eq (castmxKV _ _) (castmxK _ _)); apply/eqP. +apply/matrixP=> i j; rewrite !(castmxE, mxE) /=. +symmetry; case: splitP=> l hl; rewrite !mxE hl. + case hst: sg_tab (sg_tab_nil (size sq))=> [|s st] // _. + have sst: (size st).+1 = (3 ^ size sq)%N. + transitivity (size (sg_tab (size sq))); first by rewrite hst //. + by rewrite size_sg_tab. + rewrite /taq !big_filter !(big_mkcond (sgp_is _ _)) -sumrB. + apply: congr_big=> // x _. + rewrite cats0 !map_cat nth_cat !size_map /= sst ltn_ord /=. + rewrite !poly_comb_cons /= !comb_expE expr1z. + rewrite -!(nth_map _ 0 (fun p => p.[_])) /= ?size_map ?sst ?ltn_ord //. + rewrite -!map_comp /= hornerM. + set f := _ \o _; set g := _ \o _. + set h := fun sc => q.[x] * (poly_comb sq sc).[x]. + have hg : g =1 h. + by move=> sx; rewrite /g /h /= poly_comb_cons comb_expE expr1z hornerM. + rewrite -/(h _) -hg -[g _ :: _]/(map g (_ ::_)). + rewrite (nth_map [::]) /= ?sst ?ltn_ord // hg /h sgzM. + rewrite -![(poly_comb _ _).[_]]/(f _) -[f _ :: _]/(map f (_ ::_)). + rewrite (nth_map [::]) /= ?sst ?ltn_ord // !sgr_cp0. + by case: (sgzP q.[x]); rewrite ?(mul0r, mul1r, mulN1r, subr0, sub0r). +case: splitP=> k hk /=; rewrite !mxE hk. + case hst: sg_tab (sg_tab_nil (size sq))=> [|s st] // _. + have sst: (size st).+1 = (3 ^ size sq)%N. + transitivity (size (sg_tab (size sq))); first by rewrite hst //. + by rewrite size_sg_tab. + rewrite /taq !big_filter !(big_mkcond (sgp_is _ _)) -big_split. + apply: congr_big=> // x _. + rewrite cats0 !map_cat !nth_cat !size_map /= sst. + rewrite ltnNge leq_addr /= addKn ltn_ord /=. + rewrite !poly_comb_cons /= !comb_expE. + rewrite -!(nth_map _ 0 (fun p => p.[_])) /= ?size_map ?sst ?ltn_ord //. + rewrite -!map_comp /= hornerM. + set f := _ \o _; set g := _ \o _. + set h := fun sc => (q ^ 2).[x] * (poly_comb sq sc).[x]. + have hg : g =1 h. + by move=> sx; rewrite /g /h /= poly_comb_cons comb_expE hornerM. + rewrite -/(h _) -hg -[g _ :: _]/(map g (_ ::_)). + rewrite (nth_map [::]) /= ?sst ?ltn_ord // hg /h sgzM. + rewrite -![(poly_comb _ _).[_]]/(f _) -[f _ :: _]/(map f (_ ::_)). + rewrite (nth_map [::]) /= ?sst ?ltn_ord //. + rewrite hornerM sgzM !sgr_cp0. + by case: (sgzP q.[x]); rewrite ?(mul0r, mul1r, mulN1r, addr0, add0r). +case hst: sg_tab (sg_tab_nil (size sq))=> [|s st] // _. +have sst: (size st).+1 = (3 ^ size sq)%N. + transitivity (size (sg_tab (size sq))); first by rewrite hst //. + by rewrite size_sg_tab. +rewrite /taq !big_filter !(big_mkcond (sgp_is _ _)) -!big_split. +apply: congr_big=> // x _. +rewrite cats0 !map_cat !nth_cat !size_map /= sst. +rewrite !addKn 2!ltnNge !leq_addr /=. +rewrite !poly_comb_cons /= !comb_expE expr0z mul1r. +rewrite -!(nth_map _ 0 (fun p => p.[_])) /= ?size_map ?sst ?ltn_ord //. +rewrite -!map_comp /=. +set f := _ \o _; set g := _ \o _. +have hg : g =1 f. + by move=> sx; rewrite /g /f /= poly_comb_cons comb_expE expr0z mul1r. +rewrite -[(poly_comb _ _).[_]]/(f _) -{4}hg. +rewrite -[g s :: _]/(map _ (_ ::_)) (eq_map hg) !sgr_cp0. +by case: (sgzP q.[x])=> _; rewrite ?(addr0, add0r). +Qed. + +Lemma tvec_cvec z sq : + tvec z sq = (cvec z sq) *m (ctmat (size sq)). +Proof. +elim: sq z => [|q sq ihsq] z /=. + rewrite mulmx1; apply/rowP=> [] [i hi] /=; rewrite !mxE /=. + move: hi; rewrite expn0 ltnS leqn0; move/eqP=> -> /=. + rewrite /poly_comb big_ord0 /taq /constraints /=. + rewrite sumMz; apply: (congr_big)=> //= x _. + by rewrite hornerC sgz1 big_ord0. +rewrite /ctmat /ntensmx /=. (* simpl in trunk is "weaker" here *) +case: sq ihsq=> /= [|q' sq] ihsq; first by apply: tvec_cvec1. +rewrite cvec_rec tensmx_decl mulmxA tvec_rec. +apply/eqP; rewrite (can2_eq (castmxK _ _) (castmxKV _ _)); apply/eqP. +rewrite !castmx_mul !castmx_id [row_mx _ _ *m _]mulmx_cast. +congr (_ *m _); last by congr (castmx (_, _) _); apply: nat_irrelevance. +rewrite /=; have->: forall n, exp3S n.+1 = mul3n (3^n.+1)%N. + by move=> n; apply: nat_irrelevance. +by rewrite mul_1tensmx !ihsq. +Qed. + +Lemma cvec_tvec z sq : + zmxR (cvec z (sq)) = (zmxR (tvec z (sq))) *m (invmx (zmxR (ctmat (size (sq))))). +Proof. +apply/eqP; set A := zmxR (ctmat _). +rewrite -(@can2_eq _ _ (fun (x : 'rV_(_)) => x *m A) (fun x => x *m (invmx A))). +* by rewrite /A -map_mxM ?tvec_cvec//; apply: zinjR_morph. +* by apply: mulmxK; rewrite /A ctmat_unit. +* by apply: mulmxKV; rewrite /A ctmat_unit. +Qed. + +Lemma constraints1_tvec : forall z sq, + (constraints z (sq) (nseq (size (sq)) 1))%:~R = (castmx (erefl _, exp3n _) + (zmxR (tvec z (sq)) *m (invmx (zmxR (ctmat (size (sq))))))) ord0 ord0. +Proof. +move=> z sq. +rewrite -cvec_tvec castmxE /= cast_ord_id /= /cvec !mxE //= intz. +congr ((constraints _ _ _)%:~R); elim: sq=> //= _ s -> /=. +set l := sg_tab _; suff: size l != 0%N by case: l. +exact: size_sg_tab_neq0. +Qed. + +(* Cauchy Index, relation with Tarski query*) + +Local Notation seq_mids a s b := (pairmap (fun x y => midf x y) a (rcons s b)). +Local Notation noroot p := (forall x, ~~ root p x). +Notation lcn_neq0 := lc_expn_rscalp_neq0. + +Definition jump q p x: int := + let non_null := (q != 0) && odd (\mu_x p - \mu_x q) in + let sign := (sgp_right (q * p) x < 0)%R in + (-1) ^+ sign *+ non_null. + +Definition cindex (a b : R) (q p : {poly R}) : int := + \sum_(x <- roots p a b) jump q p x. + +Definition cindexR q p := \sum_(x <- rootsR p) jump q p x. + +Definition sjump p x : int := + ((-1) ^+ (sgp_right p x < 0)%R) *+ odd (\mu_x p). + +Definition variation (x y : R) : int := (sgz y) * (x * y < 0). + +Definition cross p a b := variation p.[a] p.[b]. + +Definition crossR p := variation (sgp_minfty p) (sgp_pinfty p). + +Definition sum_var (s : seq R) := \sum_(n <- pairmap variation 0 s) n. + +Lemma cindexEba a b : b <= a -> forall p q, cindex a b p q = 0. +Proof. by move=> le_ba p q; rewrite /cindex rootsEba ?big_nil. Qed. + +Lemma jump0p q x : jump 0 q x = 0. Proof. by rewrite /jump eqxx mulr0n. Qed. + +Lemma taq_cindex a b p q : taq (roots p a b) q = cindex a b (p^`() * q) p. +Proof. +have [lt_ab|?] := ltrP a b; last by rewrite rootsEba ?cindexEba /taq ?big_nil. +rewrite /taq /cindex !big_seq; apply: eq_bigr => x. +have [->|p_neq0 /root_roots rpx] := eqVneq p 0; first by rewrite roots0 in_nil. +have [->|q_neq0] := eqVneq q 0; first by rewrite mulr0 jump0p horner0 sgz0. +have [p'0|p'_neq0] := eqVneq p^`() 0. + move/(root_size_gt1 p_neq0): rpx. + by rewrite -subn_gt0 subn1 -size_deriv p'0 size_poly0. +have p'q0: p^`() * q != 0 by rewrite mulf_neq0. +move:(p'q0); rewrite mulf_eq0 negb_or; case/andP=> p'0 q0. +have p0: p != 0 by move: p'0; apply: contra; move/eqP->; rewrite derivC. +rewrite /jump mu_mul// {1}(@mu_deriv_root _ _ p)// addn1 p'q0 /=. +case emq: (\mu_(_) q)=> [|m]. + move/eqP: emq; rewrite -leqn0 leqNgt mu_gt0// => qxn0. + rewrite addn0 subSnn mulr1n. + rewrite !sgp_right_mul// (sgp_right_deriv rpx) mulrAC. + rewrite sgp_right_square// mul1r sgp_rightNroot//. + rewrite sgr_lt0 -sgz_cp0. + by move: qxn0; rewrite -[root q x]sgz_cp0; case: sgzP. +rewrite addnS subSS -{1}[\mu_(_) _]addn0 subnDl sub0n mulr0n. +by apply/eqP; rewrite sgz_cp0 -[_ == 0]mu_gt0// emq. +Qed. + +Lemma sum_varP s : 0 \notin s -> sum_var s = variation (head 0 s) (last 0 s). +Proof. +rewrite /sum_var /variation. +case: s => /= [_|a s]; first by rewrite big_nil sgz0 mul0r. +rewrite in_cons big_cons mul0r ltrr mulr0 add0r. +elim: s a => [|b s IHs] a; first by rewrite big_nil ler_gtF ?mulr0 ?sqr_ge0. +move=> /norP [neq_0a Hbs]; move: (Hbs); rewrite in_cons => /norP[neq_0b Hs]. +rewrite /= big_cons IHs ?negb_or ?neq_0b // -!sgz_cp0 !sgzM. +have: (last b s) != 0 by apply: contra Hbs => /eqP <-; rewrite mem_last. +by move: neq_0a neq_0b; do 3?case: sgzP => ? //. +Qed. + +Lemma jump_coprime p q : p != 0 -> coprimep p q + -> forall x, root p x -> jump q p x = sjump (q * p) x. +Proof. +move=> pn0 cpq x rpx; rewrite /jump /sjump. +have q_neq0 : q != 0; last rewrite q_neq0 /=. + apply: contraTneq cpq => ->; rewrite coprimep0. + by apply: contraL rpx => /eqp_root ->; rewrite rootC oner_eq0. +have := coprimep_root cpq rpx; rewrite -rootE -mu_eq0 => // /eqP muxq_eq0. +by rewrite mu_mul ?mulf_neq0 ?muxq_eq0 ?subn0 ?add0n. +Qed. + +Lemma sjump_neigh a b p x : p != 0 -> + {in neighpl p a x & neighpr p x b, + forall yl yr, sjump p x = cross p yl yr}. +Proof. +move=> pn0 yl yr yln yrn; rewrite /cross /variation. +rewrite -sgr_cp0 sgrM /sjump (sgr_neighpl yln) -!(sgr_neighpr yrn). +rewrite -mulrA -expr2 sqr_sg (rootPf (neighpr_root yrn)) mulr1. +rewrite sgrEz ltrz0 -[in rhs in _ = rhs]intr_sign -[X in _ == X]mulrN1z eqr_int. +by have /rootPf := neighpr_root yrn; case: sgzP; case: odd. +Qed. + +Lemma jump_neigh a b p q x : q * p != 0 -> + {in neighpl (q * p) a x & neighpr (q * p) x b, forall yl yr, + jump q p x = cross (q * p) yl yr *+ ((q != 0) && (\mu_x p > \mu_x q)%N)}. +Proof. +move=> pqn0 yl yr hyl hyr; rewrite -(sjump_neigh pqn0 hyl hyr). +rewrite /jump /sjump -mulrnA mulnb andbCA. +have [muqp|/eqnP ->] := ltnP; rewrite (andbF, andbT) //. +by rewrite mu_mul // odd_add addbC odd_sub // ltnW. +Qed. + +Lemma jump_mul2l (p q r : {poly R}) : + p != 0 -> jump (p * q) (p * r) =1 jump q r. +Proof. +move=> p0 x; rewrite /jump. +case q0: (q == 0); first by rewrite (eqP q0) mulr0 eqxx. +have ->: p * q != 0 by rewrite mulf_neq0 ?p0 ?q0. +case r0: (r == 0); first by rewrite (eqP r0) !mulr0 mu0 !sub0n. +rewrite !mu_mul ?mulf_neq0 ?andbT ?q0 ?r0 //; rewrite subnDl. +rewrite mulrAC mulrA -mulrA. +rewrite (@sgp_right_mul _ (p * p)) // sgp_right_mul // sgp_right_square //. +by rewrite mul1r mulrC /=. +Qed. + +Lemma jump_mul2r (p q r : {poly R}) : + p != 0 -> jump (q * p) (r * p) =1 jump q r. +Proof. by move=> p0 x; rewrite ![_ * p]mulrC jump_mul2l. Qed. + +Lemma jumppc p c x : jump p c%:P x = 0. +Proof. by rewrite /jump mu_polyC sub0n !andbF. Qed. + +Lemma noroot_jump q p x : ~~ root p x -> jump q p x = 0. +Proof. +have [->|p_neq0] := eqVneq p 0; first by rewrite jumppc. +by rewrite -mu_gt0 // lt0n negbK /jump => /eqP ->; rewrite andbF mulr0n. +Qed. + +Lemma jump_mulCp c p q x : jump (c *: p) q x = (sgz c) * jump p q x. +Proof. +have [->|c0] := eqVneq c 0; first by rewrite sgz0 scale0r jump0p mul0r. +have [->|p0] := eqVneq p 0; first by rewrite scaler0 jump0p mulr0. +have [->|q0] := eqVneq q 0; first by rewrite !jumppc mulr0. +(* :TODO: : rename mu_mulC *) +rewrite /jump scale_poly_eq0 mu_mulC ?negb_or ?c0 ?p0 ?andTb //. +rewrite -scalerAl sgp_right_scale //. +case: sgzP c0 => // _ _; rewrite !(mul1r, mulN1r, =^~ mulNrn) //. +by rewrite ?oppr_cp0 lt0r sgp_right_eq0 ?mulf_neq0 // andTb lerNgt signrN. +Qed. + +Lemma jump_pmulC c p q x : jump p (c *: q) x = (sgz c) * jump p q x. +Proof. +have [->|c0] := eqVneq c 0; first by rewrite sgz0 scale0r mul0r jumppc. +have [->|p0] := eqVneq p 0; first by rewrite !jump0p mulr0. +have [->|q0] := eqVneq q 0; first by rewrite scaler0 !jumppc mulr0. +rewrite /jump mu_mulC // -scalerAr sgp_right_scale //. +case: sgzP c0 => // _ _; rewrite !(mul1r, mulN1r, =^~ mulNrn) //. +by rewrite ?oppr_cp0 lt0r sgp_right_eq0 ?mulf_neq0 // andTb lerNgt signrN. +Qed. + +Lemma jump_mod p q x : + jump p q x = sgz (lead_coef q) ^+ (rscalp p q) * jump (rmodp p q) q x. +Proof. +case p0: (p == 0); first by rewrite (eqP p0) rmod0p jump0p mulr0. +case q0: (q == 0); first by rewrite (eqP q0) rmodp0 jumppc mulr0. +rewrite -sgzX; set s := sgz _. +apply: (@mulfI _ s); first by rewrite /s sgz_eq0 lcn_neq0. +rewrite mulrA mulz_sg lcn_neq0 mul1r -jump_mulCp rdivp_eq. +have [->|rpq_eq0] := altP (rmodp p q =P 0). + by rewrite addr0 jump0p -[X in jump _ X]mul1r jump_mul2r ?q0 // jumppc. +rewrite /jump. set r := _ * q + _. +have muxp : \mu_x p = \mu_x r by rewrite /r -rdivp_eq mu_mulC ?lcn_neq0. +have r_neq0 : r != 0 by rewrite /r -rdivp_eq scaler_eq0 p0 orbF lcn_neq0. +have [hpq|hpq] := leqP (\mu_x q) (\mu_x r). + rewrite 2!(_ : _ - _ = 0)%N ?andbF //; apply/eqP; rewrite -/(_ <= _)%N //. + by rewrite mu_mod_leq ?dvdpE // muxp. +rewrite mu_mod ?muxp // rpq_eq0 (negPf r_neq0); congr (_ ^+ _ *+ _). +rewrite !sgp_right_mul sgp_right_mod ?muxp // /r -rdivp_eq. +by rewrite -mul_polyC sgp_right_mul sgp_rightc sgrX. +Qed. + +Lemma cindexRP q p a b : + {in `]-oo, a], noroot p} -> {in `[b , +oo[, noroot p} -> + cindex a b q p = cindexR q p. +Proof. by rewrite /cindex => rpa rpb; rewrite rootsRP. Qed. + +Lemma cindex0p a b q : cindex a b 0 q = 0. +Proof. +have [lt_ab|le_ba] := ltrP a b; last by rewrite cindexEba. +by apply: big1_seq=> x; rewrite /jump eqxx mulr0n. +Qed. + +Lemma cindexR0p p : cindexR 0 p = 0. +Proof. by rewrite /cindexR big1 // => q _; rewrite jump0p. Qed. + +Lemma cindexpC a b p c : cindex a b p (c%:P) = 0. +Proof. +have [lt_ab|le_ba] := ltrP a b; last by rewrite cindexEba. +by rewrite /cindex /jump rootsC big_nil. +Qed. + +Lemma cindexRpC q c : cindexR q c%:P = 0. +Proof. by rewrite /cindexR rootsRC big_nil. Qed. + +Lemma cindex_mul2r a b p q r : r != 0 -> + cindex a b (p * r) (q * r) = cindex a b p q. +Proof. +have [lt_ab r0|le_ba] := ltrP a b; last by rewrite !cindexEba. +have [->|p0] := eqVneq p 0; first by rewrite mul0r !cindex0p. +have [->|q0] := eqVneq q 0; first by rewrite mul0r !cindexpC. +rewrite /cindex (eq_big_perm _ (roots_mul _ _ _))//= big_cat/=. +rewrite -[\sum_(x <- _) jump p _ _]addr0; congr (_+_). + by rewrite !big_seq; apply: congr_big => // x hx; rewrite jump_mul2r. +rewrite big1_seq//= => x hx; rewrite jump_mul2r // /jump. +suff ->: \mu_x q = 0%N by rewrite andbF. +apply/eqP; rewrite -leqn0 leqNgt mu_gt0 //. +apply/negP; rewrite root_factor_theorem => rqx; move/root_roots:hx. +case: gdcopP=> g hg; rewrite (negPf r0) orbF => cgq hdg. +rewrite root_factor_theorem=> rgx. +move/coprimepP:cgq rqx rgx=> cgq; rewrite -!dvdpE=> /cgq hgq /hgq. +by rewrite -size_poly_eq1 size_XsubC. +Qed. + +Lemma cindex_mulCp a b p q c : + cindex a b (c *: p) q = (sgz c) * cindex a b p q. +Proof. +have [lt_ab|le_ba] := ltrP a b; last by rewrite !cindexEba ?mulr0. +have [->|p0] := eqVneq p 0; first by rewrite !(cindex0p, scaler0, mulr0). +have [->|q0] := eqVneq q 0; first by rewrite !(cindexpC, scaler0, mulr0). +by rewrite /cindex big_distrr; apply: congr_big => //= x; rewrite jump_mulCp. +Qed. + +Lemma cindex_pmulC a b p q c : + cindex a b p (c *: q) = (sgz c) * cindex a b p q. +Proof. +have [lt_ab|le_ba] := ltrP a b; last by rewrite !cindexEba ?mulr0. +have [->|p0] := eqVneq p 0; first by rewrite !(cindex0p, scaler0, mulr0). +have [->|q0] := eqVneq q 0; first by rewrite !(cindexpC, scaler0, mulr0). +have [->|c0] := eqVneq c 0; first by rewrite scale0r sgz0 mul0r cindexpC. +rewrite /cindex big_distrr rootsZ //. +by apply: congr_big => // x _; rewrite jump_pmulC. +Qed. + +Lemma cindex_mod a b p q : + cindex a b p q = + (sgz (lead_coef q) ^+ rscalp p q) * cindex a b (rmodp p q) q. +Proof. +have [lt_ab|le_ba] := ltrP a b; last by rewrite !cindexEba ?mulr0. +by rewrite /cindex big_distrr; apply: congr_big => // x; rewrite jump_mod. +Qed. + +Lemma variation0r b : variation 0 b = 0. +Proof. by rewrite /variation mul0r ltrr mulr0. Qed. + +Lemma variationC a b : variation a b = - variation b a. +Proof. by rewrite /variation -!sgz_cp0 !sgzM; do 2?case: sgzP => _ //. Qed. + +Lemma variationr0 a : variation a 0 = 0. +Proof. by rewrite variationC variation0r oppr0. Qed. + +Lemma variation_pmull a b c : c > 0 -> variation (a * c) (b) = variation a b. +Proof. by move=> c_gt0; rewrite /variation mulrAC pmulr_llt0. Qed. + +Lemma variation_pmulr a b c : c > 0 -> variation a (b * c) = variation a b. +Proof. by move=> c_gt0; rewrite variationC variation_pmull // -variationC. Qed. + +Lemma congr_variation a b a' b' : sgz a = sgz a' -> sgz b = sgz b' -> + variation a b = variation a' b'. +Proof. by rewrite /variation -!sgz_cp0 !sgzM => -> ->. Qed. + +Lemma crossRP p a b : + {in `]-oo, a], noroot p} -> {in `[b , +oo[, noroot p} -> + cross p a b = crossR p. +Proof. +move=> rpa rpb; rewrite /crossR /cross. +rewrite -(@sgp_minftyP _ _ _ rpa a) ?boundr_in_itv //. +rewrite -(@sgp_pinftyP _ _ _ rpb b) ?boundl_in_itv //. +by rewrite /variation -sgrM sgr_lt0 sgz_sgr. +Qed. + +Lemma noroot_cross p a b : a <= b -> + {in `]a, b[, noroot p} -> cross p a b = 0. +Proof. +move=> le_ab noroot_ab; rewrite /cross /variation. +have [] := ltrP; last by rewrite mulr0. +rewrite mulr1 -sgr_cp0 sgrM => /eqP. +by move=> /(ivt_sign le_ab) [x /noroot_ab /negPf->]. +Qed. + +Lemma cross_pmul p q a b : q.[a] > 0 -> q.[b] > 0 -> + cross (p * q) a b = cross p a b. +Proof. +by move=> qa0 qb0; rewrite /cross !hornerM variation_pmull ?variation_pmulr. +Qed. + +Lemma cross0 a b : cross 0 a b = 0. +Proof. by rewrite /cross !horner0 variation0r. Qed. + +Lemma crossR0 : crossR 0 = 0. +Proof. +by rewrite /crossR /sgp_minfty /sgp_pinfty lead_coef0 mulr0 sgr0 variationr0. +Qed. + +Lemma cindex_seq_mids a b : a < b -> + forall p q, p != 0 -> q != 0 -> coprimep p q -> + cindex a b q p + cindex a b p q = + sum_var (map (horner (p * q)) (seq_mids a (roots (p * q) a b) b)). +Proof. +move=> hab p q p0 q0 cpq; rewrite /cindex /sum_var 2!big_seq. +have pq_neq0 : p * q != 0 by rewrite mulf_neq0. +have pq_eq0 := negPf pq_neq0. +have jumpP : forall (p q : {poly R}), p != 0 -> coprimep p q -> + forall x, x \in roots p a b -> jump q p x = sjump (q * p) x. + by move=> ? ? ? ? ?; move/root_roots=> ?; rewrite jump_coprime. +rewrite !(eq_bigr _ (jumpP _ _ _ _))// 1?coprimep_sym// => {jumpP}. +have sjumpC x : sjump (q * p) x = sjump (p * q) x by rewrite mulrC. +rewrite -!big_seq (eq_bigr _ (fun x _ => sjumpC x)). +rewrite -big_cat /= -(eq_big_perm _ (roots_mul_coprime _ _ _ _)) //=. +move: {1 2 5}a hab (erefl (roots (p * q) a b)) => //=. +elim: roots => {a} [|x s /= ihs] a hab /eqP. + by rewrite big_cons !big_nil variation0r. +rewrite roots_cons; case/and5P => _ xab /eqP hax hx /eqP hs. +rewrite !big_cons variation0r add0r (ihs _ _ hs) ?(itvP xab) // => {ihs}. +pose y := (head b s); pose ax := midf a x; pose xy := midf x y. +rewrite (@sjump_neigh a b _ _ _ ax xy) ?inE ?midf_lte//=; last 2 first. ++ by rewrite /prev_root pq_eq0 hax minr_l ?(itvP xab, midf_lte). ++ have hy: y \in `]x, b]. + rewrite /y; case: s hs {y xy} => /= [|u s] hu. + by rewrite boundr_in_itv /= ?(itvP xab). + have /roots_in: u \in roots (p * q) x b by rewrite hu mem_head. + by apply: subitvP; rewrite /= !lerr. + by rewrite /next_root pq_eq0 hs maxr_l ?(itvP hy, midf_lte). +move: @y @xy {hs}; rewrite /cross. +by case: s => /= [|y l]; rewrite ?(big_cons, big_nil, variation0r, add0r). +Qed. + +Lemma cindex_inv a b : a < b -> forall p q, + ~~ root (p * q) a -> ~~ root (p * q) b -> + cindex a b q p + cindex a b p q = cross (p * q) a b. +Proof. +move=> hab p q hpqa hpqb. +have hlab: a <= b by apply: ltrW. +wlog cpq: p q hpqa hpqb / coprimep p q => [hwlog|]. + have p0: p != 0 by apply: contraNneq hpqa => ->; rewrite mul0r rootC. + have q0: q != 0 by apply: contraNneq hpqa => ->; rewrite mulr0 rootC. + set p' := p; rewrite -(divpK (dvdp_gcdr p q)) -[p'](divpK (dvdp_gcdl p q)). + rewrite !cindex_mul2r ?gcdp_eq0 ?(negPf p0) //. + have ga0 : (gcdp p q).[a] != 0. + apply: contra hpqa; rewrite -rootE -!dvdp_XsubCl => /dvdp_trans -> //. + by rewrite dvdp_mulr ?dvdp_gcdl. + have gb0 : (gcdp p q).[b] != 0. + apply: contra hpqb; rewrite -rootE -!dvdp_XsubCl => /dvdp_trans -> //. + by rewrite dvdp_mulr ?dvdp_gcdl. + rewrite mulrACA -expr2 cross_pmul ?horner_exp ?exprn_even_gt0 ?ga0 ?gb0 //. + apply: hwlog; rewrite ?coprimep_div_gcd ?p0 // rootM. + + apply: contra hpqa; rewrite -!dvdp_XsubCl => /orP. + case=> /dvdp_trans-> //; rewrite (dvdp_trans (divp_dvd _)); + by rewrite ?(dvdp_gcdl, dvdp_gcdr) ?(dvdp_mulIl, dvdp_mulIr). + + apply: contra hpqb; rewrite -!dvdp_XsubCl => /orP. + case=> /dvdp_trans-> //; rewrite (dvdp_trans (divp_dvd _)); + by rewrite ?(dvdp_gcdl, dvdp_gcdr) ?(dvdp_mulIl, dvdp_mulIr). +have p0: p != 0 by apply: contraNneq hpqa => ->; rewrite mul0r rootC. +have q0: q != 0 by apply: contraNneq hpqa => ->; rewrite mulr0 rootC. +have pq0 : p * q != 0 by rewrite mulf_neq0. +rewrite cindex_seq_mids // sum_varP /cross. + apply: congr_variation; apply: (mulrIz (oner_neq0 R)); rewrite -!sgrEz. + case hr: roots => [|c s] /=; apply: (@sgr_neighprN _ _ a b) => //; + rewrite /neighpr /next_root ?(negPf pq0) maxr_l // hr mid_in_itv //=. + by move/eqP: hr; rewrite roots_cons => /and5P [_ /itvP ->]. + rewrite -cats1 pairmap_cat /= cats1 map_rcons last_rcons. + apply: (@sgr_neighplN _ _ a b) => //. + rewrite /neighpl /prev_root (negPf pq0) minr_l //. + by rewrite mid_in_itv //= last_roots_le. +elim: roots {-2 6}a (erefl (roots (p * q) a b)) + {hpqa hpqb} hab hlab => {a} [|c s IHs] a Hs hab hlab /=. + rewrite in_cons orbF eq_sym. (* ; set x := (X in _.[X]). *) + by rewrite -rootE (@roots_nil _ _ a b) // mid_in_itv. +move/eqP: Hs; rewrite roots_cons => /and5P [_ cab /eqP rac rc /eqP rcb]. +rewrite in_cons eq_sym -rootE negb_or (roots_nil _ rac) //=; last first. + by rewrite mid_in_itv //= (itvP cab). +by rewrite IHs // (itvP cab). +Qed. + +Definition next_mod p q := - (lead_coef q ^+ rscalp p q) *: rmodp p q. + +Lemma next_mod0p q : next_mod 0 q = 0. +Proof. by rewrite /next_mod rmod0p scaler0. Qed. + +Lemma cindex_rec a b : a < b -> forall p q, + ~~ root (p * q) a -> ~~ root (p * q) b -> + cindex a b q p = cross (p * q) a b + cindex a b (next_mod p q) q. +Proof. +move=> lt_ab p q rpqa rpqb; have [->|p0] := eqVneq p 0. + by rewrite cindexpC next_mod0p cindex0p mul0r cross0 add0r. +have [->|q0] := eqVneq q 0. + by rewrite cindex0p cindexpC mulr0 cross0 add0r. +have /(canRL (addrK _)) -> := cindex_inv lt_ab rpqa rpqb. +by rewrite cindex_mulCp cindex_mod sgzN mulNr sgzX. +Qed. + +Lemma cindexR_rec p q : + cindexR q p = crossR (p * q) + cindexR (next_mod p q) q. +Proof. +have [->|p_neq0] := eqVneq p 0. + by rewrite cindexRpC mul0r next_mod0p cindexR0p crossR0. +have [->|q_neq0] := eqVneq q 0. + by rewrite cindexR0p mulr0 crossR0 cindexRpC. +have pq_neq0 : p * q != 0 by rewrite mulf_neq0. +pose b := cauchy_bound (p * q). +have [lecb gecb] := pair (le_cauchy_bound pq_neq0) (ge_cauchy_bound pq_neq0). +rewrite -?(@cindexRP _ _ (-b) b); do ? + by [move=> x Hx /=; have: ~~ root (p * q) x by [apply: lecb|apply: gecb]; + rewrite rootM => /norP []]. +rewrite -(@crossRP _ (-b) b) 1?cindex_rec ?gt0_cp ?cauchy_bound_gt0 //. + by rewrite lecb // boundr_in_itv. +by rewrite gecb // boundl_in_itv. +Qed. + +(* Computation of cindex through changes_mods *) + +Definition mods p q := + let fix aux p q n := + if n is m.+1 + then if p == 0 then [::] else p :: (aux q (next_mod p q) m) + else [::] in aux p q (maxn (size p) (size q).+1). + +Lemma mods_rec p q : mods p q = + if p == 0 then [::] else p :: (mods q (next_mod p q)). +Proof. +rewrite /mods; set aux := fix aux _ _ n := if n is _.+1 then _ else _. +have aux0 u n : aux 0 u n = [::] by case: n => [//|n] /=; rewrite eqxx. +pose m p q := maxn (size p) (size q).+1; rewrite -!/(m _ _). +suff {p q} Hnext p q : q != 0 -> (m q (next_mod p q) < m p q)%N; last first. + rewrite /m -maxnSS leq_max !geq_max !ltnS leqnn /= /next_mod. + rewrite size_scale ?oppr_eq0 ?lcn_neq0 //=. + by move=> q_neq0; rewrite ltn_rmodp ?q_neq0 ?orbT. +suff {p q} m_gt0 p q : (0 < m p q)%N; last by rewrite leq_max orbT. +rewrite -[m p q]prednK //=; have [//|p_neq0] := altP (p =P 0). +have [->|q_neq0] := altP (q =P 0); first by rewrite !aux0. +congr (_ :: _); suff {p q p_neq0 q_neq0} Haux p q n n' : + (m p q <= n)%N -> (m p q <= n')%N -> aux p q n = aux p q n'. + by apply: Haux => //; rewrite -ltnS prednK // Hnext. +elim: n p q n' => [p q|n IHn p q n' Hn]; first by rewrite geq_max ltn0 andbF. +case: n' => [|n' Hn' /=]; first by rewrite geq_max ltn0 andbF. +have [//|p_neq0] := altP eqP; congr (_ :: _). +have [->|q_neq0] := altP (q =P 0); first by rewrite !aux0. +by apply: IHn; rewrite -ltnS (leq_trans _ Hn, leq_trans _ Hn') ?Hnext. +Qed. + +Lemma mods_eq0 p q : (mods p q == [::]) = (p == 0). +Proof. by rewrite mods_rec; have [] := altP (p =P 0). Qed. + +Lemma neq0_mods_rec p q : p != 0 -> mods p q = p :: mods q (next_mod p q). +Proof. by rewrite mods_rec => /negPf ->. Qed. + +Lemma mods0p q : mods 0 q = [::]. +Proof. by apply/eqP; rewrite mods_eq0. Qed. + +Lemma modsp0 p : mods p 0 = if p == 0 then [::] else [::p]. +Proof. by rewrite mods_rec mods0p. Qed. + +Fixpoint changes (s : seq R) : nat := + (if s is a :: q then (a * (head 0 q) < 0)%R + changes q else 0)%N. + +Definition changes_pinfty (p : seq {poly R}) := changes (map lead_coef p). +Definition changes_minfty (p : seq {poly R}) := + changes (map (fun p : {poly _} => (-1) ^+ (~~ odd (size p)) * lead_coef p) p). + +Definition changes_poly (p : seq {poly R}) := + (changes_minfty p)%:Z - (changes_pinfty p)%:Z. +Definition changes_mods p q := changes_poly (mods p q). + +Lemma changes_mods0p q : changes_mods 0 q = 0. +Proof. by rewrite /changes_mods /changes_poly mods0p. Qed. + +Lemma changes_modsp0 p : changes_mods p 0 = 0. +Proof. +rewrite /changes_mods /changes_poly modsp0; have [//|p_neq0] := altP eqP. +by rewrite /changes_minfty /changes_pinfty /= !mulr0 ltrr. +Qed. + +Lemma changes_mods_rec p q : + changes_mods p q = crossR (p * q) + changes_mods q (next_mod p q). +Proof. +have [->|p0] := eqVneq p 0. + by rewrite changes_mods0p mul0r crossR0 next_mod0p changes_modsp0. +have [->|q0] := eqVneq q 0. + by rewrite changes_modsp0 mulr0 crossR0 changes_mods0p. +rewrite /changes_mods /changes_poly neq0_mods_rec //=. +rewrite !PoszD opprD addrACA; congr (_ + _); rewrite neq0_mods_rec //=. +rewrite /crossR /variation /sgp_pinfty /sgp_minfty. +rewrite mulr_signM size_mul // !lead_coefM. +rewrite polySpred // addSn [size q]polySpred // addnS /= !negbK. +rewrite -odd_add signr_odd; set s := _ ^+ _. +rewrite -!sgz_cp0 !(sgz_sgr, sgzM). +have: s != 0 by rewrite signr_eq0. +by move: p0 q0; rewrite -!lead_coef_eq0; do 3!case: sgzP=> _. +Qed. + +Lemma changes_mods_cindex p q : changes_mods p q = cindexR q p. +Proof. +elim: mods {-2}p {-2}q (erefl (mods p q)) => [|r s IHs] {p q} p q hrpq. + move/eqP: hrpq; rewrite mods_eq0 => /eqP ->. + by rewrite changes_mods0p cindexRpC. +rewrite changes_mods_rec cindexR_rec IHs //. +by move: hrpq IHs; rewrite mods_rec; case: (p == 0) => // [] []. +Qed. + +Definition taqR p q := changes_mods p (p^`() * q). + +Lemma taq_taqR p q : taq (rootsR p) q = taqR p q. +Proof. by rewrite /taqR changes_mods_cindex taq_cindex. Qed. + +Section ChangesItvMod_USELESS. +(* Not used anymore, but the content of this section is *) +(* used in the LMCS 2012 paper and in Cyril's thesis *) + +Definition changes_horner (p : seq {poly R}) x := + changes (map (fun p => p.[x]) p). +Definition changes_itv_poly a b (p : seq {poly R}) := + (changes_horner p a)%:Z - (changes_horner p b)%:Z. + +Definition changes_itv_mods a b p q := changes_itv_poly a b (mods p q). + +Lemma changes_itv_mods0p a b q : changes_itv_mods a b 0 q = 0. +Proof. +by rewrite /changes_itv_mods /changes_itv_poly mods0p /changes_horner /= subrr. +Qed. + +Lemma changes_itv_modsp0 a b p : changes_itv_mods a b p 0 = 0. +Proof. +rewrite /changes_itv_mods /changes_itv_poly modsp0 /changes_horner /=. +by have [//|p_neq0 /=] := altP eqP; rewrite !mulr0 ltrr. +Qed. + +Lemma changes_itv_mods_rec a b : a < b -> forall p q, + ~~ root (p * q) a -> ~~ root (p * q) b -> + changes_itv_mods a b p q = cross (p * q) a b + + changes_itv_mods a b q (next_mod p q). +Proof. +move=> lt_ab p q rpqa rpqb. +have [->|p0] := eqVneq p 0. + by rewrite changes_itv_mods0p mul0r next_mod0p changes_itv_modsp0 cross0. +have [->|q0] := eqVneq q 0. + by rewrite changes_itv_modsp0 mulr0 cross0 changes_itv_mods0p. +rewrite /changes_itv_mods /changes_itv_poly /changes_horner neq0_mods_rec //=. +rewrite !PoszD opprD addrACA; congr (_ + _); rewrite neq0_mods_rec //=. +move: rpqa rpqb; rewrite -!hornerM !rootE; move: (p * q) => r {p q p0 q0}. +by rewrite /cross /variation -![_ < _]sgz_cp0 sgzM; do 2!case: sgzP => _. +Qed. + +Lemma changes_itv_mods_cindex a b : a < b -> forall p q, + all (fun p => ~~ root p a) (mods p q) -> + all (fun p => ~~ root p b) (mods p q) -> + changes_itv_mods a b p q = cindex a b q p. +Proof. +move=> hab p q. +elim: mods {-2}p {-2}q (erefl (mods p q)) => [|r s IHs] {p q} p q hrpq. + move/eqP: hrpq; rewrite mods_eq0 => /eqP ->. + by rewrite changes_itv_mods0p cindexpC. +have p_neq0 : p != 0 by rewrite -(mods_eq0 p q) hrpq. +move: hrpq IHs; rewrite neq0_mods_rec //. +move=> [_ <-] IHs /= /andP[rpa Ha] /andP[rpb Hb]. +move=> /(_ _ _ (erefl _) Ha Hb) in IHs. +have [->|q_neq0] := eqVneq q 0; first by rewrite changes_itv_modsp0 cindex0p. +move: Ha Hb; rewrite neq0_mods_rec //= => /andP[rqa _] /andP[rqb _]. +rewrite cindex_rec 1?changes_itv_mods_rec; +by rewrite ?rootM ?negb_or ?rpa ?rpb ?rqa ?rqb // IHs. +Qed. + +Definition taq_itv a b p q := changes_itv_mods a b p (p^`() * q). + +Lemma taq_taq_itv a b : a < b -> forall p q, + all (fun p => p.[a] != 0) (mods p (p^`() * q)) -> + all (fun p => p.[b] != 0) (mods p (p^`() * q)) -> + taq (roots p a b) q = taq_itv a b p q. +Proof. by move=> *; rewrite /taq_itv changes_itv_mods_cindex // taq_cindex. Qed. + +End ChangesItvMod_USELESS. + +Definition tvecR p sq := let sg_tab := sg_tab (size sq) in + \row_(i < 3^size sq) (taqR p (map (poly_comb sq) sg_tab)`_i). + +Lemma tvec_tvecR sq p : tvec (rootsR p) sq = tvecR p sq. +Proof. +by rewrite /tvec /tvecR; apply/matrixP=> i j; rewrite !mxE taq_taqR. +Qed. + +Lemma all_prodn_gt0 : forall (I : finType) r (P : pred I) (F : I -> nat), + (\prod_(i <- r | P i) F i > 0)%N -> + forall i, i \in r -> P i -> (F i > 0)%N. +Proof. +move=> I r P F; elim: r => [_|a r hr] //. +rewrite big_cons; case hPa: (P a). + rewrite muln_gt0; case/andP=> Fa0; move/hr=> hF x. + by rewrite in_cons; case/orP; [move/eqP-> | move/hF]. +move/hr=> hF x; rewrite in_cons; case/orP; last by move/hF. +by move/eqP->; rewrite hPa. +Qed. + +Definition taqsR p sq i : R := + (taqR p (map (poly_comb sq) (sg_tab (size sq)))`_i)%:~R. + +Definition ccount_weak p sq : R := + let fix aux s (i : nat) := if i is i'.+1 + then aux (taqsR p sq i' * coefs R (size sq) i' + s) i' + else s in aux 0 (3 ^ size sq)%N. + +Lemma constraints1P (p : {poly R}) (sq : seq {poly R}) : + (constraints (rootsR p) (sq) (nseq (size (sq)) 1))%:~R + = ccount_weak p sq. +Proof. +rewrite constraints1_tvec; symmetry. +rewrite castmxE mxE /= /ccount_weak. +transitivity (\sum_(0 <= i < 3 ^ size sq) taqsR p sq i * coefs R (size sq) i). + rewrite unlock /reducebig /= -foldr_map /= /index_iota subn0 foldr_map. + elim: (3 ^ size sq)%N 0%R => [|n ihn] u //. + by rewrite -[X in iota _ X]addn1 iota_add add0n /= foldr_cat ihn. +rewrite big_mkord; apply: congr_big=> // i _. +rewrite /taqsR /coefs /tvecR /=. +have o : 'I_(3 ^ size sq) by rewrite exp3n; apply: ord0. +rewrite (@nth_map _ o); last by rewrite size_enum_ord. +by rewrite !castmxE !cast_ord_id !mxE /= nth_ord_enum taq_taqR. +Qed. + +Lemma ccount_weakP p sq : p != 0 -> + reflect (exists x, (p.[x] == 0) && \big[andb/true]_(q <- sq) (q.[x] > 0)) + (ccount_weak p sq > 0). +Proof. +move=> p_neq0; rewrite -constraints1P /constraints ltr0n lt0n. +rewrite -(@pnatr_eq0 [numDomainType of int]) natr_sum psumr_eq0 //. +rewrite -has_predC /=. +apply: (iffP hasP) => [[x rpx /= prod_neq0]|[x /andP[rpx]]]. + exists x; rewrite -rootE [root _ _]roots_on_rootsR // rpx /=. + rewrite big_seq big1 => // q Hq. + move: prod_neq0; rewrite pnatr_eq0 -lt0n => /all_prodn_gt0. + have := index_mem q sq; rewrite Hq => Hoq. + pose oq := Ordinal Hoq => /(_ oq). + rewrite mem_index_enum => /(_ isT isT) /=. + by rewrite nth_nseq index_mem Hq nth_index // lt0b sgz_cp0. +rewrite big_all => /allP Hsq. +exists x => /=; first by rewrite -roots_on_rootsR. +rewrite pnatr_eq0 -lt0n prodn_gt0 => // i; rewrite nth_nseq ltn_ord lt0b. +by rewrite sgz_cp0 Hsq // mem_nth. +Qed. + +Lemma myprodf_eq0 (S : idomainType)(I : eqType) (r : seq I) P (F : I -> S) : + reflect (exists2 i, ((i \in r) && (P i)) & (F i == 0)) + (\prod_(i <- r| P i) F i == 0). +Proof. +apply: (iffP idP) => [|[i Pi /eqP Fi0]]; last first. + by case/andP: Pi => ri Pi; rewrite (big_rem _ ri) /= Pi Fi0 mul0r. +elim: r => [|i r IHr]; first by rewrite big_nil oner_eq0. +rewrite big_cons /=; have [Pi | ?] := ifP. + rewrite mulf_eq0; case/orP=> [Fi0|]; first by exists i => //; rewrite mem_head. + by case/IHr=> j /andP [rj Pj] Fj; exists j; rewrite // in_cons rj orbT. +by case/IHr=> j /andP [rj Pj] Fj; exists j; rewrite // in_cons rj orbT. +Qed. + +Definition bounding_poly (sq : seq {poly R}) := (\prod_(q <- sq) q)^`(). + +Lemma bounding_polyP (sq : seq {poly R}) : + [\/ \big[andb/true]_(q <- sq) (lead_coef q > 0), + \big[andb/true]_(q <- sq) ((-1)^+(size q).-1 * (lead_coef q) > 0) | + exists x, + ((bounding_poly sq).[x] == 0) && \big[andb/true]_(q <- sq) (q.[x] > 0)] + <-> exists x, \big[andb/true]_(q <- sq) (q.[x] > 0). +Proof. +split=> [|[x]]. + case; last by move=> [x /andP [? h]]; exists x; rewrite h. + rewrite big_all => /allP hsq. + have sqn0 : {in sq, forall q, q != 0}. + by move=> q' /= /hsq; apply: contraL=> /eqP->; rewrite lead_coef0 ltrr. + pose qq := \prod_(q <- sq) q. + have pn0 : qq != 0. + by apply/negP=> /myprodf_eq0 [] q; rewrite andbT => /sqn0 /negPf ->. + pose b := cauchy_bound qq; exists b. + rewrite big_all; apply/allP=> r hr; have:= hsq r hr. + rewrite -!sgr_cp0=> /eqP <-; apply/eqP. + apply: (@sgp_pinftyP _ b); last by rewrite boundl_in_itv. + move=> z Hz /=; have: ~~ root qq z by rewrite ge_cauchy_bound. + by rewrite /root /qq horner_prod prodf_seq_neq0 => /allP /(_ _ hr). + rewrite big_all => /allP hsq. + have sqn0 : {in sq, forall q, q != 0}. + move=> q' /= /hsq; apply: contraL=> /eqP->. + by rewrite lead_coef0 mulr0 ltrr. + pose qq := \prod_(q <- sq) q. + have pn0 : qq != 0. + by apply/negP=> /myprodf_eq0 [] q; rewrite andbT => /sqn0 /negPf ->. + pose b := - cauchy_bound qq; exists b. + rewrite big_all; apply/allP=> r hr; have:= hsq r hr. + rewrite -!sgr_cp0=> /eqP <-; apply/eqP. + apply: (@sgp_minftyP _ b); last by rewrite boundr_in_itv. + move=> z Hz /=; have: ~~ root qq z by rewrite le_cauchy_bound. + by rewrite /root /qq horner_prod prodf_seq_neq0 => /allP /(_ _ hr). +rewrite /bounding_poly; set q := \prod_(q <- _) _. +rewrite big_all => /allP hsq; set bnd := cauchy_bound q. +have sqn0 : {in sq, forall q, q != 0}. + by move=> q' /= /hsq; apply: contraL=> /eqP->; rewrite horner0 ltrr. +have [/eqP|q_neq0] := eqVneq q 0. + by rewrite prodf_seq_eq0=> /hasP [q' /= /sqn0 /negPf->]. +have genroot y : {in sq, forall r, ~~ root q y -> ~~ root r y}. + rewrite /root /q => r r_sq. + by rewrite horner_prod prodf_seq_neq0 => /allP /(_ _ r_sq). +case: (next_rootP q x bnd) q_neq0; [by move->; rewrite eqxx| |]; last first. + move=> _ q_neq0 _ Hq _. + suff -> : \big[andb/true]_(q1 <- sq) (0 < lead_coef q1) by constructor. + rewrite big_all; apply/allP=> r hr; have rxp := hsq r hr. + rewrite -sgr_cp0 -/(sgp_pinfty _). + rewrite -(@sgp_pinftyP _ x _ _ x) ?boundl_in_itv ?sgr_cp0 //. + move=> z; rewrite (@itv_splitU _ x true) /= ?boundl_in_itv //. + rewrite itv_xx /= inE => /orP [/eqP->|]; first by rewrite /root gtr_eqF. + have [x_b|b_x] := ltrP x bnd. + rewrite (@itv_splitU _ bnd false) /=; last by rewrite inE x_b. + move=> /orP [] Hz; rewrite genroot //; + by [rewrite Hq|rewrite ge_cauchy_bound]. + by move=> Hz; rewrite genroot // ge_cauchy_bound // (subitvP _ Hz) //= b_x. +move=> y1 _ rqy1 hy1xb hy1. +case: (prev_rootP q (- bnd) x); [by move->; rewrite eqxx| |]; last first. + move=> _ q_neq0 _ Hq _. (* assia : what is the use of c ? *) + suff -> : \big[andb/true]_(q1 <- sq) (0 < (-1) ^+ (size q1).-1 * lead_coef q1). + by constructor 2. + rewrite big_all; apply/allP=> r hr; have rxp := hsq r hr. + rewrite -sgr_cp0 -/(sgp_minfty _). + rewrite -(@sgp_minftyP _ x _ _ x) ?boundr_in_itv ?sgr_cp0 //. + move=> z; rewrite (@itv_splitU _ x false) /= ?boundr_in_itv //. + rewrite itv_xx => /orP [/=|/eqP->]; last by rewrite /root gtr_eqF. + have [b_x|x_b] := ltrP (- bnd) x. + rewrite (@itv_splitU _ (- bnd) true) /=; last by rewrite inE b_x. + move=> /orP [] Hz; rewrite genroot //; + by [rewrite Hq|rewrite le_cauchy_bound]. + by move=> Hz; rewrite genroot // le_cauchy_bound // (subitvP _ Hz) //= x_b. +move=> y2 _ rqy2 hy2xb hy2 q_neq0. +have lty12 : y2 < y1. + by apply: (@ltr_trans _ x); rewrite 1?(itvP hy1xb) 1?(itvP hy2xb). +have : q.[y2] = q.[y1] by rewrite rqy1 rqy2. +case/(rolle lty12) => z hz rz; constructor 3; exists z. +rewrite rz eqxx /= big_all; apply/allP => r r_sq. +have xy : x \in `]y2, y1[ by rewrite inE 1?(itvP hy1xb) 1?(itvP hy2xb). +rewrite -sgr_cp0 (@polyrN0_itv _ `]y2, y1[ _ _ x) ?sgr_cp0 ?hsq // => t. +rewrite (@itv_splitU2 _ x) // => /or3P [/hy2|/eqP->|/hy1]; do ?exact: genroot. +by rewrite rootE gtr_eqF ?hsq. +Qed. + +Lemma size_prod_eq1 (sq : seq {poly R}) : + reflect (forall q, q \in sq -> size q = 1%N) (size (\prod_(q0 <- sq) q0) == 1%N). +Proof. +apply: (iffP idP). + elim: sq => [| q sq ih]; first by move=> _ q; rewrite in_nil. + rewrite big_cons; case: (altP (q =P 0)) => [-> | qn0]. + by rewrite mul0r size_poly0. + case: (altP ((\prod_(j <- sq) j) =P 0)) => [-> | pn0]. + by rewrite mulr0 size_poly0. + rewrite size_mul //; case: (ltngtP (size q) 1). + - by rewrite ltnS leqn0 size_poly_eq0 (negPf qn0). + - case: (size q) => [|n] //; case: n => [|n] // _; rewrite !addSn /= eqSS. + by rewrite addn_eq0 size_poly_eq0 (negPf pn0) andbF. + - move=> sq1; case: (ltngtP (size (\prod_(j <- sq) j)) 1). + + by rewrite ltnS leqn0 size_poly_eq0 (negPf pn0). + + case: (size (\prod_(j <- sq) j)) => [|n] //; case: n => [|n] // _. + by rewrite !addnS /= eqSS addn_eq0 size_poly_eq0 (negPf qn0). + move=> sp1 _ p; rewrite in_cons; case/orP => [/eqP -> |] //; apply: ih. + by apply/eqP. +elim: sq => [| q sq ih] hs; first by rewrite big_nil size_poly1 eqxx. +case: (altP (q =P 0)) => [ | qn0]. + by move/eqP; rewrite -size_poly_eq0 hs ?mem_head. +case: (altP ((\prod_(q0 <- sq) q0) =P 0)) => [ | pn0]. + move/eqP; rewrite -size_poly_eq0 (eqP (ih _)) // => t ht; apply: hs. + by rewrite in_cons ht orbT. +rewrite big_cons size_mul // (eqP (ih _)) //; last first. + by move=> t ht; apply: hs; rewrite in_cons ht orbT. +rewrite addnS addn0; apply/eqP; apply: hs; exact: mem_head. +Qed. + +Definition ccount_gt0 (sp sq : seq {poly R}) := + let p := \big[@rgcdp _/0%R]_(p <- sp) p in + if p != 0 then 0 < ccount_weak p sq + else let bq := bounding_poly sq in + [|| \big[andb/true]_(q <- sq)(lead_coef q > 0) , + \big[andb/true]_(q <- sq)((-1)^+(size q).-1 *(lead_coef q) > 0) | + 0 < ccount_weak bq sq]. + +Lemma ccount_gt0P (sp sq : seq {poly R}) : + reflect (exists x, \big[andb/true]_(p <- sp) (p.[x] == 0) + && \big[andb/true]_(q <- sq) (q.[x] > 0)) + (ccount_gt0 sp sq). +Proof. +rewrite /ccount_gt0; case: (boolP (_ == 0))=> hsp /=; last first. + apply: (iffP (ccount_weakP _ _)) => // [] [x Hx]; exists x; + by move: Hx; rewrite -rootE root_bigrgcd -big_all. +apply: (@equivP (exists x, \big[andb/true]_(q <- sq) (0 < q.[x]))); last first. + split=> [] [x Hx]; exists x; rewrite ?Hx ?andbT; do ?by case/andP: Hx. + move: hsp; rewrite (big_morph _ (@rgcdp_eq0 _) (eqxx _)) !big_all. + by move=> /allP Hsp; apply/allP => p /Hsp /eqP ->; rewrite horner0. +have [|bq_neq0] := boolP (bounding_poly sq == 0). + rewrite /bounding_poly -derivn1 -derivn_poly0 => ssq_le1. + rewrite -constraints1P (size1_polyC ssq_le1) derivnC /= rootsRC. + rewrite /constraints big_nil ltrr orbF. + move: ssq_le1; rewrite leq_eqVlt ltnS leqn0 orbC. + have [|_ /=] := boolP (_ == _). + rewrite size_poly_eq0 => /eqP sq_eq0; move/eqP: (sq_eq0). + rewrite prodf_seq_eq0 => /hasP /sig2W [q /= q_sq] /eqP q_eq0. + move: q_sq; rewrite q_eq0 => sq0 _ {q q_eq0}. + set f := _ || _; suff -> : f = false; move: @f => /=. + constructor => [] [x]; rewrite big_all. + by move=> /allP /(_ _ sq0); rewrite horner0 ltrr. + apply: negbTE; rewrite !negb_or !big_all -!has_predC. + apply/andP; split; apply/hasP; + by exists 0; rewrite //= ?lead_coef0 ?mulr0 ltrr. + move=> /size_prod_eq1 Hsq. + have {Hsq} Hsq q : q \in sq -> q = (lead_coef q)%:P. + by move=> /Hsq sq1; rewrite [q]size1_polyC ?sq1 // lead_coefC. + apply: (@equivP (\big[andb/true]_(q <- sq) (0 < lead_coef q))); last first. + split; [move=> sq0; exists 0; move: sq0|move=> [x]]; + rewrite !big_all => /allP H; apply/allP => q q_sq; have:= H _ q_sq; + by rewrite [q]Hsq ?lead_coefC ?hornerC. + have [] := boolP (\big[andb/true]_(q <- _) (0 < lead_coef q)). + by constructor. + rewrite !big_all -has_predC => /hasP sq0; apply: (iffP allP) => //=. + move: sq0 => [q q_sq /= lq_gt0 /(_ _ q_sq)]. + rewrite [q]Hsq ?size_polyC ?lead_coefC //. + by case: (_ != 0); rewrite /= expr0 mul1r ?(negPf lq_gt0). +apply: (iffP or3P); rewrite -bounding_polyP; +case; do ?by [constructor 1|constructor 2]; +by move/(ccount_weakP _ bq_neq0); constructor 3. +Qed. + +End QeRcfTh. diff --git a/mathcomp/real_closed/realalg.v b/mathcomp/real_closed/realalg.v new file mode 100644 index 0000000..88d656a --- /dev/null +++ b/mathcomp/real_closed/realalg.v @@ -0,0 +1,1530 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import bigop ssralg ssrnum ssrint rat poly polydiv polyorder. +Require Import perm matrix mxpoly polyXY binomial generic_quotient. +Require Import cauchyreals separable zmodp bigenough. + +(*************************************************************************) +(* This files constructs the real closure of an archimedian field in the *) +(* way described in Cyril Cohen. Construction of real algebraic numbers *) +(* in Coq. In Lennart Beringer and Amy Felty, editors, ITP - 3rd *) +(* International Conference on Interactive Theorem Proving - 2012, *) +(* Princeton, United States, August 2012. Springer *) +(* *) +(* The only definition one may want to use in this file is the operator *) +(* {realclosure R} which constructs the real closure of the archimedian *) +(* field R (for which rat is a prefect candidate) *) +(*************************************************************************) + +Import GRing.Theory Num.Theory BigEnough. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Reserved Notation "{ 'realclosure' T }" + (at level 0, format "{ 'realclosure' T }"). +Reserved Notation "{ 'alg' T }" (at level 0, format "{ 'alg' T }"). + +Section extras. + +Local Open Scope ring_scope. +Local Notation "p ^ f" := (map_poly f p) : ring_scope. + +Lemma map_comp_poly (aR : fieldType) (rR : idomainType) + (f : {rmorphism aR -> rR}) + (p q : {poly aR}) : (p \Po q) ^ f = (p ^ f) \Po (q ^ f). +Proof. +rewrite !comp_polyE size_map_poly; apply: (big_ind2 (fun x y => x ^ f = y)). ++ by rewrite rmorph0. ++ by move=> u u' v v' /=; rewrite rmorphD /= => -> ->. +move=> /= i _; rewrite -mul_polyC rmorphM /= map_polyC mul_polyC. +by rewrite coef_map rmorphX. +Qed. + +End extras. + +Module RealAlg. + +Local Open Scope ring_scope. +Local Notation eval := horner_eval. + +Section RealAlg. + +Variable F : archiFieldType. +Local Notation m0 := (fun _ => 0%N). + +(*********************************************************************) +(* Construction of algebraic Cauchy reals : Cauchy real + polynomial *) +(*********************************************************************) + +CoInductive algcreal := AlgCReal { + creal_of_alg :> creal F; + annul_creal : {poly F}; + _ : annul_creal \is monic; + _ : (annul_creal.[creal_of_alg] == 0)%CR +}. + +Lemma monic_annul_creal x : annul_creal x \is monic. +Proof. by case: x. Qed. +Hint Resolve monic_annul_creal. + +Lemma annul_creal_eq0 x : (annul_creal x == 0) = false. +Proof. by rewrite (negPf (monic_neq0 _)). Qed. + +Lemma root_annul_creal x : ((annul_creal x).[x] == 0)%CR. +Proof. by case: x. Qed. +Hint Resolve root_annul_creal. + +Definition cst_algcreal (x : F) := + AlgCReal (monicXsubC _) (@root_cst_creal _ x). + +Local Notation zero_algcreal := (cst_algcreal 0). +Local Notation one_algcreal := (cst_algcreal 1). + +Lemma size_annul_creal_gt1 (x : algcreal) : + (1 < size (annul_creal x))%N. +Proof. +apply: (@has_root_creal_size_gt1 _ x). + by rewrite monic_neq0 // monic_annul_creal. +exact: root_annul_creal. +Qed. + +Lemma is_root_annul_creal (x : algcreal) (y : creal F) : + (x == y)%CR -> ((annul_creal x).[y] == 0)%CR. +Proof. by move <-. Qed. + +Definition AlgCRealOf (p : {poly F}) (x : creal F) + (p_neq0 : p != 0) (px_eq0 : (p.[x] == 0)%CR) := + AlgCReal (monic_monic_from_neq0 p_neq0) (root_monic_from_neq0 px_eq0). + +Lemma sub_annihilant_algcreal_neq0 (x y : algcreal) : + sub_annihilant (annul_creal x) (annul_creal y) != 0. +Proof. by rewrite sub_annihilant_neq0 ?monic_neq0. Qed. + +Lemma root_sub_algcreal (x y : algcreal) : + ((sub_annihilant (annul_creal x) (annul_creal y)).[x - y] == 0)%CR. +Proof. by rewrite root_sub_annihilant_creal ?root_annul_creal ?monic_neq0. Qed. + +Definition sub_algcreal (x y : algcreal) : algcreal := + AlgCRealOf (sub_annihilant_algcreal_neq0 x y) (@root_sub_algcreal x y). + +Lemma root_opp_algcreal (x : algcreal) : + ((annul_creal (sub_algcreal (cst_algcreal 0) x)).[- x] == 0)%CR. +Proof. by apply: is_root_annul_creal; rewrite /= add_0creal. Qed. + +Definition opp_algcreal (x : algcreal) : algcreal := + AlgCReal (@monic_annul_creal _) (@root_opp_algcreal x). + +Lemma root_add_algcreal (x y : algcreal) : + ((annul_creal (sub_algcreal x (opp_algcreal y))).[x + y] == 0)%CR. +Proof. +apply: is_root_annul_creal; apply: eq_crealP. +by exists m0=> * /=; rewrite opprK subrr normr0. +Qed. + +Definition add_algcreal (x y : algcreal) : algcreal := + AlgCReal (@monic_annul_creal _) (@root_add_algcreal x y). + +Lemma div_annihilant_algcreal_neq0 (x y : algcreal) : + (annul_creal y).[0] != 0 -> + div_annihilant (annul_creal x) (annul_creal y) != 0. +Proof. by move=> ?; rewrite div_annihilant_neq0 ?monic_neq0. Qed. + +Hint Resolve eq_creal_refl. +Hint Resolve le_creal_refl. + +Lemma simplify_algcreal (x : algcreal) (x_neq0 : (x != 0)%CR) : + {y | ((annul_creal y).[0] != 0) & ((y != 0)%CR * (x == y)%CR)%type}. +Proof. +elim: size {-3}x x_neq0 (leqnn (size (annul_creal x))) => + {x} [|n ihn] x x_neq0 hx. + by move: hx; rewrite leqn0 size_poly_eq0 annul_creal_eq0. +have [dvdX|ndvdX] := boolP ('X %| annul_creal x); last first. + by exists x=> //; rewrite -rootE -dvdp_XsubCl subr0. +have monic_p: @annul_creal x %/ 'X \is monic. + by rewrite -(monicMr _ (@monicX _)) divpK //. +have root_p: ((@annul_creal x %/ 'X).[x] == 0)%CR. + have := @eq_creal_refl _ ((annul_creal x).[x])%CR. + rewrite -{1}(divpK dvdX) horner_crealM // root_annul_creal. + by case/poly_mul_creal_eq0=> //; rewrite horner_crealX. +have [//|/=|y *] := ihn (AlgCReal monic_p root_p); last by exists y. +by rewrite size_divp ?size_polyX ?polyX_eq0 ?leq_subLR ?add1n. +Qed. + +(* Decidability of equality to 0 *) +Lemma algcreal_eq0_dec (x : algcreal) : {(x == 0)%CR} + {(x != 0)%CR}. +Proof. +pose p := annul_creal x; move: {2}(size _)%N (leqnn (size p))=> n. +elim: n x @p => [x p|n ihn x p le_sp_Sn]. + by rewrite leqn0 size_poly_eq0 /p annul_creal_eq0. +move: le_sp_Sn; rewrite leq_eqVlt; have [|//|eq_sp_Sn _] := ltngtP. + by rewrite ltnS=> /ihn ihnp _; apply: ihnp. +have px0 : (p.[x] == 0)%CR by apply: root_annul_creal. +have [cpX|ncpX] := boolP (coprimep p 'X). + by right; move: (cpX)=> /coprimep_root /(_ px0); rewrite horner_crealX. +have [eq_pX|] := altP (p =P 'X). + by left; move: px0; rewrite eq_pX horner_crealX. +rewrite -eqp_monic /p ?monicX // negb_and orbC. +have:= ncpX; rewrite coprimepX -dvdp_XsubCl subr0 => /negPf-> /= ndiv_pX. +have [r] := smaller_factor (monic_annul_creal _) px0 ndiv_pX ncpX. +rewrite eq_sp_Sn ltnS => /andP[le_r_n monic_r] rx_eq0. +exact: (ihn (AlgCReal monic_r rx_eq0)). +Qed. + +Lemma eq_algcreal_dec (x y : algcreal) : {(x == y)%CR} + {(x != y)%CR}. +Proof. +have /= [d_eq0|d_neq0] := algcreal_eq0_dec (sub_algcreal x y); [left|right]. + apply: eq_crealP; exists_big_modulus m F. + by move=> e i e_gt0 hi; rewrite (@eq0_modP _ _ d_eq0). + by close. +pose_big_enough i. + apply: (@neq_crealP _ (lbound d_neq0) i i); do ?by rewrite ?lbound_gt0. + by rewrite (@lbound0P _ _ d_neq0). +by close. +Qed. + +Definition eq_algcreal : rel algcreal := eq_algcreal_dec. + +Lemma eq_algcrealP (x y : algcreal) : reflect (x == y)%CR (eq_algcreal x y). +Proof. by rewrite /eq_algcreal; case: eq_algcreal_dec=> /=; constructor. Qed. +Implicit Arguments eq_algcrealP [x y]. + +Lemma neq_algcrealP (x y : algcreal) : reflect (x != y)%CR (~~ eq_algcreal x y). +Proof. by rewrite /eq_algcreal; case: eq_algcreal_dec=> /=; constructor. Qed. +Implicit Arguments neq_algcrealP [x y]. +Prenex Implicits eq_algcrealP neq_algcrealP. + +Fact eq_algcreal_is_equiv : equiv_class_of eq_algcreal. +Proof. +split=> [x|x y|y x z]; first by apply/eq_algcrealP. + by apply/eq_algcrealP/eq_algcrealP; symmetry. +by move=> /eq_algcrealP /eq_creal_trans h /eq_algcrealP /h /eq_algcrealP. +Qed. + +Canonical eq_algcreal_rel := EquivRelPack eq_algcreal_is_equiv. + +Lemma root_div_algcreal (x y : algcreal) (y_neq0 : (y != 0)%CR) : + (annul_creal y).[0] != 0 -> + ((div_annihilant (annul_creal x) (annul_creal y)).[x / y_neq0] == 0)%CR. +Proof. by move=> hx; rewrite root_div_annihilant_creal ?monic_neq0. Qed. + +Definition div_algcreal (x y : algcreal) := + match eq_algcreal_dec y (cst_algcreal 0) with + | left y_eq0 => cst_algcreal 0 + | right y_neq0 => + let: exist2 y' py'0_neq0 (y'_neq0, _) := simplify_algcreal y_neq0 in + AlgCRealOf (div_annihilant_algcreal_neq0 x py'0_neq0) + (@root_div_algcreal x y' y'_neq0 py'0_neq0) + end. + +Lemma root_inv_algcreal (x : algcreal) (x_neq0 : (x != 0)%CR) : + ((annul_creal (div_algcreal (cst_algcreal 1) x)).[x_neq0^-1] == 0)%CR. +Proof. +rewrite /div_algcreal; case: eq_algcreal_dec=> [/(_ x_neq0)|x_neq0'] //=. +case: simplify_algcreal=> x' px'0_neq0 [x'_neq0 eq_xx']. +apply: is_root_annul_creal;rewrite /= -(@eq_creal_inv _ _ _ x_neq0) //. +by apply: eq_crealP; exists m0=> * /=; rewrite div1r subrr normr0. +Qed. + +Definition inv_algcreal (x : algcreal) := + match eq_algcreal_dec x (cst_algcreal 0) with + | left x_eq0 => cst_algcreal 0 + | right x_neq0 => + AlgCReal (@monic_annul_creal _) (@root_inv_algcreal _ x_neq0) + end. + +Lemma div_creal_creal (y : creal F) (y_neq0 : (y != 0)%CR) : + (y / y_neq0 == 1%:CR)%CR. +Proof. +apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi; rewrite /= divff ?subrr ?normr0 //. + by rewrite (@creal_neq_always _ _ 0%CR). +by close. +Qed. + +Lemma root_mul_algcreal (x y : algcreal) : + ((annul_creal (div_algcreal x (inv_algcreal y))).[x * y] == 0)%CR. +Proof. +rewrite /div_algcreal /inv_algcreal. +case: (eq_algcreal_dec y)=> [->|y_neq0]; apply: is_root_annul_creal. + rewrite mul_creal0; case: eq_algcreal_dec=> // neq_00. + by move: (eq_creal_refl neq_00). +case: eq_algcreal_dec=> /= [yV_eq0|yV_neq0]. + have: (y * y_neq0^-1 == 0)%CR by rewrite yV_eq0 mul_creal0. + by rewrite div_creal_creal=> /eq_creal_cst; rewrite oner_eq0. +case: simplify_algcreal=> y' py'0_neq0 [y'_neq0 /= eq_yy']. +rewrite -(@eq_creal_inv _ _ _ yV_neq0) //. +by apply: eq_crealP; exists m0=> * /=; rewrite invrK subrr normr0. +Qed. + +Definition mul_algcreal (x y : algcreal) := + AlgCReal (@monic_annul_creal _) (@root_mul_algcreal x y). + +Lemma le_creal_neqVlt (x y : algcreal) : (x <= y)%CR -> {(x == y)%CR} + {(x < y)%CR}. +Proof. +case: (eq_algcreal_dec x y); first by left. +by move=> /neq_creal_ltVgt [|h /(_ h) //]; right. +Qed. + +Lemma ltVge_algcreal_dec (x y : algcreal) : {(x < y)%CR} + {(y <= x)%CR}. +Proof. +have [eq_xy|/neq_creal_ltVgt [lt_xy|lt_yx]] := eq_algcreal_dec x y; +by [right; rewrite eq_xy | left | right; apply: lt_crealW]. +Qed. + +Definition lt_algcreal : rel algcreal := ltVge_algcreal_dec. +Definition le_algcreal : rel algcreal := fun x y => ~~ ltVge_algcreal_dec y x. + +Lemma lt_algcrealP (x y : algcreal) : reflect (x < y)%CR (lt_algcreal x y). +Proof. by rewrite /lt_algcreal; case: ltVge_algcreal_dec; constructor. Qed. +Implicit Arguments lt_algcrealP [x y]. + +Lemma le_algcrealP (x y : algcreal) : reflect (x <= y)%CR (le_algcreal x y). +Proof. by rewrite /le_algcreal; case: ltVge_algcreal_dec; constructor. Qed. +Implicit Arguments le_algcrealP [x y]. +Prenex Implicits lt_algcrealP le_algcrealP. + +Definition exp_algcreal x n := iterop n mul_algcreal x one_algcreal. + +Lemma exp_algcrealE x n : (exp_algcreal x n == x ^+ n)%CR. +Proof. +case: n=> // n; rewrite /exp_algcreal /exp_creal !iteropS. +by elim: n=> //= n ->. +Qed. + +Definition horner_algcreal (p : {poly F}) x : algcreal := + \big[add_algcreal/zero_algcreal]_(i < size p) + mul_algcreal (cst_algcreal p`_i) (exp_algcreal x i). + +Lemma horner_algcrealE p x : (horner_algcreal p x == p.[x])%CR. +Proof. +rewrite horner_coef_creal. +apply: (big_ind2 (fun (u : algcreal) v => u == v)%CR)=> //. + by move=> u u' v v' /= -> ->. +by move=> i _ /=; rewrite exp_algcrealE. +Qed. + +Definition norm_algcreal (x : algcreal) := + if le_algcreal zero_algcreal x then x else opp_algcreal x. + +Lemma norm_algcrealE (x : algcreal) : (norm_algcreal x == `| x |)%CR. +Proof. +rewrite /norm_algcreal /le_algcreal; case: ltVge_algcreal_dec => /=. + move=> x_lt0; apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi /=; rewrite [`|x i|]ler0_norm ?subrr ?normr0 //. + by rewrite ltrW // [_ < 0%CR i]creal_lt_always. + by close. +move=> /(@le_creal_neqVlt zero_algcreal) /= []. + by move<-; apply: eq_crealP; exists m0=> * /=; rewrite !(normr0, subrr). +move=> x_gt0; apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi /=; rewrite [`|x i|]ger0_norm ?subrr ?normr0 //. + by rewrite ltrW // creal_gt0_always. +by close. +Qed. + +(**********************************************************************) +(* Theory of the "domain" of algebraic numbers: polynomial + interval *) +(**********************************************************************) +CoInductive algdom := AlgDom { + annul_algdom : {poly F}; + center_alg : F; + radius_alg : F; + _ : annul_algdom \is monic; + _ : radius_alg >= 0; + _ : annul_algdom.[center_alg - radius_alg] + * annul_algdom.[center_alg + radius_alg] <= 0 +}. + +Lemma radius_alg_ge0 x : 0 <= radius_alg x. Proof. by case: x. Qed. + +Lemma monic_annul_algdom x : annul_algdom x \is monic. Proof. by case: x. Qed. +Hint Resolve monic_annul_algdom. + +Lemma annul_algdom_eq0 x : (annul_algdom x == 0) = false. +Proof. by rewrite (negPf (monic_neq0 _)). Qed. + +Lemma algdomP x : (annul_algdom x).[center_alg x - radius_alg x] + * (annul_algdom x).[center_alg x + radius_alg x] <= 0. +Proof. by case: x. Qed. + +Definition algdom' := seq F. + +Definition encode_algdom (x : algdom) : algdom' := + [:: center_alg x, radius_alg x & (annul_algdom x)]. + +Definition decode_algdom (x : algdom') : option algdom := + if x is [::c, r & p'] + then let p := Poly p' in + if ((p \is monic) =P true, (r >= 0) =P true, + (p.[c - r] * p.[c + r] <= 0) =P true) + is (ReflectT monic_p, ReflectT r_gt0, ReflectT hp) + then Some (AlgDom monic_p r_gt0 hp) + else None + else None. + +Lemma encode_algdomK : pcancel encode_algdom decode_algdom. +Proof. +case=> p c r monic_p r_ge0 hp /=; rewrite polyseqK. +do 3?[case: eqP; rewrite ?monic_p ?r_ge0 ?monic_p //] => monic_p' r_ge0' hp'. +by congr (Some (AlgDom _ _ _)); apply: bool_irrelevance. +Qed. + +Definition algdom_EqMixin := PcanEqMixin encode_algdomK. +Canonical algdom_eqType := EqType algdom algdom_EqMixin. + +Definition algdom_ChoiceMixin := PcanChoiceMixin encode_algdomK. +Canonical algdom_choiceType := ChoiceType algdom algdom_ChoiceMixin. + +Fixpoint to_algcreal_of (p : {poly F}) (c r : F) (i : nat) : F := + match i with + | 0 => c + | i.+1 => + let c' := to_algcreal_of p c r i in + if p.[c' - r / 2%:R ^+ i] * p.[c'] <= 0 + then c' - r / 2%:R ^+ i.+1 + else c' + r / 2%:R ^+ i.+1 + end. + + +Lemma to_algcreal_of_recP p c r i : 0 <= r -> + `|to_algcreal_of p c r i.+1 - to_algcreal_of p c r i| <= r * 2%:R ^- i.+1. +Proof. +move=> r_ge0 /=; case: ifP=> _; rewrite addrAC subrr add0r ?normrN ger0_norm //; +by rewrite mulr_ge0 ?invr_ge0 ?exprn_ge0 ?ler0n. +Qed. + +Lemma to_algcreal_ofP p c r i j : 0 <= r -> (i <= j)%N -> + `|to_algcreal_of p c r j - to_algcreal_of p c r i| <= r * 2%:R ^- i. +Proof. +move=> r_ge0 leij; pose r' := r * 2%:R ^- j * (2%:R ^+ (j - i) - 1). +rewrite (@ler_trans _ r') //; last first. + rewrite /r' -mulrA ler_wpmul2l // ler_pdivr_mull ?exprn_gt0 ?ltr0n //. + rewrite -{2}(subnK leij) exprD mulfK ?gtr_eqF ?exprn_gt0 ?ltr0n //. + by rewrite ger_addl lerN10. +rewrite /r' subrX1 addrK mul1r -{1 2}(subnK leij); set f := _ c r. +elim: (_ - _)%N=> [|k ihk]; first by rewrite subrr normr0 big_ord0 mulr0 lerr. +rewrite addSn big_ord_recl /= mulrDr. +rewrite (ler_trans (ler_dist_add (f (k + i)%N) _ _)) //. +rewrite ler_add ?expr0 ?mulr1 ?to_algcreal_of_recP // (ler_trans ihk) //. +rewrite exprSr invfM -!mulrA !ler_wpmul2l ?invr_ge0 ?exprn_ge0 ?ler0n //. +by rewrite mulr_sumr ler_sum // => l _ /=; rewrite exprS mulKf ?pnatr_eq0. +Qed. + +Lemma alg_to_crealP (x : algdom) : + creal_axiom (to_algcreal_of (annul_algdom x) (center_alg x) (radius_alg x)). +Proof. +pose_big_modulus m F. + exists m=> e i j e_gt0 hi hj. + wlog leij : i j {hi} hj / (j <= i)%N. + move=> hwlog; case/orP: (leq_total i j)=> /hwlog; last exact. + by rewrite distrC; apply. + rewrite (ler_lt_trans (to_algcreal_ofP _ _ _ _)) ?radius_alg_ge0 //. + rewrite ltr_pdivr_mulr ?gtr0E // -ltr_pdivr_mull //. + by rewrite upper_nthrootP. +by close. +Qed. + +Definition alg_to_creal x := CReal (alg_to_crealP x). + +Lemma exp2k_crealP : @creal_axiom F (fun i => 2%:R ^- i). +Proof. +pose_big_modulus m F. + exists m=> e i j e_gt0 hi hj. + wlog leij : i j {hj} hi / (i <= j)%N. + move=> hwlog; case/orP: (leq_total i j)=> /hwlog; first exact. + by rewrite distrC; apply. + rewrite ger0_norm ?subr_ge0; last first. + by rewrite ?lef_pinv -?topredE /= ?gtr0E // ler_eexpn2l ?ltr1n. + rewrite -(@ltr_pmul2l _ (2%:R ^+ i )) ?gtr0E //. + rewrite mulrBr mulfV ?gtr_eqF ?gtr0E //. + rewrite (@ler_lt_trans _ 1) // ?ger_addl ?oppr_le0 ?mulr_ge0 ?ger0E //. + by rewrite -ltr_pdivr_mulr // mul1r upper_nthrootP. +by close. +Qed. +Definition exp2k_creal := CReal exp2k_crealP. + +Lemma exp2k_creal_eq0 : (exp2k_creal == 0)%CR. +Proof. +apply: eq_crealP; exists_big_modulus m F. + move=> e i e_gt0 hi /=. + rewrite subr0 gtr0_norm ?gtr0E // -ltf_pinv -?topredE /= ?gtr0E //. + by rewrite invrK upper_nthrootP. +by close. +Qed. + +Notation lbound0_of p := (@lbound0P _ _ p _ _ _). + +Lemma to_algcrealP (x : algdom) : ((annul_algdom x).[alg_to_creal x] == 0)%CR. +Proof. +set u := alg_to_creal _; set p := annul_algdom _. +pose r := radius_alg x; pose cr := cst_creal r. +have: ((p).[u - cr * exp2k_creal] * (p).[u + cr * exp2k_creal] <= 0)%CR. + apply: (@le_crealP _ 0%N)=> i _ /=. + rewrite -/p -/r; set c := center_alg _. + elim: i=> /= [|i]. + by rewrite !expr0 divr1 algdomP. + set c' := to_algcreal_of _ _ _=> ihi. + have [] := lerP (_ * p.[c' i]). + rewrite addrNK -addrA -opprD -mulr2n -[_ / _ *+ _]mulr_natr. + by rewrite -mulrA exprSr invfM mulfVK ?pnatr_eq0. + rewrite addrK -addrA -mulr2n -[_ / _ *+ _]mulr_natr. + rewrite -mulrA exprSr invfM mulfVK ?pnatr_eq0 // => /ler_pmul2l<-. + rewrite mulr0 mulrCA !mulrA [X in X * _]mulrAC -mulrA. + by rewrite mulr_ge0_le0 // -expr2 exprn_even_ge0. +rewrite exp2k_creal_eq0 mul_creal0 opp_creal0 add_creal0. +move=> hu pu0; apply: hu; pose e := (lbound pu0). +pose_big_enough i. + apply: (@lt_crealP _ (e * e) i i) => //. + by rewrite !pmulr_rgt0 ?invr_gt0 ?ltr0n ?lbound_gt0. + rewrite add0r [u]lock /= -!expr2. + rewrite -[_.[_] ^+ _]ger0_norm ?exprn_even_ge0 // normrX. + rewrite ler_pexpn2r -?topredE /= ?lbound_ge0 ?normr_ge0 //. + by rewrite -lock (ler_trans _ (lbound0_of pu0)). +by close. +Qed. + +Definition to_algcreal_rec (x : algdom) := + AlgCReal (monic_annul_algdom x) (@to_algcrealP x). +(* "Encoding" function from algdom to algcreal *) +Definition to_algcreal := locked to_algcreal_rec. + +(* "Decoding" function, constructed interactively *) +Lemma to_algdom_exists (x : algcreal) : + { y : algdom | (to_algcreal y == x)%CR }. +Proof. +pose p := annul_creal x. +move: {2}(size p) (leqnn (size p))=> n. +elim: n x @p=> [x p|n ihn x p le_sp_Sn]. + by rewrite leqn0 size_poly_eq0 /p annul_creal_eq0. +move: le_sp_Sn; rewrite leq_eqVlt. +have [|//|eq_sp_Sn _] := ltngtP. + by rewrite ltnS=> /ihn ihnp _; apply: ihnp. +have px0 := @root_annul_creal x; rewrite -/p -/root in px0. +have [|ncop] := boolP (coprimep p p^`()). + move/coprimep_root => /(_ _ px0) /deriv_neq0_mono [r r_gt0 [i ir sm]]. + have p_chg_sign : p.[x i - r] * p.[x i + r] <= 0. + have [/accr_pos_incr hp|/accr_neg_decr hp] := sm. + have hpxj : forall j, (i <= j)%N -> + (p.[x i - r] <= p.[x j]) * (p.[x j] <= p.[x i + r]). + move=> j hj. + suff: p.[x i - r] <= p.[x j] <= p.[x i + r] by case/andP=> -> ->. + rewrite !hp 1?addrAC ?subrr ?add0r ?normrN; + rewrite ?(gtr0_norm r_gt0) //; + do ?by rewrite ltrW ?cauchymodP ?(leq_trans _ hj). + by rewrite -ler_distl ltrW ?cauchymodP ?(leq_trans _ hj). + rewrite mulr_le0_ge0 //; apply/le_creal_cst; rewrite -px0; + by apply: (@le_crealP _ i)=> h hj /=; rewrite hpxj. + have hpxj : forall j, (i <= j)%N -> + (p.[x i + r] <= p.[x j]) * (p.[x j] <= p.[x i - r]). + move=> j hj. + suff: p.[x i + r] <= p.[x j] <= p.[x i - r] by case/andP=> -> ->. + rewrite !hp 1?addrAC ?subrr ?add0r ?normrN; + rewrite ?(gtr0_norm r_gt0) //; + do ?by rewrite ltrW ?cauchymodP ?(leq_trans _ hj). + by rewrite andbC -ler_distl ltrW ?cauchymodP ?(leq_trans _ hj). + rewrite mulr_ge0_le0 //; apply/le_creal_cst; rewrite -px0; + by apply: (@le_crealP _ i)=> h hj /=; rewrite hpxj. + pose y := (AlgDom (monic_annul_creal x) (ltrW r_gt0) p_chg_sign). + have eq_py_px: (p.[to_algcreal y] == p.[x])%CR. + rewrite /to_algcreal -lock. + by have := @to_algcrealP y; rewrite /= -/p=> ->; apply: eq_creal_sym. + exists y. + move: sm=> /strong_mono_bound [k k_gt0 hk]. + rewrite -/p; apply: eq_crealP. + exists_big_modulus m F. + move=> e j e_gt0 hj; rewrite (ler_lt_trans (hk _ _ _ _)) //. + + rewrite /to_algcreal -lock. + rewrite (ler_trans (to_algcreal_ofP _ _ _ (leq0n _))) ?(ltrW r_gt0) //. + by rewrite expr0 divr1. + + by rewrite ltrW // cauchymodP. + rewrite -ltr_pdivl_mull //. + by rewrite (@eq_modP _ _ _ eq_py_px) // ?pmulr_rgt0 ?invr_gt0. + by close. +case: (@smaller_factor _ p p^`() x); rewrite ?monic_annul_creal //. + rewrite gtNdvdp // -?size_poly_eq0 size_deriv eq_sp_Sn //=. + apply: contra ncop=> /eqP n_eq0; move: eq_sp_Sn; rewrite n_eq0. + by move=> /eqP /size_poly1P [c c_neq0 ->]; rewrite derivC coprimep0 polyC_eqp1. +move=> r /andP [hsr monic_r rx_eq0]. +apply: (ihn (AlgCReal monic_r rx_eq0))=> /=. +by rewrite -ltnS -eq_sp_Sn. +Qed. + +Definition to_algdom x := projT1 (to_algdom_exists x). + +Lemma to_algdomK x : (to_algcreal (to_algdom x) == x)%CR. +Proof. by rewrite /to_algdom; case: to_algdom_exists. Qed. + +Lemma eq_algcreal_to_algdom x : eq_algcreal (to_algcreal (to_algdom x)) x. +Proof. by apply/eq_algcrealP; apply: to_algdomK. Qed. + +(* Explicit encoding to a choice type *) +Canonical eq_algcreal_encModRel := EncModRel eq_algcreal eq_algcreal_to_algdom. + +Local Open Scope quotient_scope. + +(***************************************************************************) +(* Algebraic numbers are the quotient of algcreal by their setoid equality *) +(***************************************************************************) +Definition alg := {eq_quot eq_algcreal}. + +Definition alg_of of (phant F) := alg. +Identity Coercion type_alg_of : alg_of >-> alg. + +Notation "{ 'alg' F }" := (alg_of (Phant F)). + +(* A lot of structure is inherited *) +Canonical alg_eqType := [eqType of alg]. +Canonical alg_choiceType := [choiceType of alg]. +Canonical alg_quotType := [quotType of alg]. +Canonical alg_eqQuotType := [eqQuotType eq_algcreal of alg]. +Canonical alg_of_eqType := [eqType of {alg F}]. +Canonical alg_of_choiceType := [choiceType of {alg F}]. +Canonical alg_of_quotType := [quotType of {alg F}]. +Canonical alg_of_eqQuotType := [eqQuotType eq_algcreal of {alg F}]. + +Definition to_alg_def (phF : phant F) : F -> {alg F} := + lift_embed {alg F} cst_algcreal. +Notation to_alg := (@to_alg_def (Phant F)). +Notation "x %:RA" := (to_alg x) + (at level 2, left associativity, format "x %:RA"). +Local Notation "p ^ f" := (map_poly f p) : ring_scope. + +Canonical to_alg_pi_morph := PiEmbed to_alg. + +Local Notation zero_alg := 0%:RA. +Local Notation one_alg := 1%:RA. + +Lemma equiv_alg (x y : algcreal) : (x == y)%CR <-> (x = y %[mod {alg F}]). +Proof. +split; first by move=> /eq_algcrealP /eqquotP ->. +by move=> /eqquotP /eq_algcrealP. +Qed. + +Lemma nequiv_alg (x y : algcreal) : reflect (x != y)%CR (x != y %[mod {alg F}]). +Proof. by rewrite eqquotE; apply: neq_algcrealP. Qed. +Implicit Arguments nequiv_alg [x y]. +Prenex Implicits nequiv_alg. + +Lemma pi_algK (x : algcreal) : (repr (\pi_{alg F} x) == x)%CR. +Proof. by apply/equiv_alg; rewrite reprK. Qed. + +Definition add_alg := lift_op2 {alg F} add_algcreal. + +Lemma pi_add : {morph \pi_{alg F} : x y / add_algcreal x y >-> add_alg x y}. +Proof. by unlock add_alg=> x y; rewrite -equiv_alg /= !pi_algK. Qed. + +Canonical add_pi_morph := PiMorph2 pi_add. + +Definition opp_alg := lift_op1 {alg F} opp_algcreal. + +Lemma pi_opp : {morph \pi_{alg F} : x / opp_algcreal x >-> opp_alg x}. +Proof. by unlock opp_alg=> x; rewrite -equiv_alg /= !pi_algK. Qed. + +Canonical opp_pi_morph := PiMorph1 pi_opp. + +Definition mul_alg := lift_op2 {alg F} mul_algcreal. + +Lemma pi_mul : {morph \pi_{alg F} : x y / mul_algcreal x y >-> mul_alg x y}. +Proof. by unlock mul_alg=> x y; rewrite -equiv_alg /= !pi_algK. Qed. + +Canonical mul_pi_morph := PiMorph2 pi_mul. + +Definition inv_alg := lift_op1 {alg F} inv_algcreal. + +Lemma pi_inv : {morph \pi_{alg F} : x / inv_algcreal x >-> inv_alg x}. +Proof. +unlock inv_alg=> x; symmetry; rewrite -equiv_alg /= /inv_algcreal. +case: eq_algcreal_dec=> /= [|x'_neq0]. + by rewrite pi_algK; case: eq_algcreal_dec. +move: x'_neq0 (x'_neq0); rewrite {1}pi_algK. +case: eq_algcreal_dec=> // x'_neq0' x_neq0 x'_neq0 /=. +by apply: eq_creal_inv; rewrite pi_algK. +Qed. + +Canonical inv_pi_morph := PiMorph1 pi_inv. + +Lemma add_algA : associative add_alg. +Proof. +elim/quotW=> x; elim/quotW=> y; elim/quotW=> z; rewrite !piE -equiv_alg. +by apply: eq_crealP; exists m0=> * /=; rewrite addrA subrr normr0. +Qed. + +Lemma add_algC : commutative add_alg. +Proof. +elim/quotW=> x; elim/quotW=> y; rewrite !piE -equiv_alg /=. +by apply: eq_crealP; exists m0=> * /=; rewrite [X in _ - X]addrC subrr normr0. +Qed. + +Lemma add_0alg : left_id zero_alg add_alg. +Proof. by elim/quotW=> x; rewrite !piE -equiv_alg /= add_0creal. Qed. + +Lemma add_Nalg : left_inverse zero_alg opp_alg add_alg. +Proof. +elim/quotW=> x; rewrite !piE -equiv_alg /=. +by apply: eq_crealP; exists m0=> *; rewrite /= addNr subr0 normr0. +Qed. + +Definition alg_zmodMixin := ZmodMixin add_algA add_algC add_0alg add_Nalg. +Canonical alg_zmodType := Eval hnf in ZmodType alg alg_zmodMixin. +Canonical alg_of_zmodType := Eval hnf in ZmodType {alg F} alg_zmodMixin. + + +Lemma add_pi x y : \pi_{alg F} x + \pi_{alg F} y + = \pi_{alg F} (add_algcreal x y). +Proof. by rewrite [_ + _]piE. Qed. + +Lemma opp_pi x : - \pi_{alg F} x = \pi_{alg F} (opp_algcreal x). +Proof. by rewrite [- _]piE. Qed. + +Lemma zeroE : 0 = \pi_{alg F} zero_algcreal. +Proof. by rewrite [0]piE. Qed. + +Lemma sub_pi x y : \pi_{alg F} x - \pi_{alg F} y + = \pi_{alg F} (add_algcreal x (opp_algcreal y)). +Proof. by rewrite [_ - _]piE. Qed. + +Lemma mul_algC : commutative mul_alg. +Proof. +elim/quotW=> x; elim/quotW=> y; rewrite !piE -equiv_alg /=. +by apply: eq_crealP; exists m0=> * /=; rewrite mulrC subrr normr0. +Qed. + +Lemma mul_algA : associative mul_alg. +Proof. +elim/quotW=> x; elim/quotW=> y; elim/quotW=> z; rewrite !piE -equiv_alg /=. +by apply: eq_crealP; exists m0=> * /=; rewrite mulrA subrr normr0. +Qed. + +Lemma mul_1alg : left_id one_alg mul_alg. +Proof. by elim/quotW=> x; rewrite piE -equiv_alg /= mul_1creal. Qed. + +Lemma mul_alg_addl : left_distributive mul_alg add_alg. +Proof. +elim/quotW=> x; elim/quotW=> y; elim/quotW=> z; rewrite !piE -equiv_alg /=. +by apply: eq_crealP; exists m0=> * /=; rewrite mulrDl subrr normr0. +Qed. + +Implicit Arguments neq_creal_cst [F x y]. +Prenex Implicits neq_creal_cst. + +Lemma nonzero1_alg : one_alg != zero_alg. +Proof. by rewrite piE -(rwP neq_algcrealP) (rwP neq_creal_cst) oner_eq0. Qed. + +Definition alg_comRingMixin := + ComRingMixin mul_algA mul_algC mul_1alg mul_alg_addl nonzero1_alg. +Canonical alg_Ring := Eval hnf in RingType alg alg_comRingMixin. +Canonical alg_comRing := Eval hnf in ComRingType alg mul_algC. +Canonical alg_of_Ring := Eval hnf in RingType {alg F} alg_comRingMixin. +Canonical alg_of_comRing := Eval hnf in ComRingType {alg F} mul_algC. + +Lemma mul_pi x y : \pi_{alg F} x * \pi_{alg F} y + = \pi_{alg F} (mul_algcreal x y). +Proof. by rewrite [_ * _]piE. Qed. + +Lemma oneE : 1 = \pi_{alg F} one_algcreal. +Proof. by rewrite [1]piE. Qed. + +Lemma mul_Valg (x : alg) : x != zero_alg -> mul_alg (inv_alg x) x = one_alg. +Proof. +elim/quotW: x=> x; rewrite piE -(rwP neq_algcrealP) /= => x_neq0. +apply/eqP; rewrite piE; apply/eq_algcrealP; rewrite /= /inv_algcreal. +case: eq_algcreal_dec=> [/(_ x_neq0) //|/= x_neq0']. +apply: eq_crealP; exists_big_modulus m F. + by move=> * /=; rewrite mulVf ?subrr ?normr0 ?creal_neq0_always. +by close. +Qed. + +Lemma inv_alg0 : inv_alg zero_alg = zero_alg. +Proof. +rewrite !piE -equiv_alg /= /inv_algcreal. +by case: eq_algcreal_dec=> //= zero_neq0; move: (eq_creal_refl zero_neq0). +Qed. + +Lemma to_alg_additive : additive to_alg. +Proof. +move=> x y /=; rewrite !piE sub_pi -equiv_alg /=. +by apply: eq_crealP; exists m0=> * /=; rewrite subrr normr0. +Qed. + +Canonical to_alg_is_additive := Additive to_alg_additive. + +Lemma to_alg_multiplicative : multiplicative to_alg. +Proof. +split=> [x y |] //; rewrite !piE mul_pi -equiv_alg. +by apply: eq_crealP; exists m0=> * /=; rewrite subrr normr0. +Qed. + +Canonical to_alg_is_rmorphism := AddRMorphism to_alg_multiplicative. + +Lemma expn_pi (x : algcreal) (n : nat) : + (\pi_{alg F} x) ^+ n = \pi (exp_algcreal x n). +Proof. +rewrite /exp_algcreal; case: n=> [|n]; first by rewrite expr0 oneE. +rewrite exprS iteropS; elim: n=> /= [|n ihn]; rewrite ?expr0 ?mulr1 //. +by rewrite exprS ihn mul_pi. +Qed. + +Lemma horner_pi (p : {poly F}) (x : algcreal) : + (p ^ to_alg).[\pi_alg x] = \pi (horner_algcreal p x). +Proof. +rewrite horner_coef /horner_algcreal size_map_poly. +apply: (big_ind2 (fun x y => x = \pi_alg y)). ++ by rewrite zeroE. ++ by move=> u u' v v' -> ->; rewrite [_ + _]piE. +by move=> i /= _; rewrite expn_pi coef_map /= [_ * _]piE. +Qed. + +(* Defining annihilating polynomials for algebraics *) +Definition annul_alg : {alg F} -> {poly F} := locked (annul_creal \o repr). + +Lemma root_annul_algcreal (x : algcreal) : ((annul_alg (\pi x)).[x] == 0)%CR. +Proof. by unlock annul_alg; rewrite /= -pi_algK root_annul_creal. Qed. + +Lemma root_annul_alg (x : {alg F}) : root ((annul_alg x) ^ to_alg) x. +Proof. +apply/rootP; rewrite -[x]reprK horner_pi /= zeroE -equiv_alg. +by rewrite horner_algcrealE root_annul_algcreal. +Qed. + +Lemma monic_annul_alg (x : {alg F}) : annul_alg x \is monic. +Proof. by unlock annul_alg; rewrite monic_annul_creal. Qed. + +Lemma annul_alg_neq0 (x : {alg F}) : annul_alg x != 0. +Proof. by rewrite monic_neq0 ?monic_annul_alg. Qed. + +Definition AlgFieldUnitMixin := FieldUnitMixin mul_Valg inv_alg0. +Canonical alg_unitRing := + Eval hnf in UnitRingType alg AlgFieldUnitMixin. +Canonical alg_comUnitRing := Eval hnf in [comUnitRingType of alg]. +Canonical alg_of_unitRing := + Eval hnf in UnitRingType {alg F} AlgFieldUnitMixin. +Canonical alg_of_comUnitRing := Eval hnf in [comUnitRingType of {alg F}]. + +Lemma field_axiom : GRing.Field.mixin_of alg_unitRing. Proof. exact. Qed. + +Definition AlgFieldIdomainMixin := (FieldIdomainMixin field_axiom). +Canonical alg_iDomain := + Eval hnf in IdomainType alg (FieldIdomainMixin field_axiom). +Canonical alg_fieldType := FieldType alg field_axiom. +Canonical alg_of_iDomain := + Eval hnf in IdomainType {alg F} (FieldIdomainMixin field_axiom). +Canonical alg_of_fieldType := FieldType {alg F} field_axiom. + +Lemma inv_pi x : (\pi_{alg F} x)^-1 = \pi_{alg F} (inv_algcreal x). +Proof. by rewrite [_^-1]piE. Qed. + +Lemma div_pi x y : \pi_{alg F} x / \pi_{alg F} y + = \pi_{alg F} (mul_algcreal x (inv_algcreal y)). +Proof. by rewrite [_ / _]piE. Qed. + +Definition lt_alg := lift_fun2 {alg F} lt_algcreal. +Definition le_alg := lift_fun2 {alg F} le_algcreal. + +Lemma lt_alg_pi : {mono \pi_{alg F} : x y / lt_algcreal x y >-> lt_alg x y}. +Proof. +move=> x y; unlock lt_alg; rewrite /lt_algcreal. +by do 2!case: ltVge_algcreal_dec=> //=; rewrite !pi_algK. +Qed. + +Canonical lt_alg_pi_mono := PiMono2 lt_alg_pi. + +Lemma le_alg_pi : {mono \pi_{alg F} : x y / le_algcreal x y >-> le_alg x y}. +Proof. +move=> x y; unlock le_alg; rewrite /le_algcreal. +by do 2!case: ltVge_algcreal_dec=> //=; rewrite !pi_algK. +Qed. + +Canonical le_alg_pi_mono := PiMono2 le_alg_pi. + +Definition norm_alg := lift_op1 {alg F} norm_algcreal. + +Lemma norm_alg_pi : {morph \pi_{alg F} : x / norm_algcreal x >-> norm_alg x}. +Proof. +move=> x; unlock norm_alg; rewrite /norm_algcreal /le_algcreal. +by do 2!case: ltVge_algcreal_dec=> //=; rewrite !(pi_opp, pi_algK, reprK). +Qed. + +Canonical norm_alg_pi_morph := PiMorph1 norm_alg_pi. + +(* begin hide *) +(* Lemma norm_pi (x : algcreal) : `|\pi_{alg F} x| = \pi (norm_algcreal x). *) +(* Proof. by rewrite /norm_algcreal -lt_pi -zeroE -lerNgt fun_if -opp_pi. Qed. *) +(* end hide *) + +Lemma add_alg_gt0 x y : lt_alg zero_alg x -> lt_alg zero_alg y -> + lt_alg zero_alg (x + y). +Proof. +rewrite -[x]reprK -[y]reprK !piE -!(rwP lt_algcrealP). +move=> x_gt0 y_gt0; pose_big_enough i. + apply: (@lt_crealP _ (diff x_gt0 + diff y_gt0) i i) => //. + by rewrite addr_gt0 ?diff_gt0. + by rewrite /= add0r ler_add // ?diff0P. +by close. +Qed. + +Lemma mul_alg_gt0 x y : lt_alg zero_alg x -> lt_alg zero_alg y -> + lt_alg zero_alg (x * y). +Proof. +rewrite -[x]reprK -[y]reprK !piE -!(rwP lt_algcrealP). +move=> x_gt0 y_gt0; pose_big_enough i. + apply: (@lt_crealP _ (diff x_gt0 * diff y_gt0) i i) => //. + by rewrite pmulr_rgt0 ?diff_gt0. + rewrite /= add0r (@ler_trans _ (diff x_gt0 * (repr y) i)) //. + by rewrite ler_wpmul2l ?(ltrW (diff_gt0 _)) // diff0P. + by rewrite ler_wpmul2r ?diff0P ?ltrW ?creal_gt0_always. +by close. +Qed. + +Lemma gt0_alg_nlt0 x : lt_alg zero_alg x -> ~~ lt_alg x zero_alg. +Proof. +rewrite -[x]reprK !piE -!(rwP lt_algcrealP, rwP le_algcrealP). +move=> hx; pose_big_enough i. + apply: (@le_crealP _ i)=> j /= hj. + by rewrite ltrW // creal_gt0_always. +by close. +Qed. + +Lemma sub_alg_gt0 x y : lt_alg zero_alg (y - x) = lt_alg x y. +Proof. +rewrite -[x]reprK -[y]reprK !piE; apply/lt_algcrealP/lt_algcrealP=> /= hxy. + pose_big_enough i. + apply: (@lt_crealP _ (diff hxy) i i); rewrite ?diff_gt0 //. + by rewrite (monoLR (addNKr _) (ler_add2l _)) addrC diff0P. + by close. +pose_big_enough i. + apply: (@lt_crealP _ (diff hxy) i i); rewrite ?diff_gt0 //. + by rewrite (monoRL (addrK _) (ler_add2r _)) add0r addrC diffP. +by close. +Qed. + +Lemma lt0_alg_total (x : alg) : x != zero_alg -> + lt_alg zero_alg x || lt_alg x zero_alg. +Proof. +rewrite -[x]reprK !piE => /neq_algcrealP /= x_neq0. +apply/orP; rewrite -!(rwP lt_algcrealP). +case/neq_creal_ltVgt: x_neq0=> /= [lt_x0|lt_0x]; [right|by left]. +pose_big_enough i. + by apply: (@lt_crealP _ (diff lt_x0) i i); rewrite ?diff_gt0 ?diffP. +by close. +Qed. + +Lemma norm_algN x : norm_alg (- x) = norm_alg x. +Proof. +rewrite -[x]reprK !piE /= -equiv_alg !norm_algcrealE; apply: eq_crealP. +exists_big_modulus m F=> [e i e_gt0 hi /=|]. + by rewrite normrN subrr normr0. +by close. +Qed. + +Lemma ge0_norm_alg x : le_alg 0 x -> norm_alg x = x. +Proof. by rewrite -[x]reprK !piE /= /norm_algcreal => ->. Qed. + +Lemma lt_alg0N (x : alg) : lt_alg 0 (- x) = lt_alg x 0. +Proof. by rewrite -sub0r sub_alg_gt0. Qed. + +Lemma lt_alg00 : lt_alg zero_alg zero_alg = false. +Proof. by have /implyP := @gt0_alg_nlt0 0; case: lt_alg. Qed. + +Lemma le_alg_def x y : le_alg x y = (y == x) || lt_alg x y. +Proof. +rewrite -[x]reprK -[y]reprK eq_sym piE [lt_alg _ _]piE; apply/le_algcrealP/orP. + move=> /le_creal_neqVlt [/eq_algcrealP/eqquotP/eqP-> //|lt_xy]; first by left. + by right; apply/lt_algcrealP. +by move=> [/eqP/eqquotP/eq_algcrealP-> //| /lt_algcrealP /lt_crealW]. +Qed. + +Definition AlgNumFieldMixin := RealLtMixin add_alg_gt0 mul_alg_gt0 + gt0_alg_nlt0 sub_alg_gt0 lt0_alg_total norm_algN ge0_norm_alg le_alg_def. +Canonical alg_numDomainType := NumDomainType alg AlgNumFieldMixin. +Canonical alg_numFieldType := [numFieldType of alg]. +Canonical alg_of_numDomainType := [numDomainType of {alg F}]. +Canonical alg_of_numFieldType := [numFieldType of {alg F}]. + +Definition AlgRealFieldMixin := RealLeAxiom alg. +Canonical alg_realDomainType := RealDomainType alg AlgRealFieldMixin. +Canonical alg_realFieldType := [realFieldType of alg]. +Canonical alg_of_realDomainType := [realDomainType of {alg F}]. +Canonical alg_of_realFieldType := [realFieldType of {alg F}]. + +Lemma lt_pi x y : \pi_{alg F} x < \pi y = lt_algcreal x y. +Proof. by rewrite [_ < _]lt_alg_pi. Qed. + +Lemma le_pi x y : \pi_{alg F} x <= \pi y = le_algcreal x y. +Proof. by rewrite [_ <= _]le_alg_pi. Qed. + +Lemma norm_pi (x : algcreal) : `|\pi_{alg F} x| = \pi (norm_algcreal x). +Proof. by rewrite [`|_|]piE. Qed. + +Lemma lt_algP (x y : algcreal) : reflect (x < y)%CR (\pi_{alg F} x < \pi y). +Proof. by rewrite lt_pi; apply: lt_algcrealP. Qed. +Implicit Arguments lt_algP [x y]. + +Lemma le_algP (x y : algcreal) : reflect (x <= y)%CR (\pi_{alg F} x <= \pi y). +Proof. by rewrite le_pi; apply: le_algcrealP. Qed. +Implicit Arguments le_algP [x y]. +Prenex Implicits lt_algP le_algP. + +Lemma ler_to_alg : {mono to_alg : x y / x <= y}. +Proof. +apply: homo_mono=> x y lt_xy; rewrite !piE -(rwP lt_algP). +by apply/lt_creal_cst; rewrite lt_xy. +Qed. + +Lemma ltr_to_alg : {mono to_alg : x y / x < y}. +Proof. by apply: lerW_mono; apply: ler_to_alg. Qed. + +Lemma normr_to_alg : { morph to_alg : x / `|x| }. +Proof. +move=> x /=; have [] := ger0P; have [] := ger0P x%:RA; + rewrite ?rmorph0 ?rmorphN ?oppr0 //=. + by rewrite ltr_to_alg lerNgt => ->. +by rewrite ler_to_alg ltrNge => ->. +Qed. + +Lemma inf_alg_proof x : {d | 0 < d & 0 < x -> (d%:RA < x)}. +Proof. +have [|] := lerP; first by exists 1. +elim/quotW: x=> x; rewrite zeroE=> /lt_algP /= x_gt0. +exists (diff x_gt0 / 2%:R); first by rewrite pmulr_rgt0 ?gtr0E ?diff_gt0. +rewrite piE -(rwP lt_algP) /= => _; pose_big_enough i. + apply: (@lt_crealP _ (diff x_gt0 / 2%:R) i i) => //. + by rewrite pmulr_rgt0 ?gtr0E ?diff_gt0. + by rewrite -[_ + _](splitf 2) diff0P. +by close. +Qed. + +Definition inf_alg (x : {alg F}) : F := + let: exist2 d _ _ := inf_alg_proof x in d. + +Lemma inf_alg_gt0 x : 0 < inf_alg x. +Proof. by rewrite /inf_alg; case: inf_alg_proof. Qed. + +Lemma inf_lt_alg x : 0 < x -> (inf_alg x)%:RA < x. +Proof. by rewrite /inf_alg=> x_gt0; case: inf_alg_proof=> d _ /(_ x_gt0). Qed. + +Lemma approx_proof x e : {y | 0 < e -> `|x - y%:RA| < e}. +Proof. +elim/quotW:x => x; pose_big_enough i. + exists (x i)=> e_gt0; rewrite (ltr_trans _ (inf_lt_alg _)) //. + rewrite !piE sub_pi norm_pi -(rwP lt_algP) /= norm_algcrealE /=. + pose_big_enough j. + apply: (@lt_crealP _ (inf_alg e / 2%:R) j j) => //. + by rewrite pmulr_rgt0 ?gtr0E ?inf_alg_gt0. + rewrite /= {2}[inf_alg e](splitf 2) /= ler_add2r. + by rewrite ltrW // cauchymodP ?pmulr_rgt0 ?gtr0E ?inf_alg_gt0. + by close. +by close. +Qed. + +Definition approx := locked + (fun (x : alg) (e : alg) => projT1 (approx_proof x e) : F). + +Lemma approxP (x e e': alg) : 0 < e -> e <= e' -> `|x - (approx x e)%:RA| < e'. +Proof. +by unlock approx; case: approx_proof=> /= y hy /hy /ltr_le_trans hy' /hy'. +Qed. + +Lemma alg_archi : Num.archimedean_axiom alg_of_numDomainType. +Proof. +move=> x; move: {x}`|x| (normr_ge0 x) => x x_ge0. +pose a := approx x 1%:RA; exists (Num.bound (a + 1)). +have := @archi_boundP _ (a + 1); rewrite -ltr_to_alg rmorph_nat. +have := @approxP x _ _ ltr01 (lerr _); rewrite ltr_distl -/a => /andP [_ hxa]. +rewrite -ler_to_alg rmorphD /= (ler_trans _ (ltrW hxa)) //. +by move=> /(_ isT) /(ltr_trans _)->. +Qed. + +Canonical alg_archiFieldType := ArchiFieldType alg alg_archi. +Canonical alg_of_archiFieldType := [archiFieldType of {alg F}]. + +(**************************************************************************) +(* At this stage, algebraics form an archimedian field. We now build the *) +(* material to prove the intermediate value theorem. We first prove a *) +(* "weak version", which expresses that the extension {alg F} indeed *) +(* contains solutions of the intermediate value probelem in F *) +(**************************************************************************) + +Notation "'Y" := 'X%:P. + +Lemma weak_ivt (p : {poly F}) (a b : F) : a <= b -> p.[a] <= 0 <= p.[b] -> + { x : alg | a%:RA <= x <= b%:RA & root (p ^ to_alg) x }. +Proof. +move=> le_ab; have [-> _|p_neq0] := eqVneq p 0. + by exists a%:RA; rewrite ?lerr ?ler_to_alg // rmorph0 root0. +move=> /andP[pa_le0 pb_ge0]; apply/sig2W. +have hpab: p.[a] * p.[b] <= 0 by rewrite mulr_le0_ge0. +move=> {pa_le0 pb_ge0}; wlog monic_p : p hpab p_neq0 / p \is monic. + set q := (lead_coef p) ^-1 *: p => /(_ q). + rewrite !hornerZ mulrCA !mulrA -mulrA mulr_ge0_le0 //; last first. + by rewrite (@exprn_even_ge0 _ 2). + have mq: q \is monic by rewrite monicE lead_coefZ mulVf ?lead_coef_eq0. + rewrite monic_neq0 ?mq=> // [] [] // x hx hqx; exists x=> //. + move: hqx; rewrite /q -mul_polyC rmorphM /= rootM map_polyC rootC. + by rewrite fmorph_eq0 invr_eq0 lead_coef_eq0 (negPf p_neq0). +pose c := (a + b) / 2%:R; pose r := (b - a) / 2%:R. +have r_ge0 : 0 <= r by rewrite mulr_ge0 ?ger0E // subr_ge0. +have hab: ((c - r = a)%R * (c + r = b)%R)%type. + rewrite -mulrDl -mulrBl opprD addrA addrK opprK. + rewrite addrAC addrA [a + _ + _]addrAC subrr add0r. + by rewrite !mulrDl -. +have hp: p.[c - r] * p.[c + r] <= 0 by rewrite !hab. +pose x := AlgDom monic_p r_ge0 hp; exists (\pi_alg (to_algcreal x)). + rewrite !piE; apply/andP; rewrite -!(rwP le_algP) /=. + split; + by do [ unlock to_algcreal=> /=; apply: (@le_crealP _ 0%N)=> /= j _; + have := @to_algcreal_ofP p c r 0%N j r_ge0 isT; + rewrite ler_distl /= expr0 divr1 !hab=> /andP []]. +apply/rootP; rewrite horner_pi zeroE -equiv_alg horner_algcrealE /=. +by rewrite -(@to_algcrealP x); unlock to_algcreal. +Qed. + +(* any sequence of algebraic can be expressed as a sequence of +polynomials in a unique algebraic *) +Lemma pet_alg_proof (s : seq alg) : + { ap : {alg F} * seq {poly F} | + [forall i : 'I_(size s), (ap.2`_i ^ to_alg).[ap.1] == s`_i] + & size ap.2 = size s }. +Proof. +apply: sig2_eqW; elim: s; first by exists (0,[::])=> //; apply/forallP=> [] []. +move=> x s [[a sp] /forallP /= hs hsize]. +have:= char0_PET _ (root_annul_alg a) _ (root_annul_alg x). +rewrite !annul_alg_neq0 => /(_ isT isT (char_num _)) /= [n [[p hp] [q hq]]]. +exists (x *+ n - a, q :: [seq r \Po p | r <- sp]); last first. + by rewrite /= size_map hsize. +apply/forallP=> /=; rewrite -add1n=> i; apply/eqP. +have [k->|l->] := splitP i; first by rewrite !ord1. +rewrite add1n /= (nth_map 0) ?hsize // map_comp_poly /=. +by rewrite horner_comp hp; apply/eqP. +Qed. + +(****************************************************************************) +(* Given a sequence s of algebraics (seq {alg F}) *) +(* pet_alg == primitive algebraic *) +(* pet_alg_poly == sequence of polynomials such that when instanciated in *) +(* pet_alg gives back the sequence s (cf. pet_algK) *) +(* *) +(* Given a polynomial p on algebraic {poly {alg F}} *) +(* pet_ground == bivariate polynomial such that when the inner *) +(* variable ('Y) is instanciated in pet_alg gives back *) +(* the polynomial p. *) +(****************************************************************************) + +Definition pet_alg s : {alg F} := + let: exist2 (a, _) _ _ := pet_alg_proof s in a. +Definition pet_alg_poly s : seq {poly F}:= + let: exist2 (_, sp) _ _ := pet_alg_proof s in sp. + +Lemma size_pet_alg_poly s : size (pet_alg_poly s) = size s. +Proof. by unlock pet_alg_poly; case: pet_alg_proof. Qed. + +Lemma pet_algK s i : + ((pet_alg_poly s)`_i ^ to_alg).[pet_alg s] = s`_i. +Proof. +rewrite /pet_alg /pet_alg_poly; case: pet_alg_proof. +move=> [a sp] /= /forallP hs hsize; have [lt_is|le_si] := ltnP i (size s). + by rewrite -[i]/(val (Ordinal lt_is)); apply/eqP; apply: hs. +by rewrite !nth_default ?hsize // rmorph0 horner0. +Qed. + +Definition poly_ground := locked (fun (p : {poly {alg F}}) => + swapXY (Poly (pet_alg_poly p)) : {poly {poly F}}). + +Lemma sizeY_poly_ground p : sizeY (poly_ground p) = size p. +Proof. +unlock poly_ground; rewrite sizeYE swapXYK; have [->|p_neq0] := eqVneq p 0. + apply/eqP; rewrite size_poly0 eqn_leq leq0n (leq_trans (size_Poly _)) //. + by rewrite size_pet_alg_poly size_poly0. +rewrite (@PolyK _ 0) -?nth_last ?size_pet_alg_poly //. +have /eqP := (pet_algK p (size p).-1); apply: contraL=> /eqP->. +by rewrite rmorph0 horner0 -lead_coefE eq_sym lead_coef_eq0. +Qed. + +Lemma poly_ground_eq0 p : (poly_ground p == 0) = (p == 0). +Proof. by rewrite -sizeY_eq0 sizeY_poly_ground size_poly_eq0. Qed. + +Lemma poly_ground0 : poly_ground 0 = 0. +Proof. by apply/eqP; rewrite poly_ground_eq0. Qed. + +Lemma poly_groundK p : + ((poly_ground p) ^ (map_poly to_alg)).[(pet_alg p)%:P] = p. +Proof. +have [->|p_neq0] := eqVneq p 0; first by rewrite poly_ground0 rmorph0 horner0. +unlock poly_ground; rewrite horner_polyC /eval /= swapXY_map swapXYK. +apply/polyP=> i /=; rewrite coef_map_id0 ?horner0 // coef_map /=. +by rewrite coef_Poly pet_algK. +Qed. + +Lemma annul_from_alg_proof (p : {poly alg}) (q : {poly F}) : + p != 0 -> q != 0 -> root (q ^ to_alg) (pet_alg p) + -> {r | resultant (poly_ground p) (r ^ polyC) != 0 + & (r != 0) && (root (r ^ to_alg) (pet_alg p))}. +Proof. +move=> p_neq0; elim: (size q) {-2}q (leqnn (size q))=> {q} [|n ihn] q. + by rewrite size_poly_leq0=> ->. +move=> size_q q_neq0 hq; apply/sig2_eqW. +have [|apq_neq0] := + eqVneq (resultant (poly_ground p) (q ^ polyC)) 0; last first. + by exists q=> //; rewrite q_neq0. +move/eqP; rewrite resultant_eq0 ltn_neqAle eq_sym -coprimep_def. +move=> /andP[] /(Bezout_coprimepPn _ _) []. ++ by rewrite poly_ground_eq0. ++ by rewrite map_polyC_eq0. +move=> [u v] /and3P [] /andP [u_neq0 ltn_uq] v_neq0 ltn_vp hpq. +rewrite ?size_map_polyC in ltn_uq ltn_vp. +rewrite ?size_poly_gt0 in u_neq0 v_neq0. +pose a := pet_alg p. +have := erefl (size ((u * poly_ground p) ^ (map_poly to_alg)).[a%:P]). +rewrite {2}hpq !{1}rmorphM /= !{1}hornerM poly_groundK -map_poly_comp /=. +have /eq_map_poly-> : (map_poly to_alg) \o polyC =1 polyC \o to_alg. + by move=> r /=; rewrite map_polyC. +rewrite map_poly_comp horner_map (rootP hq) mulr0 size_poly0. +move/eqP; rewrite size_poly_eq0 mulf_eq0 (negPf p_neq0) orbF. +pose u' : {poly F} := lead_coef (swapXY u). +have [/rootP u'a_eq0|u'a_neq0] := eqVneq (u' ^ to_alg).[a] 0; last first. + rewrite horner_polyC -lead_coef_eq0 lead_coef_map_eq /=; + by do ?rewrite swapXY_map /= lead_coef_map_eq /= + ?map_poly_eq0 ?lead_coef_eq0 ?swapXY_eq0 ?(negPf u'a_neq0). +have u'_neq0 : u' != 0 by rewrite lead_coef_eq0 swapXY_eq0. +have size_u' : (size u' < size q)%N. + by rewrite /u' (leq_ltn_trans (max_size_lead_coefXY _)) // sizeYE swapXYK. +move: u'a_eq0=> /ihn [] //; first by rewrite -ltnS (leq_trans size_u'). +by move=> r; exists r. +Qed. + +Definition annul_pet_alg (p : {poly {alg F}}) : {poly F} := + if (p != 0) =P true is ReflectT p_neq0 then + let: exist2 r _ _ := + annul_from_alg_proof p_neq0 (annul_alg_neq0 _) (root_annul_alg _) in r + else 0. + +Lemma root_annul_pet_alg p : root (annul_pet_alg p ^ to_alg) (pet_alg p). +Proof. +rewrite /annul_pet_alg; case: eqP=> /=; last by rewrite ?rmorph0 ?root0. +by move=> p_neq0; case: annul_from_alg_proof=> ? ? /andP[]. +Qed. + +Definition annul_from_alg p := + if (size (poly_ground p) == 1)%N then lead_coef (poly_ground p) + else resultant (poly_ground p) (annul_pet_alg p ^ polyC). + +Lemma annul_from_alg_neq0 p : p != 0 -> annul_from_alg p != 0. +Proof. +move=> p_neq0; rewrite /annul_from_alg. +case: ifP; first by rewrite lead_coef_eq0 poly_ground_eq0. +rewrite /annul_pet_alg; case: eqP p_neq0=> //= p_neq0 _. +by case: annul_from_alg_proof. +Qed. + +Lemma annul_pet_alg_neq0 p : p != 0 -> annul_pet_alg p != 0. +Proof. +rewrite /annul_pet_alg; case: eqP=> /=; last by rewrite ?rmorph0 ?root0. +by move=> p_neq0; case: annul_from_alg_proof=> ? ? /andP[]. +Qed. + +End RealAlg. + +Notation to_alg F := (@to_alg_def _ (Phant F)). +Notation "x %:RA" := (to_alg _ x) + (at level 2, left associativity, format "x %:RA"). + +Lemma upper_nthrootVP (F : archiFieldType) (x : F) (i : nat) : + 0 < x -> (Num.bound (x ^-1) <= i)%N -> 2%:R ^- i < x. +Proof. +move=> x_gt0 hx; rewrite -ltf_pinv -?topredE //= ?gtr0E //. +by rewrite invrK upper_nthrootP. +Qed. + +Notation "{ 'alg' F }" := (alg_of (Phant F)). + +Section AlgAlg. + +Variable F : archiFieldType. + +Local Open Scope ring_scope. + +Local Notation "p ^ f" := (map_poly f p) : ring_scope. +Local Notation "'Y" := 'X%:P. +Local Notation m0 := (fun _ => 0%N). + +Definition approx2 (x : {alg {alg F}}) i := + approx (approx x (2%:R ^- i)) (2%:R ^- i). + +Lemma asympt_approx2 x : { asympt e : i / `|(approx2 x i)%:RA%:RA - x| < e }. +Proof. +exists_big_modulus m {alg {alg F}}. + move=> e i e_gt0 hi; rewrite distrC /approx2. + rewrite (@split_dist_add _ (approx x (2%:R ^- i))%:RA) //. + rewrite approxP ?gtr0E // ltrW //. + by rewrite upper_nthrootVP ?divrn_gt0 ?ltr_to_alg. + rewrite (ltr_trans _ (inf_lt_alg _)) ?divrn_gt0 //. + rewrite -rmorphB -normr_to_alg ltr_to_alg approxP ?gtr0E // ltrW //. + by rewrite upper_nthrootVP ?divrn_gt0 ?inf_alg_gt0 ?ltr_to_alg. +by close. +Qed. + +Lemma from_alg_crealP (x : {alg {alg F}}) : creal_axiom (approx2 x). +Proof. +exists_big_modulus m F. + move=> e i j e_gt0 hi hj; rewrite -2!ltr_to_alg !normr_to_alg !rmorphB /=. + rewrite (@split_dist_add _ x) // ?[`|_ - _%:RA|]distrC; + by rewrite (@asympt1modP _ _ (asympt_approx2 x)) ?divrn_gt0 ?ltr_to_alg. +by close. +Qed. + +Definition from_alg_creal := locked (fun x => CReal (from_alg_crealP x)). + +Lemma to_alg_crealP (x : creal F) : creal_axiom (fun i => to_alg F (x i)). +Proof. +exists_big_modulus m (alg F). + move=> e i j e_gt0 hi hj. + rewrite -rmorphB -normr_to_alg (ltr_trans _ (inf_lt_alg _)) //. + by rewrite ltr_to_alg cauchymodP ?inf_alg_gt0. +by close. +Qed. +Definition to_alg_creal x := CReal (to_alg_crealP x). + +Lemma horner_to_alg_creal p x : + ((p ^ to_alg F).[to_alg_creal x] == to_alg_creal p.[x])%CR. +Proof. +by apply: eq_crealP; exists m0=> * /=; rewrite horner_map subrr normr0. +Qed. + +Lemma neq_to_alg_creal x y : + (to_alg_creal x != to_alg_creal y)%CR <-> (x != y)%CR. +Proof. +split=> neq_xy. + pose_big_enough i. + apply: (@neq_crealP _ (inf_alg (lbound neq_xy)) i i) => //. + by rewrite inf_alg_gt0. + rewrite -ler_to_alg normr_to_alg rmorphB /= ltrW //. + by rewrite (ltr_le_trans (inf_lt_alg _)) ?lbound_gt0 ?lboundP. + by close. +pose_big_enough i. + apply: (@neq_crealP _ (lbound neq_xy)%:RA i i) => //. + by rewrite ltr_to_alg lbound_gt0. + by rewrite -rmorphB -normr_to_alg ler_to_alg lboundP. +by close. +Qed. + +Lemma eq_to_alg_creal x y : + (to_alg_creal x == to_alg_creal y)%CR -> (x == y)%CR. +Proof. by move=> hxy /neq_to_alg_creal. Qed. + +Lemma to_alg_creal0 : (to_alg_creal 0 == 0)%CR. +Proof. by apply: eq_crealP; exists m0=> * /=; rewrite subrr normr0. Qed. + +Import Setoid. + +Add Morphism to_alg_creal + with signature (@eq_creal _) ==> (@eq_creal _) as to_alg_creal_morph. +Proof. by move=> x y hxy /neq_to_alg_creal. Qed. +Global Existing Instance to_alg_creal_morph_Proper. + +Lemma to_alg_creal_repr (x : {alg F}) : (to_alg_creal (repr x) == x%:CR)%CR. +Proof. +apply: eq_crealP; exists_big_modulus m {alg F}. + move=> e i e_gt0 hi /=; rewrite (ler_lt_trans _ (inf_lt_alg _)) //. + rewrite -{2}[x]reprK !piE sub_pi norm_pi. + rewrite -(rwP (le_algP _ _)) norm_algcrealE /=; pose_big_enough j. + apply: (@le_crealP _ j)=> k hk /=. + by rewrite ltrW // cauchymodP ?inf_alg_gt0. + by close. +by close. +Qed. + +Local Open Scope quotient_scope. + +Lemma cst_pi (x : algcreal F) : ((\pi_{alg F} x)%:CR == to_alg_creal x)%CR. +Proof. +apply: eq_crealP; exists_big_modulus m {alg F}. + move=> e i e_gt0 hi /=; rewrite (ltr_trans _ (inf_lt_alg _)) //. + rewrite !piE sub_pi norm_pi /= -(rwP (lt_algP _ _)) norm_algcrealE /=. + pose_big_enough j. + apply: (@lt_crealP _ (inf_alg e / 2%:R) j j) => //. + by rewrite ?divrn_gt0 ?inf_alg_gt0. + rewrite /= {2}[inf_alg _](splitf 2) ler_add2r ltrW // distrC. + by rewrite cauchymodP ?divrn_gt0 ?inf_alg_gt0. + by close. +by close. +Qed. + +End AlgAlg. + +Section AlgAlgAlg. + +Variable F : archiFieldType. + +Local Open Scope ring_scope. + +Local Notation "p ^ f" := (map_poly f p) : ring_scope. +Local Notation "'Y" := 'X%:P. + +Lemma from_alg_crealK (x : {alg {alg F}}) : + (to_alg_creal (to_alg_creal (from_alg_creal x)) == x%:CR)%CR. +Proof. +apply: eq_crealP; exists_big_modulus m {alg {alg F}}. + move=> e i e_gt0 hi; unlock from_alg_creal=> /=. + by rewrite (@asympt1modP _ _ (asympt_approx2 x)). +by close. +Qed. + +Lemma root_annul_from_alg_creal (x : {alg {alg F}}) : + ((annul_from_alg (annul_alg x)).[from_alg_creal x] == 0)%CR. +Proof. +do 2!apply: eq_to_alg_creal. +rewrite -!horner_to_alg_creal from_alg_crealK !to_alg_creal0. +rewrite horner_creal_cst; apply/eq_creal_cst; rewrite -rootE. +rewrite /annul_from_alg; have [/size_poly1P [c c_neq0 hc]|sp_neq1] := boolP (_ == _). + set p := _ ^ _; suff ->: p = (annul_alg x) ^ to_alg _ by apply: root_annul_alg. + congr (_ ^ _); rewrite -{2}[annul_alg x]poly_groundK /=. + by rewrite !hc lead_coefC map_polyC /= hornerC. +have [||[u v] /= [hu hv] hpq] := @resultant_in_ideal _ + (poly_ground (annul_alg x)) (annul_pet_alg (annul_alg x) ^ polyC). ++ rewrite ltn_neqAle eq_sym sp_neq1 //= lt0n size_poly_eq0. + by rewrite poly_ground_eq0 annul_alg_neq0. ++ rewrite size_map_polyC -(size_map_poly [rmorphism of to_alg _]) /=. + rewrite (root_size_gt1 _ (root_annul_pet_alg _)) //. + by rewrite map_poly_eq0 annul_pet_alg_neq0 ?annul_alg_neq0. +move: hpq=> /(f_equal (map_poly (map_poly (to_alg _)))). +rewrite map_polyC /= => /(f_equal (eval (pet_alg (annul_alg x))%:P)). +rewrite {1}/eval hornerC !rmorphD !{1}rmorphM /= /eval /= => ->. +rewrite -map_poly_comp /=. +have /eq_map_poly->: (map_poly (@to_alg F)) \o polyC =1 polyC \o (@to_alg F). + by move=> r /=; rewrite map_polyC. +rewrite map_poly_comp horner_map /= (rootP (root_annul_pet_alg _)) mulr0 addr0. +by rewrite rmorphM /= rootM orbC poly_groundK root_annul_alg. +Qed. + +Lemma annul_alg_from_alg_creal_neq0 (x : {alg {alg F}}) : + annul_from_alg (annul_alg x) != 0. +Proof. by rewrite annul_from_alg_neq0 ?annul_alg_neq0. Qed. + +Definition from_alg_algcreal x := + AlgCRealOf (@annul_alg_from_alg_creal_neq0 x) (@root_annul_from_alg_creal x). + +Definition from_alg : {alg {alg F}} -> {alg F} := + locked (\pi%qT \o from_alg_algcreal). + +Lemma from_algK : cancel from_alg (to_alg _). +Proof. +move=> x; unlock from_alg; rewrite -{2}[x]reprK piE -equiv_alg /= cst_pi. +by apply: eq_to_alg_creal; rewrite from_alg_crealK to_alg_creal_repr. +Qed. + +Lemma ivt (p : {poly (alg F)}) (a b : alg F) : a <= b -> + p.[a] <= 0 <= p.[b] -> exists2 x : alg F, a <= x <= b & root p x. +Proof. +move=> le_ab hp; have [x /andP [hax hxb]] := @weak_ivt _ _ _ _ le_ab hp. +rewrite -[x]from_algK fmorph_root=> rpx; exists (from_alg x)=> //. +by rewrite -ler_to_alg from_algK hax -ler_to_alg from_algK. +Qed. + +Canonical alg_rcfType := RcfType (alg F) ivt. +Canonical alg_of_rcfType := [rcfType of {alg F}]. + +End AlgAlgAlg. +End RealAlg. + +Notation "{ 'realclosure' F }" := (RealAlg.alg_of (Phant F)). + +Notation annul_realalg := RealAlg.annul_alg. +Notation realalg_of F := (@RealAlg.to_alg_def _ (Phant F)). +Notation "x %:RA" := (realalg_of x) + (at level 2, left associativity, format "x %:RA"). + +Canonical RealAlg.alg_eqType. +Canonical RealAlg.alg_choiceType. +Canonical RealAlg.alg_zmodType. +Canonical RealAlg.alg_Ring. +Canonical RealAlg.alg_comRing. +Canonical RealAlg.alg_unitRing. +Canonical RealAlg.alg_comUnitRing. +Canonical RealAlg.alg_iDomain. +Canonical RealAlg.alg_fieldType. +Canonical RealAlg.alg_numDomainType. +Canonical RealAlg.alg_numFieldType. +Canonical RealAlg.alg_realDomainType. +Canonical RealAlg.alg_realFieldType. +Canonical RealAlg.alg_archiFieldType. +Canonical RealAlg.alg_rcfType. + +Canonical RealAlg.alg_of_eqType. +Canonical RealAlg.alg_of_choiceType. +Canonical RealAlg.alg_of_zmodType. +Canonical RealAlg.alg_of_Ring. +Canonical RealAlg.alg_of_comRing. +Canonical RealAlg.alg_of_unitRing. +Canonical RealAlg.alg_of_comUnitRing. +Canonical RealAlg.alg_of_iDomain. +Canonical RealAlg.alg_of_fieldType. +Canonical RealAlg.alg_of_numDomainType. +Canonical RealAlg.alg_of_numFieldType. +Canonical RealAlg.alg_of_realDomainType. +Canonical RealAlg.alg_of_realFieldType. +Canonical RealAlg.alg_of_archiFieldType. +Canonical RealAlg.alg_of_rcfType. + +Canonical RealAlg.to_alg_is_rmorphism. +Canonical RealAlg.to_alg_is_additive. + +Section RealClosureTheory. + +Variable F : archiFieldType. +Notation R := {realclosure F}. + +Local Notation "p ^ f" := (map_poly f p) : ring_scope. + +Lemma root_annul_realalg (x : R) : root ((annul_realalg x) ^ realalg_of _) x. +Proof. exact: RealAlg.root_annul_alg. Qed. +Hint Resolve root_annul_realalg. + +Lemma monic_annul_realalg (x : R) : annul_realalg x \is monic. +Proof. exact: RealAlg.monic_annul_alg. Qed. +Hint Resolve monic_annul_realalg. + +Lemma annul_realalg_neq0 (x : R) : annul_realalg x != 0%R. +Proof. exact: RealAlg.annul_alg_neq0. Qed. +Hint Resolve annul_realalg_neq0. + +Lemma realalg_algebraic : integralRange (realalg_of F). +Proof. by move=> x; exists (annul_realalg x). Qed. + +End RealClosureTheory. + +Definition realalg := {realclosure rat}. +Canonical realalg_eqType := [eqType of realalg]. +Canonical realalg_choiceType := [choiceType of realalg]. +Canonical realalg_zmodType := [zmodType of realalg]. +Canonical realalg_ringType := [ringType of realalg]. +Canonical realalg_comRingType := [comRingType of realalg]. +Canonical realalg_unitRingType := [unitRingType of realalg]. +Canonical realalg_comUnitRingType := [comUnitRingType of realalg]. +Canonical realalg_idomainType := [idomainType of realalg]. +Canonical realalg_fieldTypeType := [fieldType of realalg]. +Canonical realalg_numDomainType := [numDomainType of realalg]. +Canonical realalg_numFieldType := [numFieldType of realalg]. +Canonical realalg_realDomainType := [realDomainType of realalg]. +Canonical realalg_realFieldType := [realFieldType of realalg]. +Canonical realalg_archiFieldType := [archiFieldType of realalg]. +Canonical realalg_rcfType := [rcfType of realalg]. + +Module RatRealAlg. +Canonical RealAlg.algdom_choiceType. +Definition realalgdom_CountMixin := + PcanCountMixin (@RealAlg.encode_algdomK [archiFieldType of rat]). +Canonical realalgdom_countType := + CountType (RealAlg.algdom [archiFieldType of rat]) realalgdom_CountMixin. +Definition realalg_countType := [countType of realalg]. +End RatRealAlg. + +Canonical RatRealAlg.realalg_countType. + +(* Require Import countalg. *) +(* Canonical realalg_countZmodType := [countZmodType of realalg]. *) +(* Canonical realalg_countRingType := [countRingType of realalg]. *) +(* Canonical realalg_countComRingType := [countComRingType of realalg]. *) +(* Canonical realalg_countUnitRingType := [countUnitRingType of realalg]. *) +(* Canonical realalg_countComUnitRingType := [countComUnitRingType of realalg]. *) +(* Canonical realalg_countIdomainType := [countIdomainType of realalg]. *) +(* Canonical realalg_countFieldTypeType := [countFieldType of realalg]. *) |
