diff --git a/auxresults.v b/auxresults.v index b62c786..1208299 100644 --- a/auxresults.v +++ b/auxresults.v @@ -288,7 +288,7 @@ Lemma set_nth_catl (i : nat) (e1 e2 : seq T) (x y : T) : (i < size e1)%N -> set_nth x (e1 ++ e2) i y = set_nth x e1 i y ++ e2. Proof. move: i e1 y; elim/last_ind : e2 => [i e1| e2 z ih i e1] y h; rewrite ?cats0 //. -rewrite -rcons_cat set_nth_rcons ?size_cat ?(leq_trans h) // ?leq_addr //. +rewrite -rcons_cat set_nth_rcons ?size_cat ?(leq_trans h) // ?leq_addr //. by rewrite ih // rcons_cat //. Qed. @@ -516,8 +516,9 @@ Qed. End WithEqType. Lemma subseq_nth_iota (T : eqType) (x : T) (s1 s2 : seq T) : - reflect (exists t, subseq t (iota 0 (size s2)) /\ s1 = [seq nth x s2 i | i <- t]) - (subseq s1 s2). + reflect + (exists t, subseq t (iota 0 (size s2)) /\ s1 = [seq nth x s2 i | i <- t]) + (subseq s1 s2). Proof. elim: s1 s2 => [|x1 s1 IHs1] s2/=. rewrite sub0seq; apply/Bool.ReflectT. @@ -1556,7 +1557,8 @@ Qed. End MoreNumDomainTheory. -Lemma sgz_prod (R : realDomainType) (I : Type) (r : seq I) (P : pred I) (F : I -> R) : +Lemma sgz_prod (R : realDomainType) (I : Type) + (r : seq I) (P : pred I) (F : I -> R) : sgz (\prod_(x <- r | P x) F x) = \prod_(x <- r | P x) sgz (F x). Proof. elim: r => [|x r IHr]; first by rewrite !big_nil sgz1. @@ -1596,10 +1598,16 @@ under eq_bigr => y _. rewrite -rmorph_prod. rewrite [\prod_(_ <- dec_roots _ | _) _](bigID (fun x => 0 < complex.Im x))/=. have im_conj: forall (z : F[i]), complex.Im z^* = - complex.Im z by case. -have pE: map_poly Num.conj_op (map_poly (real_complex F) p) = (map_poly (real_complex F) p). +have pE: map_poly Num.conj_op (map_poly (real_complex F) p) + = (map_poly (real_complex F) p). by rewrite -map_poly_comp; apply/eq_map_poly => z/=; rewrite oppr0. rewrite -[\prod_(_ <- _ | _ && ~~ _) _]big_filter. -have iE: perm_eq [seq i <- dec_roots (map_poly (real_complex F) p) | complex.Im i != 0 & ~~ (0 < complex.Im i)] [seq x^* | x <- [seq i <- dec_roots (map_poly (real_complex F) p) | complex.Im i != 0 & (0 < complex.Im i)]]. +have iE: perm_eq + [seq i <- dec_roots (map_poly (real_complex F) p) | + complex.Im i != 0 & ~~ (0 < complex.Im i)] + [seq x^* | + x <- [seq i <- dec_roots (map_poly (real_complex F) p) | + complex.Im i != 0 & (0 < complex.Im i)]]. apply/uniq_perm. - exact/filter_uniq/uniq_dec_roots. - rewrite map_inj_uniq; last exact/(inv_inj conjCK). @@ -1655,7 +1663,8 @@ Lemma open_subspace_setT (T : topologicalType) (A : set T) : open (A : set (subspace setT)) = open A. Proof. rewrite !openE/=; congr (A `<=` _)%classic. -by apply/seteqP; split; apply/subsetP => x; rewrite /interior !inE nbhs_subspaceT. +by apply/seteqP; split; apply/subsetP => x; + rewrite /interior !inE nbhs_subspaceT. Qed. Lemma open_bigcap (T : topologicalType) (I : Type) (r : seq I) (P : pred I) @@ -1790,8 +1799,7 @@ Lemma horner_continuous {K : numFieldType} (p : {poly K}) : Proof. apply/(eq_continuous (f:=fun x : K => \sum_(i < size p) p`_i * x ^+ i)) => x. by rewrite -[p in RHS]coefK horner_poly. -apply/(@continuous_sum K - (GRing_regular__canonical__normedtype_PseudoMetricNormedZmod K)). +apply/(@continuous_sum _ K^o). move=> /= i _. apply/continuousM; first exact/cst_continuous. exact/continuousX/id_continuous. @@ -1804,24 +1812,24 @@ apply/(eq_continuous (f:=fun x : 'rV[K]_n => \sum_(m <- msupp p) p@_m * \prod_i x ord0 i ^+ m i)) => x. exact/mevalE. -apply/(@continuous_sum K - (GRing_regular__canonical__normedtype_PseudoMetricNormedZmod K)). +apply/(@continuous_sum K K^o). move=> /= i _. apply/continuousM; first exact/cst_continuous. apply/continuous_prod => j _. exact/continuousX/coord_continuous. Qed. -Lemma mx_continuous (T : topologicalType) (K : realFieldType) m n (f : T -> 'M[K]_(m.+1, n.+1)) : +Lemma mx_continuous (T : topologicalType) (K : realFieldType) m n + (f : T -> 'M[K]_(m.+1, n.+1)) : (forall i j, continuous (fun x => f x i j)) -> continuous f. Proof. move=> fc x; apply/cvg_ballP => e e0. near=> y. -rewrite -[X in X (f x)]ball_normE/= [X in X < _]mx_normrE bigmax_lt//= => -[] i j _. -rewrite !mxE/=. +rewrite -[X in X (f x)]ball_normE/= [X in X < _]mx_normrE bigmax_lt//=. +move=> -[] i j _; rewrite !mxE/=. suff: ball (f x i j) e (f y i j). - by rewrite -(@ball_normE _ (GRing_regular__canonical__normedtype_PseudoMetricNormedZmod K)). + by rewrite -(@ball_normE _ K^o). move: i j. near: y. apply: filter_forall => i. @@ -1914,7 +1922,8 @@ have ifE (x : {poly {mpoly S[n]}}): (if (i < size x)%N then x`_i else 0) = x`_i. by rewrite 3!ifE coefB mwidenB mulrBl. Qed. -Lemma mcoeff_muni (A : ringType) (p : {mpoly A[n.+1]}) (i : nat) (m : 'X_{1.. n}) : +Lemma mcoeff_muni (A : ringType) (p : {mpoly A[n.+1]}) + (i : nat) (m : 'X_{1.. n}) : (muni p)`_i@_m = p@_(Multinom (rcons_tuple m i)). Proof. rewrite /muni coef_sum/= raddf_sum/=. @@ -1926,7 +1935,8 @@ rewrite (big_neq_0 (i:=Multinom (rcons_tuple m i))) => //; first last. + exact/index_enum_uniq. rewrite mem_index_enum/= coefC eqxx mcoeffM. have m0: (@mdeg n 0 < (mdeg m).+1)%N by rewrite mdeg0. - rewrite (big_neq_0 (i:=(BMultinom m0, BMultinom (leqnn (mdeg m).+1)))) => //=; first last. + rewrite (big_neq_0 + (i:=(BMultinom m0, BMultinom (leqnn (mdeg m).+1)))) => //=; first last. + case=> /= ul ur /eqP mE; rewrite xpair_eqE !bmeqP/= mcoeffC. move: mE; have [->|_ _ _] := eqVneq (bmnm ul) 0%MM. by rewrite add0m => <-; rewrite eqxx. @@ -1976,13 +1986,17 @@ by rewrite coefXn eqxx mulr1 mcoeffX eqxx mulr1. Qed. Lemma mcoeff_mwiden (A : ringType) (p : {mpoly A[n]}) (m : 'X_{1.. n.+1}) : - (mwiden p)@_m = p@_(Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i)))) *+ (m ord_max == 0). + (mwiden p)@_m + = p@_(Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i)))) + *+ (m ord_max == 0). Proof. rewrite /mwiden/= /mmap raddf_sum/=. -rewrite (big_neq_0 (i:=(Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i)))))); first last. +rewrite (big_neq_0 + (i:=(Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i)))))); first last. - move=> /= u _ um. rewrite mul_mpolyC mcoeffZ /mmap1. - have ->: (\prod_(i < n) 'X_(widen_ord (leqnSn n) i) ^+ u i) = 'X_[Multinom (rcons_tuple u 0)] :> {mpoly A[n.+1]}. + have ->: (\prod_(i < n) 'X_(widen_ord (leqnSn n) i) ^+ u i) + = 'X_[Multinom (rcons_tuple u 0)] :> {mpoly A[n.+1]}. rewrite [RHS]mpolyXE_id big_ord_recr/= multinomE (tnth_nth 0)/= -cats1. rewrite nth_cat size_tuple ltnn subnn/= expr0 mulr1. apply/eq_bigr => i _. @@ -2031,8 +2045,10 @@ apply/anti_leq/andP; split=> //. by move: (ltn_ord i); rewrite ltnS. Qed. -Lemma mcoeff_mmulti (A : ringType) (p : {poly {mpoly A[n]}}) (m : 'X_{1.. n.+1}) : - (mmulti p)@_m = p`_(m ord_max)@_(Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i)))). +Lemma mcoeff_mmulti (A : ringType) (p : {poly {mpoly A[n]}}) + (m : 'X_{1.. n.+1}) : + (mmulti p)@_m + = p`_(m ord_max)@_(Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i)))). Proof. rewrite raddf_sum/=. rewrite (big_ord_iota _ _ xpredT @@ -2059,7 +2075,9 @@ rewrite (big_neq_0 (i:=m ord_max)); last first. rewrite mem_iota/=; case: (ltnP (m ord_max) (size p)) => [mp|]; last first. by move=> /leq_sizeP/(_ (m ord_max) (leqnn _)) ->; rewrite mcoeff0. rewrite mcoeffM. -have mm: (mdeg (mnmwiden (Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i))))) < (mdeg m).+1)%N. +have mm: + (mdeg (mnmwiden (Multinom (mktuple (fun i => m (widen_ord (leqnSn n) i))))) + < (mdeg m).+1)%N. rewrite ltnS !mdegE; apply/leq_sum => /= i _. rewrite mnmwidenE; case: (ltnP i n) => [ilt|//]. by rewrite mnmE; move: (leqnn (m i)); congr (m _ <= _)%N; apply/val_inj. @@ -2082,7 +2100,8 @@ have mE: m = rewrite (big_neq_0 (i:=(BMultinom mm, BMultinom mn))). - by rewrite /= -mE eqxx mem_index_enum/= mwiden_mnmwiden mcoeffXn eqxx mulr1. - exact/index_enum_uniq. -case=> /= l r /eqP; rewrite mcoeffXn xpair_eqE -![_ == BMultinom _](inj_eq val_inj). +case=> /= l r /eqP; rewrite mcoeffXn xpair_eqE + -![_ == BMultinom _](inj_eq val_inj). have [<-|i _ _] := eqVneq (U_(ord_max) *+ m ord_max)%MM r; last exact/mulr0. move=> mlr /negP; rewrite andbT; elim. apply/eqP/(@addIm _ (U_(ord_max) *+ m ord_max)%MM). diff --git a/continuity_roots.v b/continuity_roots.v index 96602be..3704d7c 100644 --- a/continuity_roots.v +++ b/continuity_roots.v @@ -1,9 +1,11 @@ From HB Require Import structures. -From mathcomp Require Import ssreflect ssrfun ssrbool eqtype choice seq ssrnat bigop fintype tuple order ssralg ssrnum poly polydiv complex polyorder. -From mathcomp Require Import matrix topology normedtype signed classical_sets. +From mathcomp Require Import ssreflect ssrfun ssrbool eqtype choice seq ssrnat. +From mathcomp Require Import bigop fintype tuple order ssralg ssrnum poly. +From mathcomp Require Import polydiv complex polyorder matrix topology. +From mathcomp Require Import normedtype signed classical_sets. -Import Order.POrderTheory Order.TotalTheory GRing.Theory Num.Theory Num.Def complex. -Import numFieldTopology.Exports. +Import Order.POrderTheory Order.TotalTheory GRing.Theory Num.Theory Num.Def. +Import complex numFieldTopology.Exports. Require Import auxresults. @@ -26,7 +28,9 @@ Definition deformp eps (p q : {poly R[i]}) := (size q <= size p)%N /\ forall i : 'I_(size p), `|p`_i - q`_i| < eps. Lemma close_root_deformp (eps : R[i]) (p : {poly R[i]}) : 0 < eps -> - exists delta : R[i], 0 < delta /\ forall q, deformp delta p q -> forall x, root p x -> exists y, root q y /\ `|x - y| < eps. + exists delta : R[i], 0 < delta /\ + forall q, deformp delta p q -> + forall x, root p x -> exists y, root q y /\ `|x - y| < eps. Proof. wlog : eps / eps <= 1 => [H|eps1 eps0]. case: eps => eR eI. @@ -40,7 +44,8 @@ wlog : eps / eps <= 1 => [H|eps1 eps0]. exists y; split=> //; apply/(lt_le_trans xy). by rewrite lecE /minr -(fun_if (fun x => x%:C))/= ltcE/= eqxx/= ge_min lexx. case sp: (size p) => [|n]. - exists 1; split=> [//|] q [sq] _ x _; exists x; split; last by rewrite subrr normr0. + exists 1; split=> [//|] q [sq] _ x _; exists x; split; last first. + by rewrite subrr normr0. by move: sq; rewrite sp size_poly_leq0 => /eqP ->; apply/root0. pose M := (\big[maxr/1]_(x <- dec_roots p) Re `|x|)%:C. have M1 : 1 <= M. @@ -58,11 +63,11 @@ split=> [|q [sq]]. rewrite sp => qcoef; move: (qcoef ord_max) => /(le_lt_trans (lerB_dist _ _)). rewrite ltrBlDl -ltrBlDr -[X in X - _]mulr1 -mulrBr. have: 1 / 2 <= (1 - (eps / M) ^+ n / (n.+1%:R *+ 4)). - rewrite ler_pdivrMr ?ltr0n// mulrBl lerBrDr -lerBrDl [1 * 2]mulr_natr -{2}[2]natr1 -addrA subrr addr0. - - rewrite -[2]divr1 mulf_div mulr1. - rewrite -[4%N]/(2 * 2)%N mulrnA -[X in _ / X]mulr_natr -mulf_div divff ?pnatr_eq0// mulr1. - rewrite -mulr_natr -natrM ler_pdivrMr ?ltr0n ?muln_gt0// mul1r. + rewrite ler_pdivrMr ?ltr0n// mulrBl lerBrDr -lerBrDl [1 * 2]mulr_natr. + rewrite -{2}[2]natr1 -addrA subrr addr0 -[2]divr1 mulf_div mulr1. + rewrite -[4%N]/(2 * 2)%N mulrnA -[X in _ / X]mulr_natr -mulf_div. + rewrite divff ?pnatr_eq0// mulr1 -mulr_natr -natrM. + rewrite ler_pdivrMr ?ltr0n ?muln_gt0// mul1r. have le1: (eps / M) ^+ n <= 1. apply/exprn_ile1. apply/divr_ge0; first exact/ltW. @@ -71,15 +76,16 @@ have: 1 / 2 <= (1 - (eps / M) ^+ n / (n.+1%:R *+ 4)). exact/(le_trans eps1). by apply/(le_trans le1); rewrite ler1n muln_gt0. move=> /(ler_pM _ _ (lexx (normr p`_ord_max))) => /(_ n (normr_ge0 _)). -rewrite div1r invr_ge0 => /(_ (ler0n _ _)) => le2 /(le_lt_trans le2) {le2} /= pqn x px. +rewrite div1r invr_ge0 => /(_ (ler0n _ _)). +move=> le2 /(le_lt_trans le2) {le2} /= pqn x px. have n0: (0 < n)%N. rewrite lt0n; apply/eqP => n0; move: sp; rewrite n0 => ps. have /size1_polyC pE : (size p <= 1)%N by rewrite ps. by move: px ps; rewrite pE rootC => /eqP ->; rewrite size_polyC eqxx. have qn0: q`_n != 0. apply/eqP => qn. - move: pqn; rewrite qn normr0 ltr_pdivrMr ?ltr0n// mul0r => /(le_lt_trans (normr_ge0 _)). - by rewrite ltxx. + move: pqn; rewrite qn normr0 ltr_pdivrMr ?ltr0n// mul0r. + by move=> /(le_lt_trans (normr_ge0 _)); rewrite ltxx. have {}sq : size q = n.+1. apply/eqP; rewrite eqn_leq; apply/andP; split; first by rewrite -sp. exact/gt_size. @@ -87,7 +93,8 @@ case: (closed_field_poly_normal q) => /= r; rewrite lead_coefE sq -pred_Sn => qE move: (sq); rewrite qE size_scale// size_prod_seq; last first. by move=> i _; rewrite polyXsubC_eq0. under eq_bigr do rewrite size_XsubC. -rewrite big_const_seq count_predT iter_addn_0 subSn ?leq_pmull// mul2n -addnn subDnCA// subnn addn0 => /eqP. +rewrite big_const_seq count_predT iter_addn_0 subSn ?leq_pmull// mul2n -addnn. +rewrite subDnCA// subnn addn0 => /eqP. rewrite eqSS => /eqP sr. pose m := (\big[Order.min/Re `|x - head 0 r|]_(z <- r) Re `|x - z|)%:C. have m0: 0 <= m. @@ -99,8 +106,8 @@ have: `|p`_n| / 2 * m ^+ n <= `|q.[x]|. - apply/divr_ge0; [apply/normr_ge0|apply/ler0n]. - exact/exprn_ge0. - exact/ltW. - rewrite -sr -prodr_const_seq big_seq [X in _ <= X]big_seq; apply/ler_prod => y yr. - apply/andP; split=> //. + rewrite -sr -prodr_const_seq big_seq [X in _ <= X]big_seq. + apply/ler_prod => y yr; apply/andP; split=> //. rewrite !hornerE -[`|x - y|]RRe_real ?normr_real// /m lecR. rewrite bigmin_le; apply/orP; right; apply/hasP; exists y => //=. rewrite -[q.[x]]subr0 -{2}(rootP px) distrC -hornerN -hornerD. @@ -108,7 +115,8 @@ rewrite -[p - q]coefK horner_poly => mle. move: (le_trans mle (ler_norm_sum _ _ _)). under eq_bigr do rewrite normrM normrX coefD coefN; move=> {}mle. have: normr (p`_n) / 2 * m ^+ n <= - \sum_(i < size (p - q)) normr p`_n * ((eps / M) ^+ n / (n.+1%:R *+ 4)) * M ^+ n. + \sum_(i < size (p - q)) + normr p`_n * ((eps / M) ^+ n / (n.+1%:R *+ 4)) * M ^+ n. apply/(le_trans mle)/ler_sum; case=> i/= ipq _. have ilt : (i < n.+1)%N. rewrite -[n.+1]maxnn -{1}sp -sq -[size q]size_opp. @@ -121,7 +129,8 @@ have: normr (p`_n) / 2 * m ^+ n <= apply/ler_pM. - exact/exprn_ge0/normr_ge0. - exact/ler01. - - rewrite -[M ^+ i]mul1r -ler_pdivrMr; last exact/exprn_gt0/(lt_le_trans ltr01). + - rewrite -[M ^+ i]mul1r -ler_pdivrMr; last first. + exact/exprn_gt0/(lt_le_trans ltr01). rewrite -expr_div_n; apply/exprn_ile1. by apply/divr_ge0; [apply/normr_ge0|apply/(le_trans ler01)]. rewrite ler_pdivrMr ?mul1r /M; last exact/(lt_le_trans ltr01). @@ -129,26 +138,32 @@ have: normr (p`_n) / 2 * m ^+ n <= rewrite le_bigmax; apply/orP; right; apply/hasP; exists x => //=. by rewrite mem_dec_roots -size_poly_leq0 sp. - exact/exprn_ege1. -rewrite sumr_const card_ord -[(_ * _) *+ _]mulr_natr -!mulrA -subr_ge0 -mulrBr pmulr_rge0; last first. -rewrite normr_gt0; apply/eqP => pn. +rewrite sumr_const card_ord -[(_ * _) *+ _]mulr_natr -!mulrA -subr_ge0 -mulrBr. +rewrite pmulr_rge0; last first. + rewrite normr_gt0; apply/eqP => pn. suff: (size p <= n)%N by rewrite sp ltnn. apply/leq_sizeP => i; rewrite leq_eqVlt => /orP[/eqP <- //|]. by rewrite -sp => /leq_sizeP/(_ i (leqnn i)). -rewrite subr_ge0 mulrC expr_div_n -[_ *+ 4]mulr_natr [_^-1 * _]mulrC [_ * 4]mulrC -mulf_div. -rewrite [X in _ <= X]mulrA mulf_div [_ * 4]mulrC -mulf_div divff; last first. - exact/expf_neq0/lt0r_neq0/(lt_le_trans ltr01). +rewrite subr_ge0 mulrC expr_div_n -[_ *+ 4]mulr_natr [_^-1 * _]mulrC. +rewrite [_ * 4]mulrC -mulf_div [X in _ <= X]mulrA mulf_div [_ * 4]mulrC. +rewrite -mulf_div divff; last exact/expf_neq0/lt0r_neq0/(lt_le_trans ltr01). rewrite mulr1 => {}mle. -have /(le_trans mle) {}mle: eps ^+ n / 4 * ((size (p - q))%:R / n.+1%:R) <= eps ^+ n / 4. +have /(le_trans mle) {}mle: + eps ^+ n / 4 * ((size (p - q))%:R / n.+1%:R) <= eps ^+ n / 4. rewrite -[X in _ <= X]mulr1; apply/ler_pM. - apply/mulr_ge0; first exact/exprn_ge0/ltW. by rewrite invr_ge0. - exact/divr_ge0. - exact/lexx. - by rewrite mulrC ler_pdivrMl ?ltr0n// mulr1 ler_nat -(maxnn n.+1) -{1}sp -sq -(size_opp q) size_add. + rewrite mulrC ler_pdivrMl ?ltr0n// mulr1 ler_nat -(maxnn n.+1) -{1}sp -sq. + by rewrite -(size_opp q) size_add. have /(le_lt_trans mle) : eps ^+ n / 4 < eps ^+ n / 2. - by rewrite mulrC ltr_pdivrMl ?ltr0n// -[4%N]/((2 * 2)%N) natrM mulrACA divff ?pnatr_eq0// mulr1 mulrC mulr_natr mulr2n -subr_gt0 -addrA subrr addr0 exprn_gt0. -rewrite -subr_gt0 -(@pmulr_rgt0 _ 2%:R)// mulrBr subr_gt0 mulrCA divff ?pnatr_eq0// mulr1. -rewrite mulrCA divff ?pnatr_eq0// -ltr_pdivrMl ?exprn_gt0// mulrC -expr_div_n expr_lt1//; last first. + rewrite mulrC ltr_pdivrMl ?ltr0n// -[4%N]/((2 * 2)%N) natrM mulrACA. + rewrite divff ?pnatr_eq0// mulr1 mulrC mulr_natr mulr2n -subr_gt0 -addrA. + by rewrite subrr addr0 exprn_gt0. +rewrite -subr_gt0 -(@pmulr_rgt0 _ 2%:R)// mulrBr subr_gt0 mulrCA. +rewrite divff ?pnatr_eq0// mulr1 mulrCA divff ?pnatr_eq0//. +rewrite -ltr_pdivrMl ?exprn_gt0// mulrC -expr_div_n expr_lt1//; last first. by apply/mulr_ge0 => //; rewrite invr_ge0 ltW. rewrite mulrC ltr_pdivrMl// mulr1 /m -[eps]RRe_real ?gtr0_real// ltcR. rewrite bigmin_lt; case: r {qE m m0 mle} sr n0 => [<- //|] y r sr n0/=. @@ -159,7 +174,10 @@ by rewrite root_XsubC. Qed. Lemma rm_root_poly (p : {poly R[i]}) (x : R[i]) : - x != 0 -> root p x -> p %/ ('X - x%:P) = \poly_(i < (size p).-1) (- x ^- (i.+1) * \sum_(j < i.+1) p`_j * x ^+ j). + x != 0 -> + root p x -> + p %/ ('X - x%:P) + = \poly_(i < (size p).-1) (- x ^- (i.+1) * \sum_(j < i.+1) p`_j * x ^+ j). Proof. move=> x0 /factor_theorem [q] ->. have X0 : ('X - x%:P != 0) by rewrite polyXsubC_eq0. @@ -170,20 +188,28 @@ rewrite size_mul// size_XsubC addn2 -!pred_Sn. apply/polyP => i; rewrite coef_poly. case: (ltnP i (size q)) => [|/leq_sizeP/(_ i (leqnn i))//] iq. rewrite mulNr -mulrN -sumrN. -under eq_bigr do rewrite mulrBr coefD coefN mulrBl coefMC -mulrA -exprS opprB coefMX {1}[X in q`_X]pred_Sn. +under eq_bigr do rewrite mulrBr coefD coefN mulrBl coefMC -mulrA -exprS opprB. +under eq_bigr do rewrite coefMX {1}[X in q`_X]pred_Sn. pose f i := (if i == 0%N then 0 else q`_i.-1) * x ^+ i. -rewrite -(big_mkord xpredT (fun i => f (i.+1) - f i)) telescope_sumr// /f/= mul0r subr0 mulrCA [X in _ * X]mulrC divff ?mulr1//. +rewrite -(big_mkord xpredT (fun i => f (i.+1) - f i)) telescope_sumr// /f/=. +rewrite mul0r subr0 mulrCA [X in _ * X]mulrC divff ?mulr1//. exact/expf_neq0. Qed. -Lemma close_rm_root (eps : R[i]) (p : {poly R[i]}) (xp : R[i]) : 0 < eps -> xp != 0 -> root p xp -> - exists delta, 0 < delta /\ forall (q : {poly R[i]}) (xq : R[i]), root q xq -> deformp delta p q -> `|xp - xq| < delta -> deformp eps (p %/ ('X - xp%:P)) (q %/ ('X - xq%:P)). +Lemma close_rm_root (eps : R[i]) (p : {poly R[i]}) (xp : R[i]) : + 0 < eps -> xp != 0 -> root p xp -> + exists delta, 0 < delta /\ + forall (q : {poly R[i]}) (xq : R[i]), + root q xq -> deformp delta p q -> `|xp - xq| < delta -> + deformp eps (p %/ ('X - xp%:P)) (q %/ ('X - xq%:P)). Proof. move=> e0 xp0 pxp /=. have [->|] := poly0Vpos p. - by exists 1; split=> [//|] q xq qxq []; rewrite size_poly0 => /size_poly_leq0P -> _ _; rewrite !div0p; split=> //; case; rewrite size_poly0. + exists 1; split=> [//|] q xq qxq []; rewrite size_poly0 => /size_poly_leq0P. + by move=> -> _ _; rewrite !div0p; split=> //; case; rewrite size_poly0. move sp: (size p) => n; case: n sp => // n sp _. -pose f := fun i (x : 'rV[R[i]^o]_n.+1 * (R[i]^o)) => - (x.2 ^- i.+1) * \sum_(j < n.+1 | (j <= i)%N) x.1 ord0 j * x.2 ^+ j. +pose f := fun i (x : 'rV[R[i]^o]_n.+1 * (R[i]^o)) => + - (x.2 ^- i.+1) * \sum_(j < n.+1 | (j <= i)%N) x.1 ord0 j * x.2 ^+ j. have cont : forall i, {for (\row_(i < n.+1) p`_i, xp), continuous (f i)}. move=> i /=. apply/(@continuousM R[i] ('rV[R[i]^o]_n.+1 * (R[i]^o))%type). @@ -191,16 +217,19 @@ have cont : forall i, {for (\row_(i < n.+1) p`_i, xp), continuous (f i)}. apply/continuousV; first by rewrite expf_eq0. apply/continuousX. apply/cvg_snd. - apply (@continuous_sum R[i] R[i]^o ('rV[R[i]^o]_n.+1 * (R[i]^o))%type) => j ji. + apply/(@continuous_sum R[i] R[i]^o ('rV[R[i]^o]_n.+1 * (R[i]^o))%type). + move=> j ji. apply/(@continuousM R[i] ('rV[R[i]^o]_n.+1 * (R[i]^o))%type); last first. exact/continuousX/cvg_snd. - apply/(@eq_continuous_at ('rV[R[i]^o]_n.+1 * (R[i]^o))%type _ ((fun x : 'rV[R[i]^o]_n.+1 => x ord0 j) \o fst)) => //. + apply/(@eq_continuous_at ('rV[R[i]^o]_n.+1 * (R[i]^o))%type _ + ((fun x : 'rV[R[i]^o]_n.+1 => x ord0 j) \o fst)) => //. apply/continuous_comp; first exact/cvg_fst. exact/coord_continuous. -have /fin_all_exists /=: forall i : 'I_n.+1, exists delta, 0 < delta /\ forall (q : {poly R[i]}) (xq : R[i]), - deformp delta p q -> - normr (xp - xq) < delta -> - `|f i (\row_(i < n.+1) p`_i, xp) - f i (\row_(i < n.+1) q`_i, xq)| < eps. +have /fin_all_exists /=: forall i : 'I_n.+1, + exists delta, 0 < delta /\ + forall (q : {poly R[i]}) (xq : R[i]), + deformp delta p q -> normr (xp - xq) < delta -> + `|f i (\row_(i < n.+1) p`_i, xp) - f i (\row_(i < n.+1) q`_i, xq)| < eps. move=> i. move/(_ i): cont => /= /cvgr_dist_lt /(_ eps e0) [][]/= Vp Vx []/=. move=> /nbhs_ballP [dp/= +] dVp /nbhs_ballP [dx/= +] dVx Vpx. @@ -225,7 +254,8 @@ split=> [|q xq qxq [] spq]. apply/leq_sizeP => j; rewrite leq_eqVlt => /orP[/eqP <- //|nj]. by have/leq_sizeP/(_ j nj): (size p <= n.+1)%N by rewrite sp. rewrite sp => dpq. -rewrite -(RRe_real (normr_real _)) ltcR lt_bigmin lt_min -!ltcR !(RRe_real (normr_real _)) => /andP[] /andP[xpq] xqpn /allP xdelta. +rewrite -(RRe_real (normr_real _)) ltcR lt_bigmin lt_min -!ltcR. +rewrite !(RRe_real (normr_real _)) => /andP[] /andP[xpq] xqpn /allP xdelta. rewrite /deformp !size_divp ?polyXsubC_eq0// !size_XsubC/= !subn1. split. move: spq; rewrite sp succnK -[(_ <= n)%N]ltnS. @@ -236,7 +266,8 @@ rewrite rm_root_poly// coef_poly sp ltn_ord rm_root_poly//; last first. rewrite coef_poly. have sq: size q = size p. apply/anti_leq; rewrite spq/= sp; apply/gt_size/eqP => qn0. - by move/(_ ord_max): dpq; rewrite -(RRe_real (normr_real _)) ltcR lt_bigmin lt_min qn0 subr0 ltxx andbF. + move/(_ ord_max): dpq; rewrite -(RRe_real (normr_real _)) ltcR lt_bigmin. + by rewrite lt_min qn0 subr0 ltxx andbF. rewrite sq sp ltn_ord. move/(_ (lift ord_max i)): deltaP => [d0] /(_ q xq) /=. have /[swap]/[apply]: deformp (delta (lift ord_max i)) p q. @@ -251,10 +282,15 @@ have /[swap]/[apply]: normr (xp - xq) < delta (lift ord_max i). rewrite /lift /= /bump leqNgt ltn_ord add0n /f/=. congr (`|_ * _ - _ * _| < _); under eq_bigr do rewrite mxE. - rewrite (big_ord_iota _ n.+1 (fun i0 => (i0 <= i)%N) (fun i => p`_i * xp ^+ i)) -big_filter -{1}[i : nat]add0n filter_iota_leq; last by apply/ltnW; rewrite ltnS. + rewrite (big_ord_iota _ n.+1 (fun i0 => (i0 <= i)%N) + (fun i => p`_i * xp ^+ i)). + rewrite -big_filter -{1}[i : nat]add0n filter_iota_leq. + by rewrite -big_ord_iota. + by rewrite ltnS; apply/ltnW. +rewrite (big_ord_iota _ n.+1 (fun i0 => (i0 <= i)%N) (fun i => q`_i * xq ^+ i)). +rewrite -big_filter -{1}[i : nat]add0n filter_iota_leq. by rewrite -big_ord_iota. -rewrite (big_ord_iota _ n.+1 (fun i0 => (i0 <= i)%N) (fun i => q`_i * xq ^+ i)) -big_filter -{1}[i : nat]add0n filter_iota_leq; last by apply/ltnW; rewrite ltnS. -by rewrite -big_ord_iota. +by rewrite ltnS; apply/ltnW. Qed. Lemma deformpW (e e' : R[i]) (p q : {poly R[i]}) : @@ -262,7 +298,8 @@ Lemma deformpW (e e' : R[i]) (p q : {poly R[i]}) : Proof. by move=> ee [spq pqe]; split=> // i; apply/(lt_le_trans (pqe i)). Qed. Lemma aligned_deformed (eps : R[i]) (p : {poly R[i]}) : - 0 < eps -> exists delta, 0 < delta /\ forall q, deformp delta p q -> alignp eps p q. + 0 < eps -> + exists delta, 0 < delta /\ forall q, deformp delta p q -> alignp eps p q. Proof. wlog : eps / eps < 1 => [H|e1 e0]. case: eps => eR eI. @@ -272,7 +309,8 @@ wlog : eps / eps < 1 => [H|e1 e0]. - rewrite ltcR lt_min e0/=; apply/mulr_gt0; first exact: ltr01. by rewrite invr_gt0 ltr0n. exists delta; split=> // q /dP pq i /pq ple. - apply/(leq_trans ple); rewrite complexr0 big_mkcond [X in (_ <= X)%N]big_mkcond/=. + apply/(leq_trans ple). + rewrite complexr0 big_mkcond [X in (_ <= X)%N]big_mkcond/=. apply/leq_sum => x _; rewrite -(RRe_real (normr_real _)) !ltcR lt_min andbC. by case: (Re `|x - i| < 1 / 2)%R. have [->|sp] := poly0Vpos p. @@ -284,11 +322,14 @@ elim: n p => /= [|n IHn] p sp _. have p0: p != 0 by apply/eqP => p0; move: sp; rewrite p0 size_poly0. case/boolP: (all (fun x => x == 0) (dec_roots p)) => [root0|/allPn[x]]. have r0: dec_roots p = [:: 0]. - case: (dec_roots p) (uniq_dec_roots p) (mem_dec_roots p) root0 => [|x r] ru memr /allP r0. + case: (dec_roots p) (uniq_dec_roots p) (mem_dec_roots p) root0 + => [|x r] ru memr /allP r0. have /closed_rootP [x]: size p != 1 by rewrite sp. by move: memr; rewrite p0/= => <-. move: (r0 x); rewrite in_cons eqxx/= => /(_ isT) /eqP x0; rewrite x0. - by case: r ru {memr} r0 => // y r /= /andP[+ _] /(_ y); rewrite !in_cons negb_or eqxx orbT/= x0 eq_sym => /andP[/negP y0 _] /(_ isT). + case: r ru {memr} r0 => // y r /= /andP[+ _] /(_ y). + rewrite !in_cons negb_or eqxx orbT/= x0 eq_sym. + by move=> /andP[/negP y0 _] /(_ isT). move: (dec_roots_closedP p). rewrite r0 big_seq1 subr0 => pE. have pn: p`_(size p).-1 != 0 by rewrite -lead_coefE lead_coef_eq0. @@ -309,7 +350,8 @@ case/boolP: (all (fun x => x == 0) (dec_roots p)) => [root0|/allPn[x]]. move: (ltn_ord i); rewrite ltn_neqAle => /andP[] /negPf -> _. by rewrite mulr0 add0r normrN. move=> x; rewrite rootE -rootE rootZ// rootE !hornerE expf_eq0/= => /eqP ->. - rewrite mu_mulC// mu_exp -['X]subr0 mu_XsubC eqxx mul1n big_mkcond big_seq big_mkcond/=. + rewrite mu_mulC// mu_exp -['X]subr0 mu_XsubC eqxx mul1n big_mkcond big_seq. + rewrite big_mkcond/=. have da2 : d < `|a| / 2. rewrite /d -muln2 natrM -mulf_div -[X in _ < X]mul1r -subr_gt0 -mulrBl. apply/mulr_gt0; last by apply/mulr_gt0 => //; rewrite normr_gt0. @@ -317,12 +359,16 @@ case/boolP: (all (fun x => x == 0) (dec_roots p)) => [root0|/allPn[x]]. apply/(lt_le_trans (y:=1)); last by rewrite ler1n. by rewrite exprn_ilt1// ltW. have da : d < `|a|. - by apply/(lt_le_trans da2); rewrite ler_pdivrMr// ler_peMr// -[1]/(1%:R) ler_nat. + apply/(lt_le_trans da2); rewrite ler_pdivrMr// ler_peMr// -[1]/(1%:R). + by rewrite ler_nat. have aqn : `|a| / 2 < `|q`_n.+1|. move: da2 => /(lt_trans qn)/(le_lt_trans (lerB_dist _ _)). - by rewrite -[_ - _ < _]subr_gt0 opprD opprK addrA -{2}[`|a|]mulr1 -{2}(divff (x:=2))// mulrCA mulr_natl [_ / _ *+ _]mulr2n opprD addrA subrr add0r addrC subr_gt0. + rewrite -[_ - _ < _]subr_gt0 opprD opprK addrA -{2}[`|a|]mulr1. + rewrite -{2}(divff (x:=2))// mulrCA mulr_natl [_ / _ *+ _]mulr2n opprD. + by rewrite addrA subrr add0r addrC subr_gt0. have {sq} : size q = n.+2. - apply/anti_leq/andP; split=> //; rewrite ltnNge; apply/negP => /leq_sizeP /( _ n.+1 (leqnn _)) q0. + apply/anti_leq/andP; split=> //; rewrite ltnNge; apply/negP. + move=> /leq_sizeP /( _ n.+1 (leqnn _)) q0. by move: qn; rewrite q0 subr0 => /(lt_trans da); rewrite ltxx. have [->|q0 sq] := eqVneq q 0; first by rewrite size_poly0. have qn0 : q`_(size q).-1 != 0 by rewrite -lead_coefE lead_coef_eq0. @@ -357,9 +403,11 @@ case/boolP: (all (fun x => x == 0) (dec_roots p)) => [root0|/allPn[x]]. by rewrite normrX ler_p1X// -ltnS. rewrite exprS mulrA -subr_gt0 -mulrBl pmulr_lgt0; last first. exact/(lt_le_trans ltr01)/exprn_ege1. - rewrite subr_gt0 /d -mul2n natrM invrM ?unitfE// [_ / 2]mulrC -!mulrA mulVf// mulr1 mulrA mulrC -subr_gt0 -mulrBl pmulr_lgt0 ?subr_gt0; last first. - by rewrite divr_gt0// normr_gt0. - by move=> /(le_lt_trans y1); rewrite expr_gt1// ?ltW// => /(lt_trans e1); rewrite ltxx. + rewrite subr_gt0 /d -mul2n natrM invrM ?unitfE// [_ / 2]mulrC -!mulrA. + rewrite mulVf// mulr1 mulrA mulrC -subr_gt0 -mulrBl. + rewrite pmulr_lgt0 ?subr_gt0; last by rewrite divr_gt0// normr_gt0. + move=> /(le_lt_trans y1); rewrite expr_gt1// ?ltW// => /(lt_trans e1). + by rewrite ltxx. have: 1 < n.+1.*2%:R * d / `|a| / `|y| ^+ n.+1. rewrite -muln2 natrM -[X in X%:R * _]sum1_ord natr_sum !mulr_suml. move/(congr1 (fun x => `|x / (q`_n.+1 * y ^+ n.+1)|)): yq. @@ -371,13 +419,15 @@ case/boolP: (all (fun x => x == 0) (dec_roots p)) => [root0|/allPn[x]]. by apply/hasP; exists ord0. rewrite subr_gt0 mul1r !normrM normrV; last first. by rewrite unitfE mulf_eq0 negb_or expf_eq0 qn0. - rewrite normrM !normrX [2 * _]mulrC -mulf_div mulrA -subr_gt0 -mulrBl pmulr_lgt0; last first. - by rewrite invr_gt0 exprn_gt0// normr_gt0. + rewrite normrM !normrX [2 * _]mulrC -mulf_div mulrA -subr_gt0 -mulrBl. + rewrite pmulr_lgt0; last by rewrite invr_gt0 exprn_gt0// normr_gt0. rewrite subr_gt0 -2!mulrA ltr_pM//. - move: aqn; rewrite ltr_pdivrMr// -ltr_pdivrMl ?normr_gt0// -ltr_pdivlMr ?normr_gt0// => aqn. + move: aqn; rewrite ltr_pdivrMr// -ltr_pdivrMl ?normr_gt0//. + rewrite -ltr_pdivlMr ?normr_gt0// => aqn. have [->|i0] := (posnP i); first by rewrite expr0 mulr1. by rewrite -[X in _ < X]mulr1 ltr_pM// expr_lt1. - rewrite /d mulrCA divff// mulr1 mulrC -!mulrA divff ?normr_eq0// mulr1 mulrC -expr_div_n expr_gt1//. + rewrite /d mulrCA divff// mulr1 mulrC -!mulrA divff ?normr_eq0// mulr1 mulrC. + rewrite -expr_div_n expr_gt1//. by rewrite ltr_pdivlMr ?normr_gt0// mul1r. exact/divr_ge0/normr_ge0/ltW. rewrite mem_dec_roots => /andP[_] px x0. @@ -388,14 +438,17 @@ have de0 : 0 < (minr (Re eps) (Re d'))%:C. by rewrite ltcR lt_min -!ltcR !(RRe_real (gtr0_real _))// e0. move: (close_root_deformp p de0) => [d''][d''0]d''P. exists (minr (Re d') (minr (Re `|lead_coef p|) (Re d'')))%:C; split=> [|q]. - by rewrite ltcR !lt_min -!ltcR (RRe_real (gtr0_real d'0)) d'0 (RRe_real (gtr0_real d''0)) (RRe_real (normr_real _)) normr_gt0 lead_coef_eq0 p0. + rewrite ltcR !lt_min -!ltcR (RRe_real (gtr0_real d'0)) d'0. + rewrite (RRe_real (gtr0_real d''0)) (RRe_real (normr_real _)) normr_gt0. + by rewrite lead_coef_eq0 p0. have [-> [_]|q0 pq] := eqVneq q 0. rewrite sp => /(_ ord_max); rewrite coef0 subr0 -(RRe_real (normr_real _)). by rewrite ltcR !lt_min lead_coefE sp -pred_Sn ltxx andbF. have /d''P/(_ _ px) [y [qy]] : deformp d'' p q. apply/(deformpW _ pq). by rewrite -{2}(RRe_real (gtr0_real d''0)) lecR !ge_min lexx !orbT. -rewrite -(RRe_real (normr_real _)) ltcR lt_min -!ltcR (RRe_real (normr_real _)) !(RRe_real (gtr0_real _))// => /andP[] xye xyd. +rewrite -(RRe_real (normr_real _)) ltcR lt_min -!ltcR (RRe_real (normr_real _)). +rewrite !(RRe_real (gtr0_real _))// => /andP[] xye xyd. have /(d'P _ _ qy)/(_ xyd)/dP pqe : deformp d' p q. apply/(deformpW _ pq). by rewrite -{2}(RRe_real (gtr0_real d'0)) lecR ge_min lexx. diff --git a/cylinder.v b/cylinder.v index 9b4b4fc..05d38b3 100644 --- a/cylinder.v +++ b/cylinder.v @@ -1,4 +1,7 @@ -From mathcomp Require Import freeg ssreflect ssrfun ssrbool eqtype choice seq ssrnat prime binomial bigop tuple order fintype finfun path ssralg ssrnum ssrint poly matrix finmap mpoly complex interval. +From mathcomp Require Import freeg ssreflect ssrfun ssrbool eqtype choice seq. +From mathcomp Require Import ssrnat prime binomial bigop tuple order fintype. +From mathcomp Require Import finfun path ssralg ssrnum ssrint poly matrix. +From mathcomp Require Import finmap mpoly complex interval. From mathcomp Require Import polydiv polyrcf polyorder qe_rcf qe_rcf_th. (* TODO: the following imports should not be needed after cleanup. *) @@ -13,7 +16,8 @@ Import GRing.Theory Num.Theory Num.Def. Import GRing. Import numFieldTopology.Exports. -Require Import auxresults formula subresultant semialgebraic topology continuity_roots. +Require Import auxresults formula subresultant semialgebraic topology. +Require Import continuity_roots. Local Open Scope type_scope. Local Open Scope classical_set_scope. @@ -25,7 +29,8 @@ Local Open Scope sa_scope. Section CylindricalDecomposition. Variables (R : rcfType). -(* How do I assert that the restriction of a function is continuous, and not that the function is continuous at every point in the set? *) +(* How do I assert that the restriction of a function is continuous, and not + that the function is continuous at every point in the set? *) Fixpoint isCylindricalDecomposition n (S : {fset {SAset R^n}}) := SAset_partition S /\ match n with | 0 => Logic.True | n.+1 => @@ -35,12 +40,14 @@ Fixpoint isCylindricalDecomposition n (S : {fset {SAset R^n}}) := exists m, exists xi : m.-tuple {SAfun R^n -> R^1}, (forall i, {within [set` val s'], continuous (tnth xi i)}) /\ sorted (@SAfun_lt _ _) xi - /\ [fset s in S | SAset_cast n s == val s'] = [fset SAset_cast _ x | x in partition_of_graphs_above (val s') xi] + /\ [fset s in S | SAset_cast n s == val s'] + = [fset SAset_cast _ x | x in partition_of_graphs_above (val s') xi] end. Local Notation isCD := isCylindricalDecomposition. -Lemma isCylindricalDecomposition_restrict n m S (mn : (m <= n)%N) : @isCD n S -> isCD [fset SAset_cast m s | s in S]. +Lemma isCylindricalDecomposition_restrict n m S (mn : (m <= n)%N) : + @isCD n S -> isCD [fset SAset_cast m s | s in S]. Proof. move: (n - m)%N mn (subnKC mn) S => + _ <-; elim=> [|d IHd]. rewrite addn0 => S. @@ -49,21 +56,27 @@ move: (n - m)%N mn (subnKC mn) S => + _ <-; elim=> [|d IHd]. by rewrite SAset_cast_id. rewrite addnS => S /= [_] [/IHd] + _; congr isCD. have md: (m <= m + d)%N by rewrite -{1}[m]addn0 leq_add2l. -by apply/fsetP => s; rewrite -imfset_comp; apply/imfsetP/imfsetP => -[x]/= xS ->; +apply/fsetP => s; rewrite -imfset_comp. +by apply/imfsetP/imfsetP => -[x]/= xS ->; exists x => //; rewrite SAset_cast_trans// geq_min md orbT. Qed. Definition poly_invariant n (p : {mpoly R[n]}) (s : {SAset R^n}) := - {in s &, forall x y, (sgz (meval (tnth (ngraph x)) p) = sgz (meval (tnth (ngraph y)) p))%R}. + {in s &, + forall x y, + (sgz (meval (tnth (ngraph x)) p) = sgz (meval (tnth (ngraph y)) p))%R}. -Definition poly_adapted n p (S : {fset {SAset R^n}}) := forall s : S, poly_invariant p (val s). +Definition poly_adapted n p (S : {fset {SAset R^n}}) := + forall s : S, poly_invariant p (val s). -Definition evalpmp {n} (x : 'rV[R]_n) (P : {poly {mpoly R[n]}}) := map_poly (meval (tnth (ngraph x))) P. +Definition evalpmp {n} (x : 'rV[R]_n) (P : {poly {mpoly R[n]}}) := + map_poly (meval (tnth (ngraph x))) P. Definition truncate (T : ringType) (p : {poly T}) (d : nat) := (\poly_(i < minn d (size p)) p`_i)%R. -Definition truncations n (p : {poly {mpoly R[n]}}) : {fset {poly {mpoly R[n]}}} := +Definition truncations n + (p : {poly {mpoly R[n]}}) : {fset {poly {mpoly R[n]}}} := (fix F p n := match n with | 0 => [fset p] @@ -97,7 +110,9 @@ Qed. Lemma truncationsE n (p : {poly {mpoly R[n]}}) : truncations p = - [fset truncate p d | d in [seq d <- iota 0 (size p).+1 | all (fun i => msize p`_i != 1) (iota d ((size p).+1 - d))]]. + [fset truncate p d | + d in [seq d <- iota 0 (size p).+1 | + all (fun i => msize p`_i != 1) (iota d ((size p).+1 - d))]]. Proof. have t00 k: truncate 0 k = 0 :> {poly {mpoly R[n]}}. by apply/eqP; rewrite truncate_ge_sizeP size_poly0. @@ -144,7 +159,8 @@ apply/orP/imfsetP => /= [[/eqP ->|/imfsetP]|]. rewrite mem_filter truncate_trans -iotanS mem_iota iotanS/= add0n. move=> /andP[] /allP ks kd ->. exists (minn (size p).-1 k) => //. - rewrite mem_filter -!iotanS mem_iota/= add0n ltnS geq_min (ltnW kd) orbT andbT. + rewrite mem_filter -!iotanS mem_iota/= add0n ltnS geq_min (ltnW kd) orbT. + rewrite andbT. apply/allP => i; rewrite mem_iota subnKC geq_min; last first. by rewrite (leq_trans (ltnW kd))// orbT. case: (ltnP i (size p).-1) => /= [ip /andP[] ki id|+ _]; last first. @@ -195,8 +211,12 @@ Ltac mp := Import ordered_qelim.ord. Theorem roots2_on n (P : {poly {mpoly R[n]}}) (d : nat) (s : {SAset R^n}) : - {in s, forall x, size (rootsR (evalpmp x P)) = d} - -> {xi : d.-tuple {SAfun R^n -> R^1} | sorted (@SAfun_lt R n) xi /\ {in s, forall x, [seq (xi : {SAfun R^n -> R^1}) x ord0 ord0 | xi <- xi] = (rootsR (evalpmp x P))}}. + {in s, forall x, size (rootsR (evalpmp x P)) = d} -> + {xi : d.-tuple {SAfun R^n -> R^1} | + sorted (@SAfun_lt R n) xi + /\ {in s, forall x, + [seq (xi : {SAfun R^n -> R^1}) x ord0 ord0 | xi <- xi] + = (rootsR (evalpmp x P))}}. Proof. case: (set0Vmem s) => [-> {s}|[x xs] dE]. exists (mktuple (fun i => SAfun_const n (\row__ i%:R))). @@ -222,23 +242,33 @@ pose G_graph (i : 'I_d) : {SAset R^(n+1)} := [set | \/ (s /\ nquantify n.+1 d Exists ( (\big[And/True]_(j < d.-1) ('X_(n.+1+j) <% 'X_(n.+1+j.+1))) - /\ \big[And/True]_(j < d) (term_mpoly (mmulti P) (fun k => if (k : nat) == n then 'X_(n.+1+j) else 'X_k) == 0) - /\ ('forall 'X_(n.+1+d), (term_mpoly (mmulti P) (fun k => if (k : nat) == n then 'X_(n.+1+d) else 'X_k) == 0) ==> \big[Or/False]_(j < d) ('X_(n.+1+d) == 'X_(n.+1+j))) + /\ \big[And/True]_(j < d) + (term_mpoly (mmulti P) + (fun k => if (k : nat) == n then 'X_(n.+1+j) else 'X_k) == 0) + /\ ('forall 'X_(n.+1+d), + (term_mpoly (mmulti P) + (fun k => if (k : nat) == n then 'X_(n.+1+d) else 'X_k) == 0) + ==> \big[Or/False]_(j < d) ('X_(n.+1+d) == 'X_(n.+1+j))) /\ 'X_n == 'X_(n.+1+i)))]. have GP i (x0 : 'rV[R]_n) (y : 'rV[R]_1) : - row_mx x0 y \in G_graph i = if x0 \in s then y ord0 ord0 == (rootsR (evalpmp x0 P))`_i + row_mx x0 y \in G_graph i + = if x0 \in s + then y ord0 ord0 == (rootsR (evalpmp x0 P))`_i else y == \row__ i%:R. rewrite pi_form /cut rcf_sat_subst. rewrite -[X in subst_env _ X]cats0 subst_env_iota_catl ?size_ngraph//. rewrite !simp_rcf_sat -rcf_sat_take ngraph_cat take_size_cat ?size_ngraph//. - move/(_ x0) : dE; rewrite inE; case/boolP: (rcf_sat (ngraph x0) s) => /= x0s dE; last first. + move/(_ x0) : dE; rewrite inE. + case/boolP: (rcf_sat (ngraph x0) s) => /= x0s dE; last first. rewrite nth_cat size_map size_enum_ord ltnn subnn enum_ordSl/= orbF. apply/eqP/eqP => [y0|/rowP ->]; last by rewrite !mxE. apply/rowP; case; case=> [|//] j0. by rewrite !mxE -[RHS]y0; congr (y _ _); apply/val_inj. move: dE => /(_ isT) /eqP dE. - rewrite -ngraph_cat -[X in nquantify X]addn1 -[X in nquantify X](size_ngraph (row_mx x0 y)). - apply/rcf_satP/eqP => [/nexistsP [z]/= []/holdsAnd zlt []/holdsAnd z0 []z0P|yE]. + rewrite -ngraph_cat -[X in nquantify X]addn1. + rewrite -[X in nquantify X](size_ngraph (row_mx x0 y)). + apply/rcf_satP/eqP => [|yE]. + move=> /nexistsP [z]/= []/holdsAnd zlt []/holdsAnd z0 []z0P. rewrite nth_cat size_map size_enum_ord {1}addn1 leqnn enum_ordD map_cat. rewrite nth_cat 2!size_map size_enum_ord ltnn subnn enum_ordSl/=. rewrite enum_ord0/= nth_cat size_cat 2!size_map size_enum_ord/=. @@ -291,7 +321,8 @@ have GP i (x0 : 'rV[R]_n) (y : 'rV[R]_1) : have: (rootsR (evalpmp x0 P))`_j \in rootsR (evalpmp x0 P). by apply/mem_nth; rewrite (eqP dE). move=> /root_roots/rootP. - rewrite -(mevalC (tnth (ngraph x0)) ((rootsR (evalpmp x0 P))`_j)) horner_map/= => Pj. + rewrite -(mevalC (tnth (ngraph x0)) ((rootsR (evalpmp x0 P))`_j)). + rewrite horner_map/= => Pj. rewrite -[RHS]Pj/= mevalC; apply/meval_eq => k. rewrite (ltn_eqF (ltn_ord k))/= nth_cat size_map size_enum_ord. rewrite -[nat_of_ord k]/(nat_of_ord (@unsplit n 1 (inl k))) ltn_ord. @@ -325,7 +356,8 @@ have GP i (x0 : 'rV[R]_n) (y : 'rV[R]_1) : rewrite -[X in _`_X]addn0 -[X in _`_X]/(nat_of_ord (@unsplit n 1 (inr ord0))). rewrite (nth_map (unsplit (inr ord0))) ?size_enum_ord// nth_enum_ord. by rewrite mxE unsplitK subDnCA// subnn addn0. -have SAfun_G (i : 'I_d) : (G_graph i \in @SAfunc _ n 1) && (G_graph i \in @SAtot _ n 1). +have SAfun_G (i : 'I_d) : + (G_graph i \in @SAfunc _ n 1) && (G_graph i \in @SAtot _ n 1). apply/andP; split. apply/inSAfunc => x0 y1 y2; rewrite !GP. case: (x0 \in s); last by move=> /eqP -> /eqP. @@ -353,11 +385,15 @@ rewrite (nth_map (@SAfun_const R n 1 0)) ?size_tuple//. by rewrite -[ i ]/(nat_of_ord (Ordinal id)) nth_mktuple GE mxE ys. Qed. -Lemma rootsR_continuous n (p : {poly {mpoly R[n]}}) (s : {SAset R^n}) (x : 'rV[R]_n) i : +Lemma rootsR_continuous n (p : {poly {mpoly R[n]}}) (s : {SAset R^n}) + (x : 'rV[R]_n) i : x \in s -> {in s, forall y, size (evalpmp y p) = size (evalpmp x p)} -> - {in s, forall y, size (gcdp (evalpmp y p) (evalpmp y p)^`()) = size (gcdp (evalpmp x p) (evalpmp x p)^`())} -> - {in s, forall y, size (rootsR (evalpmp y p)) = size (rootsR (evalpmp x p))} -> + {in s, forall y, + size (gcdp (evalpmp y p) (evalpmp y p)^`()) + = size (gcdp (evalpmp x p) (evalpmp x p)^`())} -> + {in s, forall y, + size (rootsR (evalpmp y p)) = size (rootsR (evalpmp x p))} -> {within [set` s], continuous (fun y => (rootsR (evalpmp y p))`_i)}. Proof. case: (ltnP i (size (rootsR (evalpmp x p)))) => ir; last first. @@ -365,7 +401,8 @@ case: (ltnP i (size (rootsR (evalpmp x p)))) => ir; last first. apply(@subspace_eq_continuous _ _ R (fun=> 0)); last exact/cst_continuous. move=> /= u; rewrite mem_setE => us. by apply/esym/nth_default; rewrite (r_const u us). -case: n p s x ir => [|n] p s x ir xs s_const s'_const r_const/=; apply/continuousP => /= A; +case: n p s x ir => [|n] p s x ir xs s_const s'_const r_const/=; + apply/continuousP => /= A; rewrite openE/= => /subsetP Aopen; apply/open_subspace_ballP => /= y; rewrite in_setI mem_setE => /andP[] {}/Aopen; @@ -381,24 +418,30 @@ have [p0|px0] := eqVneq (size (evalpmp x p)) 0. pose q z := map_poly (real_complex R) (evalpmp z p). have q0 z : z \in s -> q z != 0. by move=> zs; rewrite map_poly_eq0 -size_poly_eq0 s_const. -set e' := \big[Order.min/e]_(u <- dec_roots (q y)) \big[Order.min/e]_(v <- dec_roots (q y) | u != v) (complex.Re `|u - v| / 2). +set e' := \big[Order.min/e]_(u <- dec_roots (q y)) + \big[Order.min/e]_(v <- dec_roots (q y) | u != v) (complex.Re `|u - v| / 2). have e'0: 0 < e'%:C%C. rewrite ltcR lt_bigmin e0/=; apply/allP => u _. rewrite lt_bigmin e0/=; apply/allP => v _. apply/implyP => uv; apply/divr_gt0; last exact/ltr0Sn. by rewrite -ltcR (normr_gt0 (u - v)) subr_eq0. -have: exists d : R, 0 < d /\ forall z, z \in s -> `|z - y| < d -> alignp e'%:C%C (q y) (q z). +have: exists d : R, 0 < d /\ + forall z, z \in s -> `|z - y| < d -> alignp e'%:C%C (q y) (q z). case: (aligned_deformed (q y) e'0) => /= [[]] a aI []. rewrite ltcE/= => /andP[/eqP ->] a0; rewrite complexr0 => ad. - have /fin_all_exists /=: forall i : 'I_(size (val p)).+1, exists delta, 0 < delta /\ forall (z : 'rV[R]_n.+1), y \in s -> `|y - z| < delta -> `|(q y)`_i - (q z)`_i| < a%:C%C. + have /fin_all_exists /= : forall (i : 'I_(size (val p)).+1), + exists delta, 0 < delta /\ + forall (z : 'rV[R]_n.+1), y \in s -> `|y - z| < delta -> + `|(q y)`_i - (q z)`_i| < a%:C%C. move=> j. move: (@meval_continuous _ _ (val p)`_j y). rewrite /= /continuous_at. - move=> /(@cvgr_dist_lt _ (GRing_regular__canonical__normedtype_PseudoMetricNormedZmod R)). + move=> /(@cvgr_dist_lt _ R^o). move=> /(_ _ a0) /nbhs_ballP[] d'/= d'0 /subsetP xd'. exists d'; split=> // z zs yz. move: xd' => /(_ z); mp; first by rewrite -ball_normE inE/=. - rewrite inE/= !coef_map/= -rmorphB/= normc_def/= expr0n/= addr0 sqrtr_sqr ltcR. + rewrite inE/= !coef_map/= -rmorphB/= normc_def/= expr0n/= addr0 sqrtr_sqr. + rewrite ltcR. by congr (normr (_ - _) < a); apply/meval_eq => k; rewrite tnth_mktuple. move=> [d'] d'P; exists (\big[minr/1]_i d' i). split; first by rewrite lt_bigmin ltr01; apply/allP => j _ /=; case: (d'P j). @@ -412,7 +455,7 @@ move=> [] d [] d0 dP. exists d; split=> // z/=. rewrite -ball_normE/= -opprB normrN => -[] yz zs; apply/yeA. move: dP => /(_ z zs yz) yze. -rewrite -(@ball_normE R (GRing_regular__canonical__normedtype_PseudoMetricNormedZmod R))/=. +rewrite -(@ball_normE R R^o)/=. have: exists (fyz : [fset x in dec_roots (q y)] -> [fset x in dec_roots (q z)]), forall u, `|val u - val (fyz u)| < e'%:C%C. apply/(fin_all_exists (P:=fun u v => `|val u - val v| < e'%:C%C)). @@ -420,7 +463,8 @@ have: exists (fyz : [fset x in dec_roots (q y)] -> [fset x in dec_roots (q z)]), move: yze => /(_ u pu). rewrite -big_filter; case rsy: (seq.filter _ _) => [|v l]. by rewrite big_nil leqn0 mu_eq0 ?pu// map_poly_eq0 -size_poly_eq0 s_const. - move: (mem_head v l); rewrite -rsy mem_filter -normrN opprB => /andP[] uv pv _. + move: (mem_head v l). + rewrite -rsy mem_filter -normrN opprB => /andP[] uv pv _. suff vP: v \in [fset x in dec_roots (q z)]. by exists [` vP]. by rewrite mem_imfset. @@ -428,7 +472,8 @@ move=> [/=] fyz fyze. have eP (u v : [fset x | x in dec_roots (q y)]) : `|val u - val v| < 2 * e'%:C%C -> u = v. move=> uve; apply/eqP/negP => /negP uv; move: uve. - rewrite -(RRe_real (normr_real _)) mulrC mulr_natr -rmorphMn/= ltcR -mulr_natr. + rewrite -(RRe_real (normr_real _)) mulrC mulr_natr -rmorphMn/= ltcR. + rewrite -mulr_natr. rewrite -ltr_pdivrMr; last exact/ltr0Sn. rewrite lt_bigmin => /andP[_] /allP-/(_ (val u))/=. move: (fsvalP u); rewrite (mem_imfset _ _ (@inj_id _))/= => /[swap]/[apply]. @@ -449,7 +494,8 @@ have fyzb: bijective fyz. rewrite size_divp; last by rewrite gcdp_eq0 negb_and q0. rewrite ![(q _)^`()]deriv_map -!gcdp_map !size_map_poly -!/(evalpmp _ _). by rewrite s_const// s_const// s'_const// s'_const. -have pyrP j: (j < size (rootsR (evalpmp y p)))%N -> ((rootsR (evalpmp y p))`_j)%:C%C \in [fset x | x in dec_roots (q y)]. +have pyrP j: (j < size (rootsR (evalpmp y p)))%N -> + ((rootsR (evalpmp y p))`_j)%:C%C \in [fset x | x in dec_roots (q y)]. rewrite (mem_imfset _ _ (@inj_id _))/= mem_dec_roots q0//=. move=> /(mem_nth 0); rewrite in_rootsR => /andP[_] jr. exact/rmorph_root. @@ -461,7 +507,9 @@ suff ->: ((rootsR (evalpmp z p))`_i)%:C%C = val (fyz [` pyrP i ir]). move: (fyze [` pyrP i ir]) => /= pye. apply/(lt_le_trans pye). by rewrite lecR; apply/bigmin_le_id. -have perm_eqC a: perm_eq [seq u <- dec_roots (q a) | u \is Num.real] [seq x%:C%C | x <- rootsR (evalpmp a p)]. +have perm_eqC a: perm_eq + [seq u <- dec_roots (q a) | u \is Num.real] + [seq x%:C%C | x <- rootsR (evalpmp a p)]. apply/uniq_perm. - exact/filter_uniq/uniq_dec_roots. - by rewrite map_inj_uniq ?uniq_roots//; apply/complexI. @@ -500,7 +548,9 @@ have {}fyzr (u : [fset x | x in dec_roots (q y)]) : apply/idP/idP; last exact/fyzr. move=> ur; apply/negP => /negP fur. pose sr y := [fset x : [fset x in dec_roots (q y)] | val x \is Num.real]. - have srE a: [fset val x | x in sr a] = [fset x | x in dec_roots (q a) & x \is Num.real]. + have srE a: + [fset val x | x in sr a] + = [fset x | x in dec_roots (q a) & x \is Num.real]. apply/eqP; rewrite eqEfsubset; apply/andP; split; apply/fsubsetP => b; rewrite (mem_imfset _ _ (@inj_id _))/=. move=> /imfsetP[/=] c /imfsetP[/=] c' cr -> ->. @@ -511,11 +561,14 @@ have {}fyzr (u : [fset x | x in dec_roots (q y)]) : by rewrite mem_imfset. exists [` bP] => //. by rewrite (mem_imfset _ _ (@inj_id _))/=. - suff: (#|` [fset x | x in dec_roots (q z) & x \is Num.real]| < #|` [fset x | x in dec_roots (q y) & x \is Num.real]|)%N. + suff: (#|` [fset x | x in dec_roots (q z) & x \is Num.real]| + < #|` [fset x | x in dec_roots (q y) & x \is Num.real]|)%N. rewrite [X in (X < _)%N]card_imfset//= [X in (_ < X)%N]card_imfset//=. do 2 rewrite undup_id ?uniq_dec_roots//. - rewrite (@perm_size _ _ [seq x%:C%C | x <- rootsR (evalpmp z p)]); last exact/perm_eqC. - rewrite [X in (_ < X)%N](@perm_size _ _ [seq x%:C%C | x <- rootsR (evalpmp y p)]); last exact/perm_eqC. + rewrite (@perm_size _ _ [seq x%:C%C | x <- rootsR (evalpmp z p)]); + last exact/perm_eqC. + rewrite [X in (_ < X)%N](@perm_size _ _ + [seq x%:C%C | x <- rootsR (evalpmp y p)]); last exact/perm_eqC. by rewrite !size_map r_const// r_const// ltnn. rewrite -2!srE [X in (X < _)%N](card_imfset _ _ val_inj)/=. rewrite [X in (_ < X)%N](card_imfset _ _ val_inj)/=. @@ -549,9 +602,11 @@ move: (fiR); rewrite -index_mem => fiRs. rewrite -[RHS](count_lt_nth 0 (sorted_roots _ _ _) fiRs) -!/(rootsR _). rewrite (nth_index 0 fiR). pose sr y z := [fset x : [fset x in dec_roots (q y)] | val x < z]. -have srE a b: [fset val x | x in sr a b] = [fset x | x in dec_roots (q a) & x < b]. +have srE a b: + [fset val x | x in sr a b] + = [fset x | x in dec_roots (q a) & x < b]. apply/eqP; rewrite eqEfsubset; apply/andP; split; apply/fsubsetP => b'; - rewrite (mem_imfset _ _ (@inj_id _))/=. + rewrite (mem_imfset _ _ (@inj_id _))/=. move=> /imfsetP[/=] c /imfsetP[/=] c' cr -> ->. apply/andP; split=> //=. by move: (fsvalP c'); rewrite (mem_imfset _ _ (@inj_id _))/=. @@ -576,15 +631,24 @@ have {}perm_eqC a b: perm_eq rewrite mem_filter in_rootsR => /andP[] vb /andP[] pa0 pv; split. by rewrite ltcR. by rewrite pa0/=; apply/rmorph_root. -suff: (#|` [fset x | x in dec_roots (q y) & (x < ((rootsR (evalpmp y p))`_i)%:C%C)%R]| = #|` [fset x | x in dec_roots (q z) & (x < val (fyz [` pyrP i ir]))%R]|)%N. +suff: (#|` [fset x | x in dec_roots (q y) + & (x < ((rootsR (evalpmp y p))`_i)%:C%C)%R]| + = #|` [fset x | x in dec_roots (q z) + & (x < val (fyz [` pyrP i ir]))%R]|)%N. rewrite [LHS]card_imfset//= [RHS]card_imfset//=. do 2 rewrite undup_id ?uniq_dec_roots//. - rewrite (@perm_size _ _ [seq x%:C%C | x <- [seq x <- rootsR (evalpmp y p) | (x < (rootsR (evalpmp y p))`_i)%R]]); last exact/perm_eqC. + rewrite (@perm_size _ _ [seq x%:C%C | x <- + [seq x <- rootsR (evalpmp y p) | (x < (rootsR (evalpmp y p))`_i)%R]]); + last exact/perm_eqC. rewrite -{1}(RRe_real fir). - rewrite [RHS](@perm_size _ _ [seq x%:C%C | x <- [seq x <- rootsR (evalpmp z p) | (x < complex.Re (val (fyz [` pyrP i ir])))%R]]); last exact/perm_eqC. + rewrite [RHS](@perm_size _ _ [seq x%:C%C | x <- + [seq x <- rootsR (evalpmp z p) | + (x < complex.Re (val (fyz [` pyrP i ir])))%R]]); last exact/perm_eqC. by rewrite !size_map !size_filter. -rewrite -2!srE [LHS](card_imfset _ _ val_inj)/= [RHS](card_imfset _ _ val_inj)/=. -suff ->: sr z (val (fyz [` pyrP i ir])) = [fset fyz x | x in sr y (((rootsR (evalpmp y p))`_i)%:C)%C]. +rewrite -2!srE [LHS](card_imfset _ _ val_inj)/=. +rewrite [RHS](card_imfset _ _ val_inj)/=. +suff ->: sr z (val (fyz [` pyrP i ir])) + = [fset fyz x | x in sr y (((rootsR (evalpmp y p))`_i)%:C)%C]. by rewrite [RHS](card_imfset _ _ (bij_inj fyzb)). apply/eqP; rewrite eqEfsubset; apply/andP; split; apply/fsubsetP => /= u. rewrite [X in _ X -> _](mem_imfset _ _ (@inj_id _))/= => ui. @@ -597,12 +661,14 @@ apply/eqP; rewrite eqEfsubset; apply/andP; split; apply/fsubsetP => /= u. suff: val (fzy u) < ((rootsR (evalpmp y p))`_i)%:C%C by []. rewrite -(RRe_real fur) ltcR ltNge; apply/negP => iu. suff: [` pyrP i ir] = fzy u by move=> iuE; move: ui; rewrite iuE fzyK ltxx. - apply/eP; rewrite /= -(RRe_real fur) -rmorphB/= normcR mulrC mulr_natr -rmorphMn/= ltcR. + apply/eP. + rewrite /= -(RRe_real fur) -rmorphB/= normcR mulrC mulr_natr -rmorphMn/= ltcR. apply/ltr_normlP; split; last first. rewrite -subr_le0 in iu; apply/(le_lt_trans iu). by rewrite pmulrn_lgt0// -ltcR. rewrite opprB -(subrBB (complex.Re (val u))) opprB mulr2n; apply/ltrD. - apply/ltr_normlW; rewrite -ltcR -normcR rmorphB/= (RRe_real fur) (RRe_real ur). + apply/ltr_normlW. + rewrite -ltcR -normcR rmorphB/= (RRe_real fur) (RRe_real ur). by rewrite -{2}(fzyK u); apply/fyze. rewrite -(subrBB (complex.Re (val (fyz [` pyrP i ir])))) opprB -(add0r e'). apply/ltrD; first by rewrite subr_lt0; move: ui; rewrite ltcE => /andP[_]. @@ -620,11 +686,13 @@ suff: val (fyz v) < val (fyz [` pyrP i ir]) by []. have fvr: val (fyz v) \is Num.real by rewrite -fyzr. rewrite -(RRe_real fvr) -(RRe_real fir) ltcR ltNge; apply/negP => iv. suff vE: v = [` pyrP i ir] by rewrite vE/= ltxx in vi. -apply/eP; rewrite /= -(RRe_real vr) -rmorphB/= normcR mulrC mulr_natr -rmorphMn/= ltcR. +apply/eP. +rewrite /= -(RRe_real vr) -rmorphB/= normcR mulrC mulr_natr -rmorphMn/= ltcR. apply/ltr_normlP; split; last first. rewrite -(RRe_real vr) ltcR -subr_lt0 in vi; apply/(lt_trans vi). by rewrite pmulrn_lgt0// -ltcR. -rewrite opprB -(subrBB (complex.Re (val (fyz [`pyrP i ir])))) opprB mulr2n; apply/ltrD. +rewrite opprB -(subrBB (complex.Re (val (fyz [`pyrP i ir])))) opprB mulr2n. +apply/ltrD. apply/ltr_normlW; rewrite -ltcR -normcR rmorphB/= (RRe_real fir). exact/(fyze [` pyrP i ir]). rewrite -(subrBB (complex.Re (val (fyz v)))) opprB -(add0r e'). @@ -633,11 +701,13 @@ apply/ltr_normlW; rewrite -ltcR -normcR rmorphB/= (RRe_real fvr) (RRe_real vr). rewrite -normrN opprB; apply/fyze. Qed. -Definition SAevalpmp_graph n (p : {poly {mpoly R[n]}}) : {SAset R^(n + (size p))} := +Definition SAevalpmp_graph n + (p : {poly {mpoly R[n]}}) : {SAset R^(n + (size p))} := [set| \big[And/True]_(i < size p) subst_formula (rcons (iota 0 n) (n + i)%N) (SAmpoly p`_i)]. -Lemma SAevalpmp_graphP n (p : {poly {mpoly R[n]}}) (u : 'rV[R]_n) (v : 'rV[R]_(size p)) : +Lemma SAevalpmp_graphP n (p : {poly {mpoly R[n]}}) (u : 'rV[R]_n) + (v : 'rV[R]_(size p)) : (row_mx u v \in SAevalpmp_graph p) = (v == \row_i (evalpmp u p)`_i). Proof. apply/SAin_setP/eqP => [/holdsAnd /= vE|->]. @@ -665,7 +735,8 @@ by apply/meval_eq => j; rewrite tnth_mktuple. Qed. Fact SAfun_SAevalpmp n (p : {poly {mpoly R[n]}}) : - (SAevalpmp_graph p \in @SAfunc _ n (size p)) && (SAevalpmp_graph p \in @SAtot _ n (size p)). + (SAevalpmp_graph p \in @SAfunc _ n (size p)) + && (SAevalpmp_graph p \in @SAtot _ n (size p)). Proof. apply/andP; split. by apply/inSAfunc => u y1 y2; rewrite !SAevalpmp_graphP => /eqP -> /eqP. @@ -689,7 +760,8 @@ by rewrite mevalM !coef_map. Qed. (* TODO: subsumed by `rmorph_prod` (with occurence) *) -Lemma evalpmp_prod [I : Type] n (x : 'rV_n) (r : seq I) (F : I -> {poly {mpoly R[n]}}) (P : pred I) : +Lemma evalpmp_prod [I : Type] n (x : 'rV_n) (r : seq I) + (F : I -> {poly {mpoly R[n]}}) (P : pred I) : evalpmp x (\prod_(i <- r | P i) F i) = \prod_(i <- r | P i) evalpmp x (F i). Proof. elim: r => [|i r IHr]. @@ -700,27 +772,42 @@ Qed. Lemma evalpmp_prod_const n (P : {fset {poly {mpoly R[n]}}}) (s : {SAset R^n}) : SAconnected s -> - {in P, forall p, {in s &, forall x y, size (evalpmp x p) = size (evalpmp y p)}} -> - {in P, forall p, {in s &, forall x y, size (gcdp (evalpmp x p) (evalpmp x p)^`()) = size (gcdp (evalpmp y p) (evalpmp y p)^`())}} -> - {in P &, forall p q, {in s &, forall x y, size (gcdp (evalpmp x p) (evalpmp x q)) = size (gcdp (evalpmp y p) (evalpmp y q))}} -> - {in s &, forall x y, size (gcdp (evalpmp x (\prod_(p : P) (val p))) (evalpmp x (\prod_(p : P) (val p)))^`()) = size (gcdp (evalpmp y (\prod_(p : P) (val p))) (evalpmp y (\prod_(p : P) (val p)))^`())} /\ - {in s &, forall x y, size (rootsR (evalpmp x (\prod_(p : P) (val p)))) = size (rootsR (evalpmp y (\prod_(p : P) (val p))))}. + {in P, forall p, + {in s &, forall x y, size (evalpmp x p) = size (evalpmp y p)}} -> + {in P, forall p, {in s &, forall x y, + size (gcdp (evalpmp x p) (evalpmp x p)^`()) + = size (gcdp (evalpmp y p) (evalpmp y p)^`())}} -> + {in P &, forall p q, {in s &, forall x y, + size (gcdp (evalpmp x p) (evalpmp x q)) + = size (gcdp (evalpmp y p) (evalpmp y q))}} -> + {in s &, forall x y, + size (gcdp (evalpmp x (\prod_(p : P) (val p))) + (evalpmp x (\prod_(p : P) (val p)))^`()) + = size (gcdp (evalpmp y (\prod_(p : P) (val p))) + (evalpmp y (\prod_(p : P) (val p)))^`())} /\ + {in s &, forall x y, + size (rootsR (evalpmp x (\prod_(p : P) (val p)))) + = size (rootsR (evalpmp y (\prod_(p : P) (val p))))}. Proof. move=> Scon psize proots pqsize. apply/all_and2 => x; apply/all_and2 => y; apply/all_and2 => xs; apply/all_and2. -case: n P s Scon psize proots pqsize x y xs => [|n] P s Scon psize proots pqsize x y xS yS. +case: n P s Scon psize proots pqsize x y xs + => [|n] P s Scon psize proots pqsize x y xS yS. by have ->: x = y by apply/rowP => -[]. case: (eqVneq (evalpmp x (\prod_(p : P) val p)) 0) => px0. rewrite px0. move: px0; rewrite !evalpmp_prod => /eqP/prodf_eq0/= [p] _. - rewrite -size_poly_eq0 (psize (val p) (fsvalP p) x y xS yS) size_poly_eq0 => py0. + rewrite -size_poly_eq0 (psize (val p) (fsvalP p) x y xS yS) size_poly_eq0. + move=> py0. suff ->: \prod_(p : P) evalpmp y (val p) = 0 by []. by apply/eqP/prodf_eq0; exists p. have p0: {in P, forall p, {in s, forall x, size (evalpmp x p) != 0}}. move=> p pP z zS; rewrite (psize _ pP z x zS xS) size_poly_eq0. by move: px0; rewrite evalpmp_prod => /prodf_neq0/(_ [` pP] isT). have R0: [char R[i]] =i pred0 by apply/char_num. -have pmsize: {in s &, forall x y, size (evalpmp x (\prod_(p : P) \val p)) = size (evalpmp y (\prod_(p : P) \val p))}. +have pmsize: {in s &, forall x y, + size (evalpmp x (\prod_(p : P) \val p)) + = size (evalpmp y (\prod_(p : P) \val p))}. move=> {px0} {}x {}y {}xS {}yS. rewrite !evalpmp_prod size_prod; last first. by move=> /= p _; rewrite -size_poly_eq0 (p0 _ (fsvalP p) x xS). @@ -731,7 +818,8 @@ have pmsize: {in s &, forall x y, size (evalpmp x (\prod_(p : P) \val p)) = size over. by []. have rE (u : 'rV[R]_n.+1) : - (size (rootsR (evalpmp u (\prod_(p : P) \val p))))%:R = SAcomp (SAnbroots _ _) (SAevalpmp (\prod_(p : P) \val p)) u ord0 ord0. + (size (rootsR (evalpmp u (\prod_(p : P) \val p))))%:R + = SAcomp (SAnbroots _ _) (SAevalpmp (\prod_(p : P) \val p)) u ord0 ord0. rewrite SAcompE/= SAevalpmpE SAnbrootsE mxE. congr (size (rootsR _))%:R. apply/polyP => i; rewrite coef_poly. @@ -739,7 +827,9 @@ have rE (u : 'rV[R]_n.+1) : exact/nth_default/(leq_trans (size_poly _ _) ilt). by rewrite -/(nat_of_ord (Ordinal ilt)) nth_map_ord mxE. have cE (u : 'rV[R]_n.+1) : - (size (dec_roots (map_poly (real_complex R) (evalpmp u (\prod_(p : P) \val p)))))%:R = SAcomp (SAnbrootsC _ _) (SAevalpmp (\prod_(p : P) \val p)) u ord0 ord0. + (size (dec_roots + (map_poly (real_complex R) (evalpmp u (\prod_(p : P) \val p)))))%:R + = SAcomp (SAnbrootsC _ _) (SAevalpmp (\prod_(p : P) \val p)) u ord0 ord0. rewrite SAcompE/= SAevalpmpE SAnbrootsCE mxE. congr (size (dec_roots _))%:R. apply/polyP => i; rewrite !coef_poly. @@ -748,8 +838,10 @@ have cE (u : 'rV[R]_n.+1) : by rewrite (nth_mktuple _ _ (Ordinal ilt')) mxE nth_default. case: ltnP => [|//] ilt'. by rewrite (nth_mktuple _ _ (Ordinal ilt')) mxE coef_map/=. -suff: locally_constant (SAcomp (SAnbroots _ _) (SAevalpmp (\prod_(p : P) \val p))) [set` s] /\ - locally_constant (SAcomp (SAnbrootsC _ _) (SAevalpmp (\prod_(p : P) \val p))) [set` s]. +suff: locally_constant (SAcomp (SAnbroots _ _) + (SAevalpmp (\prod_(p : P) \val p))) [set` s] + /\ locally_constant (SAcomp (SAnbrootsC _ _) + (SAevalpmp (\prod_(p : P) \val p))) [set` s]. move=> [] rc cc; split; last first. apply/(mulrIn (@oner_neq0 R)). rewrite !rE; congr (_ _ ord0 ord0). @@ -780,9 +872,11 @@ have ex_and: forall T (P Q : T -> Prop), move=> T P0 Q [] a [] Pa Qa. by split; exists a. apply/ex_and. -pose pc := fun (p : P) (x : 'rV[R]_n.+1) => map_poly (real_complex R) (evalpmp x (\val p)). +pose pc := fun (p : P) (x : 'rV[R]_n.+1) => + map_poly (real_complex R) (evalpmp x (\val p)). pose rx := dec_roots (\prod_(p : P) pc p x). -pose d := (\big[Order.min/1]_(x <- rx) \big[Order.min/1]_(y <- rx | y != x) (complex.Re `|x - y| / 2))%:C%C. +pose d := (\big[Order.min/1]_(x <- rx) \big[Order.min/1]_(y <- rx | y != x) + (complex.Re `|x - y| / 2))%:C%C. have d0 : 0 < d. rewrite ltcE/= eqxx/= lt_bigmin ltr01/=; apply/allP => u _. rewrite -big_filter lt_bigmin ltr01/=; apply/allP => v. @@ -792,23 +886,31 @@ have d0 : 0 < d. have /= dP: {in rx &, forall u v, (`|u - v| < 2*d)%R -> u = v}. move=> u v ux vx uvd; apply/eqP; rewrite -[_ == _]negbK; apply/negP => uv. move: uvd. - rewrite mulrC mulr_natr -rmorphMn/= -(RRe_real (normr_real _)) ltcR -mulr_natr. + rewrite mulrC mulr_natr -rmorphMn/= -(RRe_real (normr_real _)) ltcR. + rewrite -mulr_natr. rewrite -ltr_pdivrMr ?ltr0Sn// lt_bigmin => /andP[_] /allP-/(_ u ux) /=. rewrite lt_bigmin => /andP[_] /allP-/(_ v vx) /implyP. by rewrite eq_sym ltxx => /(_ uv). -have /fin_all_exists /=: forall p : P, exists delta, 0 < delta /\ forall (y : 'rV[R]_n.+1), y \in s -> `|x - y| < delta -> alignp d (pc p x) (pc p y). +have /fin_all_exists /=: + forall p : P, exists delta, 0 < delta + /\ forall (y : 'rV[R]_n.+1), y \in s -> `|x - y| < delta -> + alignp d (pc p x) (pc p y). move=> p. case: (aligned_deformed (pc p x) d0) => /= [[]] e eI []. rewrite ltcE/= => /andP[/eqP ->] e0; rewrite complexr0 => ed. - have /fin_all_exists /=: forall i : 'I_(size (val p)).+1, exists delta, 0 < delta /\ forall (y : 'rV[R]_n.+1), y \in s -> `|x - y| < delta -> `|(pc p x)`_i - (pc p y)`_i| < e%:C%C. + have /fin_all_exists /=: + forall i : 'I_(size (val p)).+1, exists delta, 0 < delta + /\ forall (y : 'rV[R]_n.+1), y \in s -> `|x - y| < delta -> + `|(pc p x)`_i - (pc p y)`_i| < e%:C%C. move=> i. move: (@meval_continuous _ _ (val p)`_i x). rewrite /= /continuous_at. - move=> /(@cvgr_dist_lt _ (GRing_regular__canonical__normedtype_PseudoMetricNormedZmod R)). + move=> /(@cvgr_dist_lt _ R^o). move=> /(_ _ e0) /nbhs_ballP[] d'/= d'0 /subsetP xd'. exists d'; split=> // y ys xy. move: xd' => /(_ y); mp; first by rewrite -ball_normE inE/=. - rewrite inE/= !coef_map/= -rmorphB/= normc_def/= expr0n/= addr0 sqrtr_sqr ltcR. + rewrite inE/= !coef_map/= -rmorphB/= normc_def/= expr0n/= addr0 sqrtr_sqr. + rewrite ltcR. by congr (normr (_ - _) < e); apply/meval_eq => j; rewrite tnth_mktuple. move=> [d'] d'P; exists (\big[minr/1]_i d' i). split; first by rewrite lt_bigmin ltr01; apply/allP => i _ /=; case: (d'P i). @@ -833,8 +935,9 @@ apply/all_and2 => y; rewrite in_setI. apply/all_and2 => /andP[]; rewrite inE/= => ys. rewrite -ball_normE inE/= lt_bigmin => /andP[_] /allP/= xy. pose rs := fun x => [fset x in (rootsR (evalpmp x (\prod_(p : P) \val p)))]. -have card_rs : forall x, #|` rs x | = size (rootsR (evalpmp x (\prod_(p : P) \val p))). - by move=> z; rewrite /rs card_imfset//= undup_id// uniq_roots. +have card_rs z : + #|` rs z | = size (rootsR (evalpmp z (\prod_(p : P) \val p))). + by rewrite /rs card_imfset//= undup_id// uniq_roots. have pc0 p z: z \in s -> pc p z != 0. by rewrite map_poly_eq0 -size_poly_eq0; apply/p0 => //; apply/fsvalP. have pcM0 z: z \in s -> \prod_(p : P) pc p z != 0. @@ -849,7 +952,8 @@ have: exists (fxy : forall p, move: xd' => /(_ p)[_] /(_ y ys) /[apply] /(_ u pu). rewrite -big_filter; case rsy: (seq.filter _ _) => [|v l]. by rewrite big_nil leqn0 mu_eq0 ?pu// pc0. - move: (mem_head v l); rewrite -rsy mem_filter -normrN opprB => /andP[] uv pv _. + move: (mem_head v l). + rewrite -rsy mem_filter -normrN opprB => /andP[] uv pv _. suff vP: v \in [fset x in dec_roots (pc p y)]. by exists [` vP]. by rewrite mem_imfset//= mem_dec_roots pc0. @@ -873,7 +977,8 @@ have fxy_bij: forall p, bijective (fxy p). rewrite mem_dec_roots => /andP[_] pv. rewrite /rx mem_dec_roots pcM0//= root_bigmul/=. by apply/hasP; exists p => //; apply/mem_index_enum. - - rewrite -(subrBB (val (fxy p u))) {2}uv; apply/(le_lt_trans (ler_normB _ _)). + - rewrite -(subrBB (val (fxy p u))) {2}uv. + apply/(le_lt_trans (ler_normB _ _)). by rewrite mulr2n mulrDl mul1r; apply/ltrD; apply/fxyd. have: exists (fyx : forall p, [fset x in dec_roots (pc p y)] -> [fset x in dec_roots (pc p x)]), @@ -896,33 +1001,44 @@ have lift p z (u : [fset x in dec_roots (pc p z)]) : by apply/hasP; exists p => //; apply/mem_index_enum. have unlift z (u : [fset x in dec_roots (\prod_(p : P) pc p z)]) : {p : P | val u \in [fset x in dec_roots (pc p z)]}. - case: u => /= u; rewrite mem_imfset//= mem_dec_roots root_bigmul prodf_seq_neq0. - move=> /andP[+] /hasP/sig2W[/=] p _ pu => /allP/(_ p (mem_index_enum _)) /= pz0. + case: u => /= u. + rewrite mem_imfset//= mem_dec_roots root_bigmul prodf_seq_neq0. + move=> /andP[+] /hasP/sig2W[/=] p _ pu. + move=> /allP/(_ p (mem_index_enum _)) /= pz0. by exists p; rewrite mem_imfset//= mem_dec_roots pz0. -have /fin_all_exists/=: forall (u : [fset x in dec_roots (\prod_(p : P) pc p x)]), exists (v : [fset x in dec_roots (\prod_(p : P) pc p y)]), `|val u - val v| < d. +have /fin_all_exists/=: + forall (u : [fset x in dec_roots (\prod_(p : P) pc p x)]), + exists (v : [fset x in dec_roots (\prod_(p : P) pc p y)]), + `|val u - val v| < d. move=> u; case: (unlift x u) => p pu. by exists [` (lift p y (fxy p [` pu]) ys)] => /=; apply/fxyd. move=> []gxy gxyd. -have /fin_all_exists/=: forall (u : [fset x in dec_roots (\prod_(p : P) pc p y)]), exists (v : [fset x in dec_roots (\prod_(p : P) pc p x)]), `|val u - val v| < d. +have /fin_all_exists/=: + forall (u : [fset x in dec_roots (\prod_(p : P) pc p y)]), + exists (v : [fset x in dec_roots (\prod_(p : P) pc p x)]), + `|val u - val v| < d. move=> u; case: (unlift y u) => p pu. by exists [` (lift p x (fyx p [` pu]) xs)] => /=; apply/fyxd. move=> []gyx gyxd. have gyxE p (u : [fset x in dec_roots (pc p y)]) : gyx [` lift p y u ys] = [` lift p x (fyx p u) xs]. apply/val_inj/dP => /=. - - by move: (fsvalP (gyx [`lift p y u ys])); rewrite (mem_imfset _ _ (@inj_id _))/=. + - move: (fsvalP (gyx [`lift p y u ys])). + by rewrite (mem_imfset _ _ (@inj_id _))/=. - by move: (lift p x (fyx p u) xs); rewrite (mem_imfset _ _ (@inj_id _))/=. rewrite -(subrBB (val u)) opprB -normrN opprD opprB. apply/(le_lt_trans (ler_normB _ _)). rewrite mulr2n mulrDl mul1r; apply/ltrD; first exact/gyxd. exact/fyxd. -have size_pc (p : {poly R[i]}) : size p = ((\sum_(u <- dec_roots p) \mu_u p) + (p != 0%R))%N. +have size_pc (p : {poly R[i]}) : + size p = ((\sum_(u <- dec_roots p) \mu_u p) + (p != 0%R))%N. have [->|{}p0] := eqVneq p 0; first by rewrite size_poly0 dec_roots0 big_nil. move: (dec_roots_closedP p) => /(congr1 (fun p : {poly _} => size p)). rewrite size_scale; last by rewrite -lead_coefE lead_coef_eq0 p0. rewrite size_prod_seq => /= [| w _]; last first. by rewrite expf_eq0 polyXsubC_eq0 andbF. - under eq_bigr do rewrite my_size_exp ?polyXsubC_eq0// size_XsubC/= mul1n -addn1. + under eq_bigr do rewrite my_size_exp ?polyXsubC_eq0// size_XsubC/= mul1n. + under eq_bigr do rewrite -addn1. by rewrite big_split/= sum1_size -addSn subDnAC// subnn -addn1. have dP' p u: (count (fun v => (`|u - v| < d)%R) (dec_roots (pc p x)) <= 1)%N. rewrite -size_filter. @@ -946,15 +1062,18 @@ have ball_root1 (p : P) (u : [fset x | x in dec_roots (\prod_p pc p y)]) : root (pc p y) (val u) -> [seq v <- dec_roots (pc p y) | `|v - val (gyx u)| < d] = [:: val u]. move=> pu. - have: all (fun v => v == val u) [seq v <- dec_roots (pc p y) | `|v - val (gyx u)| < d]. + have: all (fun v => v == val u) + [seq v <- dec_roots (pc p y) | `|v - val (gyx u)| < d]. apply/allP => v; rewrite mem_filter => /andP[] vu vp. have uP: val u \in [fset x | x in dec_roots (pc p y)]. by rewrite (mem_imfset _ _ (@inj_id _))/= mem_dec_roots pc0. have vP: v \in [fset x | x in dec_roots (pc p y)] by rewrite mem_imfset. apply/eqP; rewrite -[val u]/(val [` uP]) -[v]/(val [` vP]) ; congr val. apply/(can_inj (fyxK p))/val_inj/dP. - - by move: (fsvalP [` lift p x (fyx p [` vP]) xs]); rewrite (mem_imfset _ _ (@inj_id _))/=. - - by move: (fsvalP [` lift p x (fyx p [` uP]) xs]); rewrite (mem_imfset _ _ (@inj_id _))/=. + - move: (fsvalP [` lift p x (fyx p [` vP]) xs]). + by rewrite (mem_imfset _ _ (@inj_id _))/=. + - move: (fsvalP [` lift p x (fyx p [` uP]) xs]). + by rewrite (mem_imfset _ _ (@inj_id _))/=. rewrite -(subrBB v) opprB -normrN opprD opprB. apply/(le_lt_trans (ler_normB _ _)). rewrite mulr2n mulrDl mul1r; apply/ltrD; first exact/fyxd. @@ -1000,7 +1119,8 @@ have mu_gyx p (u : [fset x | x in dec_roots (\prod_p pc p y)]) : move: (size_pc (pc p x)) (size_pc (pc p y)). rewrite !size_map_poly => -> -> /eqP. rewrite !pc0// !addn1 eqSS => /eqP. - rewrite -[RHS](big_rmcond_in (fun u => has (fun v => `|u - v| < d) (dec_roots (pc p x))))/=; last first. + rewrite -[RHS](big_rmcond_in (fun u => + has (fun v => `|u - v| < d) (dec_roots (pc p x))))/=; last first. move=> v pv. have vP : v \in [fset x | x in dec_roots (pc p y)] by rewrite mem_imfset//=. rewrite -all_predC => /allP. @@ -1026,23 +1146,27 @@ have gyxK: cancel gyx gxy. apply/eqP; rewrite -[_ == _]negbK; apply/negP => /eqP wv. move: (pqsize (val p) (val q) (fsvalP p) (fsvalP q) x y xs ys). move: (size_pc (gcdp (pc p x) (pc q x))) (size_pc (gcdp (pc p y) (pc q y))). - rewrite !gcdp_eq0 !negb_and !pc0//= !addn1 -!gcdp_map !size_map_poly => -> -> /eqP. + rewrite !gcdp_eq0 !negb_and !pc0//= !addn1 -!gcdp_map !size_map_poly. + move=> -> -> /eqP. rewrite eqSS !gcdp_map -!/(pc _ _) => /eqP/esym. under eq_bigr do rewrite mu_gcdp ?pc0//. under [in RHS]eq_bigr do rewrite mu_gcdp ?pc0//. apply/ltnn_ne. - rewrite -(big_rmcond_in (fun u => has (fun v => `|u - v| < d) (dec_roots (gcdp (pc p x) (pc q x)))))/=; last first. + rewrite -(big_rmcond_in (fun u => has (fun v => `|u - v| < d) + (dec_roots (gcdp (pc p x) (pc q x)))))/=; last first. move=> a; rewrite mem_dec_roots root_gcd => /andP[_] /andP[] pa qa. rewrite -all_predC => /allP/=. have aP: a \in [fset x in dec_roots (\prod_(p : P) pc p y)]. rewrite mem_imfset//= mem_dec_roots pcM0//= root_bigmul. by apply/hasP; exists p => //; apply/mem_index_enum. - suff /[swap]/[apply]: val (gyx [` aP]) \in dec_roots (gcdp (pc p x) (pc q x)). + suff /[swap]/[apply]: + val (gyx [` aP]) \in dec_roots (gcdp (pc p x) (pc q x)). by rewrite gyxd. by rewrite mem_dec_roots gcdp_eq0 negb_and !pc0//= root_gcd !gyx_root//. rewrite (@big_hasE _ _ _ _ _ _ xpredT)/=; last first. move=> a _; rewrite -size_filter. - move: (filter_uniq (fun v => `|a - v| < d) (uniq_dec_roots (gcdp (pc p x) (pc q x)))). + move: (filter_uniq (fun v => + `|a - v| < d) (uniq_dec_roots (gcdp (pc p x) (pc q x)))). case rsdE: (seq.filter _ _) => [//|b rsd]. case: rsd rsdE => [//|c rsd] rsdE /= /andP[] + _. rewrite in_cons negb_or => /andP[] /eqP + _. @@ -1066,7 +1190,8 @@ have gyxK: cancel gyx gxy. by rewrite mem_imfset//= mem_dec_roots pc0. have aQ: a \in [fset x | x in dec_roots (pc q x)]. by rewrite mem_imfset//= mem_dec_roots pc0. - have: uniq [seq i <- dec_roots (gcdp (pc p y) (pc q y)) | normr (i - a) < d]. + have: + uniq [seq i <- dec_roots (gcdp (pc p y) (pc q y)) | normr (i - a) < d]. exact/filter_uniq/uniq_dec_roots. have: all (fun b => b == val (fxy p [` aP])) [seq i <- dec_roots (gcdp (pc p y) (pc q y)) | (normr (i - a) < d)%R]. @@ -1204,7 +1329,8 @@ have gR (u : [fset x | x in dec_roots (\prod_p pc p x)]) : exact/gxyd. split; last first; apply/eqP; rewrite rowPE forall_ord1. rewrite -!cE eqr_nat; apply/eqP. - pose cs (z : 'rV[R]_n.+1) := dec_roots (map_poly (real_complex R) (evalpmp z (\prod_(p : P) val p))). + pose cs (z : 'rV[R]_n.+1) := + dec_roots (map_poly (real_complex R) (evalpmp z (\prod_(p : P) val p))). have card_cs z: #|` [fset x in cs z]| = size (cs z). by rewrite /rs card_imfset//= undup_id// uniq_dec_roots. rewrite -(card_cs x) -(card_cs y). @@ -1223,9 +1349,10 @@ have liftRP: forall z, z \in s -> move: (fun z zs => fin_all_exists (liftRP z zs)) => {}liftRP. case: (liftRP x xs) => liftxR liftxRE. case: (liftRP y ys) => liftyR liftyRE {liftRP}. -have /fin_all_exists: forall (u : [fset x | x in rootsR (\prod_(p : P) (evalpmp x (val p)))]), - exists (v : [fset x | x in rootsR (\prod_(p : P) (evalpmp y (val p)))]), - (val v)%:C%C = val (gxy (liftxR u)). +have /fin_all_exists: + forall (u : [fset x | x in rootsR (\prod_(p : P) (evalpmp x (val p)))]), + exists (v : [fset x | x in rootsR (\prod_(p : P) (evalpmp y (val p)))]), + (val v)%:C%C = val (gxy (liftxR u)). move=> u. have: val (liftxR u) \is Num.real. by apply/complex_realP; exists (val u); apply/liftxRE. @@ -1238,9 +1365,10 @@ have /fin_all_exists: forall (u : [fset x | x in rootsR (\prod_(p : P) (evalpmp rewrite -uE (mem_imfset _ _ (@inj_id _))/= mem_dec_roots. by rewrite -{1 2}rmorph_prod/= fmorph_root map_poly_eq0. move=> [] hxy hxyE. -have /fin_all_exists: forall (u : [fset x | x in rootsR (\prod_(p : P) (evalpmp y (val p)))]), - exists (v : [fset x | x in rootsR (\prod_(p : P) (evalpmp x (val p)))]), - (val v)%:C%C = val (gyx (liftyR u)). +have /fin_all_exists: + forall (u : [fset x | x in rootsR (\prod_(p : P) (evalpmp y (val p)))]), + exists (v : [fset x | x in rootsR (\prod_(p : P) (evalpmp x (val p)))]), + (val v)%:C%C = val (gyx (liftyR u)). move=> u. have: val (liftyR u) \is Num.real. by apply/complex_realP; exists (val u); apply/liftyRE. @@ -1270,7 +1398,8 @@ Definition elimp_subdef1 n (P : {fset {mpoly R[n.+1]}}) := Definition elimp_subdef2 n (P : {fset {mpoly R[n.+1]}}) := \big[fsetU/fset0]_(p : elimp_subdef1 P | (3 <= size (val p))%N) - [fset subresultant (val j) (val p) (val p)^`() | j : 'I_(size (val p)).-2]. + [fset subresultant (val j) (val p) (val p)^`() | + j : 'I_(size (val p)).-2]. Definition elimp_subdef3 n (P : {fset {mpoly R[n.+1]}}) := \big[fsetU/fset0]_(p : elimp_subdef1 P) @@ -1289,9 +1418,11 @@ Definition elimp_subdef5 n (P : {fset {mpoly R[n.+1]}}) := [fset lead_coef (val p) | p : elimp_subdef1 P]. Definition elimp n (P : {fset {mpoly R[n.+1]}}) := - [fset p in elimp_subdef2 P `|` elimp_subdef3 P (* `|` elimp_subdef4 P *) `|` elimp_subdef5 P | (1 < msize (val p))%N]. + [fset p in elimp_subdef2 P `|` elimp_subdef3 P + (* `|` elimp_subdef4 P *) `|` elimp_subdef5 P | (1 < msize (val p))%N]. -Lemma map_poly_truncate (A B : ringType) (f : {rmorphism A -> B}) d (p : {poly A}) : +Lemma map_poly_truncate (A B : ringType) (f : {rmorphism A -> B}) d + (p : {poly A}) : map_poly f (truncate p d) = truncate (map_poly f p) d. Proof. apply/polyP => i. @@ -1328,7 +1459,8 @@ have: val s \in [fset s by rewrite !inE/= eqxx andbT; apply/fsvalP. rewrite SE => /imfsetP[/=] t txi ->. apply/SAcast_connected. -apply/(@SAconnected_partition_of_graphs_above _ _ (SAset_cast n (val s)) (val xi)). +apply/(@SAconnected_partition_of_graphs_above _ _ + (SAset_cast n (val s)) (val xi)). - exact/(IHn _ [` sP] SCD). - exact/xisort. - move=> /= i; rewrite size_tuple => im; move: (xicont (Ordinal im)). @@ -1366,12 +1498,14 @@ have nth_const (s' : S') (p : P) x y : forall i, ((size (evalpmp x (muni (val p)))).-1 <= i)%N -> sgz (evalpmp x (muni (val p)))`_i = sgz (evalpmp y (muni (val p)))`_i. move=> xs ys i xi. - have iE: forall z, (evalpmp z (muni (\val p)))`_i = (truncate (evalpmp z (muni (\val p))) i.+1)`_i. - move=> z; rewrite [RHS]coef_poly ltn_min leqnn/=. + have iE z: (evalpmp z (muni (\val p)))`_i + = (truncate (evalpmp z (muni (\val p))) i.+1)`_i. + rewrite [RHS]coef_poly ltn_min leqnn/=. case: ifP => [//|] /negP/negP; rewrite -leqNgt => {}zi. by rewrite nth_default. rewrite !iE -!map_poly_truncate/= !coef_map/=. - case: (ltnP 1 (msize ((truncate (muni (val p)) i.+1)`_i))) => [pi1|]; last first. + case: (ltnP 1 (msize ((truncate (muni (val p)) i.+1)`_i))) + => [pi1|]; last first. by move=> /msize1_polyC ->; rewrite !mevalC. move: xi; rewrite -ltnS => /(leq_trans (leqSpred _)) xi. suff iP: (truncate (muni (fsval p)) i.+1)`_i \in elimp P. @@ -1389,8 +1523,11 @@ have nth_const (s' : S') (p : P) x y : apply/bigfcupP; exists p; first by rewrite mem_index_enum. exact/(truncations_witness xi). have S'size (s' : S') (p : P) : - {in val s', forall x, size (evalpmp x (muni (val p))) = size (evalpmp (proj1_sig (pick s')) (muni (val p)))}. - suff: {in val s' &, forall x y, (size (evalpmp x (muni (val p))) <= size (evalpmp y (muni (val p))))%N}. + {in val s', forall x, + size (evalpmp x (muni (val p))) + = size (evalpmp (proj1_sig (pick s')) (muni (val p)))}. + suff: {in val s' &, forall x y, + (size (evalpmp x (muni (val p))) <= size (evalpmp y (muni (val p))))%N}. move=> S'size x xS; apply/anti_leq/andP. split; apply/S'size => //; exact/(proj2_sig (pick s')). move=> x y xs ys; apply/leq_sizeP => i yi. @@ -1407,8 +1544,10 @@ have res_const (s' : S') (p q : P) (x y : 'rV_n): forall i, (i <= (size (evalpmp (val (pick s')) (muni (val p)))).-1)%N -> (i <= (size (evalpmp (val (pick s')) (muni (val q)))).-1)%N -> - sgz (subresultant i (evalpmp x (muni (\val p))) (evalpmp x (muni (\val q)))) = - sgz (subresultant i (evalpmp y (muni (\val p))) (evalpmp y (muni (\val q)))). + sgz (subresultant i (evalpmp x (muni (\val p))) + (evalpmp x (muni (\val q)))) = + sgz (subresultant i (evalpmp y (muni (\val p))) + (evalpmp y (muni (\val q)))). move=> xs ys i; rewrite {1}leq_eqVlt => /orP[/eqP -> _|ip]. rewrite -{1}(S'size s' p x xs) -(S'size s' p y ys). rewrite !subresultant_last !sgzX; congr (_ ^+ (_.-1 - _.-1)); last first. @@ -1429,7 +1568,8 @@ have res_const (s' : S') (p q : P) (x y : 'rV_n): - by rewrite (S'size s' p x xs) (S'size s' p y ys). rewrite !lead_coefE (S'size s' q y ys) -(S'size s' q x xs). by apply/(nth_const s'); last exact/leqnn. - pose Q (r : P) := truncate (muni (\val r)) (size (evalpmp (val (pick s')) (muni (\val r)))). + pose Q (r : P) := + truncate (muni (\val r)) (size (evalpmp (val (pick s')) (muni (\val r)))). wlog: p q ip iq / (size (Q q) <= size (Q p))%N => qp. move/orP: (leq_total (size (Q q)) (size (Q p))). case=> [/(qp p q ip iq)//|] /(qp q p iq ip) {}qp. @@ -1455,14 +1595,17 @@ have res_const (s' : S') (p q : P) (x y : 'rV_n): by rewrite -lead_coef_eq0 lead_coefE coef_map/= r0 meval0. by move: s0; rewrite r0 size_poly0 ltnn. (* N.B. Why does Coq stop responding if I do not give the location? *) - rewrite [X in subresultant i X](QE p x xs) [X in _ = sgz (subresultant i X _)](QE p y ys). - rewrite [X in subresultant i _ X](QE q x xs) [X in _ = sgz (subresultant i _ X)](QE q y ys). + rewrite [X in subresultant i X](QE p x xs). + rewrite [X in _ = sgz (subresultant i X _)](QE p y ys). + rewrite [X in subresultant i _ X](QE q x xs). + rewrite [X in _ = sgz (subresultant i _ X)](QE q y ys). have Q0 (r : P) z : z \in val s' -> (i < (size (evalpmp (\val (pick s')) (muni (\val r)))).-1)%N -> (lead_coef (Q r)).@[tnth (ngraph z)] != 0. move=> zs ir. rewrite lead_coefE -coef_map -(Qsize r z zs) -lead_coefE lead_coef_eq0. - rewrite -size_poly_eq0 (Qsize r z zs) -(Qsize r _ (proj2_sig (pick s'))) -(QE _ _ (proj2_sig (pick s'))). + rewrite -size_poly_eq0 (Qsize r z zs) -(Qsize r _ (proj2_sig (pick s'))). + rewrite -(QE _ _ (proj2_sig (pick s'))). by apply/eqP => s0; rewrite s0 in ir. rewrite !subresultant_map_poly/=; first last. - exact/Q0. @@ -1490,22 +1633,26 @@ have res'_const (s' : S') (p : P) (x y : 'rV_n): y \in \val s' -> forall i, (i <= (size (evalpmp (val (pick s')) (muni (val p)))).-2)%N -> - sgz (subresultant i (evalpmp x (muni (\val p))) (evalpmp x (muni (\val p)))^`()) = - sgz (subresultant i (evalpmp y (muni (\val p))) (evalpmp y (muni (\val p)))^`()). + sgz (subresultant i (evalpmp x (muni (\val p))) + (evalpmp x (muni (\val p)))^`()) = + sgz (subresultant i (evalpmp y (muni (\val p))) + (evalpmp y (muni (\val p)))^`()). move=> xs ys i. rewrite leq_eqVlt => /orP[/eqP ->|ilt]. rewrite -{1}(S'size s' p x xs) -(S'size s' p y ys) -(size_deriv _ R0). rewrite -[in RHS](size_deriv _ R0). - rewrite subresultantC subresultant_last (size_deriv _ R0) (S'size s' p x xs). - rewrite subresultantC subresultant_last (size_deriv _ R0) (S'size s' p y ys). - rewrite !sgzM; congr (_ * _). + rewrite subresultantC subresultant_last (size_deriv _ R0). + rewrite (S'size s' p x xs). + rewrite subresultantC subresultant_last (size_deriv _ R0). + rewrite (S'size s' p y ys) !sgzM; congr (_ * _). case: (_.-1) => [|j]; first by rewrite !expr0. rewrite succnK subSn; last by []. rewrite subnn !expr1 !(lead_coef_deriv _ R0). rewrite -mulr_natr -[in RHS]mulr_natr !lead_coefE !sgzM. rewrite (S'size s' p y ys) -(S'size s' p x xs); congr (_ * _). exact/(nth_const s'). - set q := truncate (muni (\val p)) (size (evalpmp (val (pick s')) (muni (\val p)))). + set q := + truncate (muni (\val p)) (size (evalpmp (val (pick s')) (muni (\val p)))). rewrite -!/(evalpmp _ _). have qE z : z \in val s' -> (evalpmp z (muni (val p))) = evalpmp z q. move=> zs. @@ -1549,8 +1696,11 @@ have res'_const (s' : S') (p : P) (x y : 'rV_n): apply/bigfcupP; exists [` qP]; first by rewrite mem_index_enum/=. by apply/imfsetP => /=; exists (Ordinal iq). have S'constR (s' : S') (p : P) : - {in val s', forall x, size (rootsR (evalpmp x (muni (val p)))) = size (rootsR (evalpmp (proj1_sig (pick s')) (muni (val p))))}. - move=> x xs; move: (proj2_sig (pick s')); set x' := proj1_sig (pick s') => x's. + {in val s', forall x, + size (rootsR (evalpmp x (muni (val p)))) + = size (rootsR (evalpmp (proj1_sig (pick s')) (muni (val p))))}. + move=> x xs; move: (proj2_sig (pick s')). + set x' := proj1_sig (pick s') => x's. have [p0|x'0] := eqVneq (evalpmp x' (muni (val p))) 0. rewrite p0; suff ->: evalpmp x (muni (val p)) = 0 by []. by apply/eqP; rewrite -size_poly_eq0 (S'size s')// size_poly_eq0; apply/eqP. @@ -1570,7 +1720,8 @@ have S'constR (s' : S') (p : P) : exact/(nth_const s'). rewrite add0n => /leq_predn; rewrite succnK => ilt; apply/eqP. exact/(res'_const s'). -pose P' (s : S') := [fset muni (val p) | p : P & evalpmp (val (pick s)) (muni (val p)) != 0]. +pose P' (s : S') := + [fset muni (val p) | p : P & evalpmp (val (pick s)) (muni (val p)) != 0]. have size_gcd_const (s' : S') : {in P' s', forall p : {poly {mpoly R[n]}}, {in \val s' &, @@ -1580,9 +1731,11 @@ have size_gcd_const (s' : S') : {in P' s', move=> q /imfsetP[/=] p _ -> {q} x y xs ys. have [px0|px0] := eqVneq (evalpmp x (muni (val p)))^`() 0. rewrite px0; move/eqP: px0. - rewrite -size_poly_eq0 (size_deriv _ R0) (S'size s' p x xs) -(S'size s' p y ys) -(size_deriv _ R0) size_poly_eq0 => /eqP ->. + rewrite -size_poly_eq0 (size_deriv _ R0) (S'size s' p x xs). + rewrite -(S'size s' p y ys) -(size_deriv _ R0) size_poly_eq0 => /eqP ->. by rewrite !gcdp0 (S'size s' p x xs) (S'size s' p y ys). - move: (px0); rewrite -size_poly_eq0 (size_deriv _ R0) (S'size s' p x xs) -(S'size s' p y ys) -(size_deriv _ R0) size_poly_eq0 => py0. + move: (px0); rewrite -size_poly_eq0 (size_deriv _ R0) (S'size s' p x xs). + rewrite -(S'size s' p y ys) -(size_deriv _ R0) size_poly_eq0 => py0. rewrite -[LHS]prednK; last first. rewrite ltnNge leqn0 size_poly_eq0 gcdp_eq0; apply/andP => -[_ p0]. by rewrite p0 in px0. @@ -1590,7 +1743,8 @@ have size_gcd_const (s' : S') : {in P' s', rewrite ltnNge leqn0 size_poly_eq0 gcdp_eq0; apply/andP => -[_ p0]. by rewrite p0 in py0. apply/esym/eqP; rewrite eqSS. - move: (eqxx (size (gcdp (evalpmp x (muni (val p))) (evalpmp x (muni (val p)))^`())).-1). + move: (eqxx + (size (gcdp (evalpmp x (muni (val p))) (evalpmp x (muni (val p)))^`())).-1). rewrite gcdp_subresultant; first last. - apply/leq_predn/leq_gcdpr/negP => p0. by rewrite p0 in px0. @@ -1599,13 +1753,15 @@ have size_gcd_const (s' : S') : {in P' s', - by []. - by apply/eqP => p0; rewrite p0 deriv0 eqxx in px0. rewrite gcdp_subresultant; first last. - - rewrite (size_deriv _ R0) (S'size s' p y ys) -(S'size s' p x xs) -[X in (_ <= X.-1)%N](size_deriv _ R0). + - rewrite (size_deriv _ R0) (S'size s' p y ys) -(S'size s' p x xs). + rewrite -[X in (_ <= X.-1)%N](size_deriv _ R0). apply/leq_predn/leq_gcdpr/negP => p0. by rewrite p0 in px0. - rewrite (S'size s' p y ys) -(S'size s' p x xs). apply/leq_predn/leq_gcdpl/eqP => p0. by rewrite p0 deriv0 eqxx in px0. - - by rewrite -size_poly_eq0 (size_deriv _ R0) (S'size s' p y ys) -(S'size s' p x xs) -(size_deriv _ R0) size_poly_eq0. + - rewrite -size_poly_eq0 (size_deriv _ R0) (S'size s' p y ys). + by rewrite -(S'size s' p x xs) -(size_deriv _ R0) size_poly_eq0. - rewrite -size_poly_eq0 (S'size s' p y ys) -(S'size s' p x xs) size_poly_eq0. by apply/eqP => p0; rewrite p0 deriv0 eqxx in px0. move=> /andP[] /forallP si sl; apply/andP; split. @@ -1628,20 +1784,24 @@ have size_gcdpq_const (s' : S') : {in P' s' &, move=> q r /imfsetP[/=] p _ -> {q} /imfsetP[/=] q _ -> {r} x y xs ys. have [p0|/negP p0] := eqVneq (evalpmp x (muni (val p))) 0. rewrite {1}p0; move/eqP: p0. - rewrite -size_poly_eq0 (S'size s' p x xs) -(S'size s' p y ys) size_poly_eq0 => /eqP {1}->. + rewrite -size_poly_eq0 (S'size s' p x xs) -(S'size s' p y ys). + rewrite size_poly_eq0 => /eqP {1}->. by rewrite !gcd0p (S'size s' q x xs) (S'size s' q y ys). have [q0|/negP q0] := eqVneq (evalpmp x (muni (val q))) 0. rewrite [X in gcdp _ X]q0; move/eqP: q0. - rewrite -size_poly_eq0 (S'size s' q x xs) -(S'size s' q y ys) size_poly_eq0 => /eqP {}q0. + rewrite -size_poly_eq0 (S'size s' q x xs) -(S'size s' q y ys). + rewrite size_poly_eq0 => /eqP {}q0. rewrite [X in _ = (size (gcdp _ X))]q0. by rewrite !gcdp0 (S'size s' p x xs) (S'size s' p y ys). rewrite -[LHS]prednK; last first. by rewrite ltnNge leqn0 size_poly_eq0 gcdp_eq0; apply/andP => -[_]. rewrite -[RHS]prednK; last first. rewrite ltnNge leqn0 size_poly_eq0 gcdp_eq0; apply/andP => -[_]. - by rewrite -size_poly_eq0 (S'size s' q y ys) -(S'size s' q x xs) size_poly_eq0. + rewrite -size_poly_eq0 (S'size s' q y ys) -(S'size s' q x xs). + by rewrite size_poly_eq0. apply/esym/eqP; rewrite eqSS. - move: (eqxx (size (gcdp (evalpmp x (muni (val p))) (evalpmp x (muni (val q))))).-1). + move: (eqxx + (size (gcdp (evalpmp x (muni (val p))) (evalpmp x (muni (val q))))).-1). rewrite gcdp_subresultant; first last. - exact/leq_predn/leq_gcdpr/negP. - exact/leq_predn/leq_gcdpl/negP/p0. @@ -1652,15 +1812,19 @@ rewrite gcdp_subresultant; first last. by apply/leq_predn/leq_gcdpr/negP. - rewrite (S'size s' p y ys) -(S'size s' p x xs). exact/leq_predn/leq_gcdpl/negP/p0. - - by rewrite -size_poly_eq0 (S'size s' q y ys) -(S'size s' q x xs) size_poly_eq0; apply/negP. - - by rewrite -size_poly_eq0 (S'size s' p y ys) -(S'size s' p x xs) size_poly_eq0; apply/negP/p0. + - rewrite -size_poly_eq0 (S'size s' q y ys) -(S'size s' q x xs). + by rewrite size_poly_eq0; apply/negP. + - rewrite -size_poly_eq0 (S'size s' p y ys) -(S'size s' p x xs). + by rewrite size_poly_eq0; apply/negP/p0. congr (_ && _). apply/eq_forallb => /= i. rewrite -sgz_eq0 -[RHS]sgz_eq0 (res_const s' p q x y xs ys)//. rewrite -[_ i]succnK -(S'size s' p x xs). - exact/leq_predn/(leq_trans (ltn_ord i))/(leq_trans (leq_pred _))/leq_gcdpl/negP/p0. + apply/leq_predn/(leq_trans (ltn_ord i))/(leq_trans (leq_pred _)). + exact/leq_gcdpl/negP/p0. rewrite -[_ i]succnK -(S'size s' q x xs). - exact/leq_predn/(leq_trans (ltn_ord i))/(leq_trans (leq_pred _))/leq_gcdpr/negP. + apply/leq_predn/(leq_trans (ltn_ord i))/(leq_trans (leq_pred _))/leq_gcdpr. + exact/negP. rewrite -sgz_eq0 -[in RHS]sgz_eq0 (res_const s' p q x y xs ys)//. by apply/leq_predn; rewrite -(S'size s' p x xs); apply/leq_gcdpl/negP/p0. by apply/leq_predn; rewrite -(S'size s' q x xs); apply/leq_gcdpr/negP. @@ -1679,8 +1843,10 @@ have S'const (s' : S') : have size_gcdpm_const (s' : S') : {in \val s', forall x : 'rV_n, - size (gcdp (evalpmp x (\prod_(p : P' s') \val p)) (evalpmp x (\prod_(p : P' s') \val p))^`()) = - size (gcdp (evalpmp (val (pick s')) (\prod_(p : P' s') \val p)) (evalpmp (val (pick s')) (\prod_(p : P' s') \val p))^`())}. + size (gcdp (evalpmp x (\prod_(p : P' s') \val p)) + (evalpmp x (\prod_(p : P' s') \val p))^`()) = + size (gcdp (evalpmp (val (pick s')) (\prod_(p : P' s') \val p)) + (evalpmp (val (pick s')) (\prod_(p : P' s') \val p))^`())}. case: (@evalpmp_prod_const _ (P' s') (val s')). - exact/S'con. - move=> q /imfsetP[/=] p _ -> {q} x y xs ys. @@ -1688,7 +1854,9 @@ have size_gcdpm_const (s' : S') : - exact/size_gcd_const. - exact/size_gcdpq_const. - move=> cc _ x xs; exact/(cc x _ xs (proj2_sig (pick s'))). -pose S := [fset SAset_cast n.+1 s' | s' in \big[fsetU/fset0]_(s' : S') partition_of_graphs_above (val s') (proj1_sig (roots2_on (S'const s')))]. +pose S := [fset SAset_cast n.+1 s' | + s' in \big[fsetU/fset0]_(s' : S') + partition_of_graphs_above (val s') (proj1_sig (roots2_on (S'const s')))]. have {}Scast : [fset SAset_cast n s | s in S] = S'. rewrite /S 2!imfset_bigfcup. apply/fsetP => x; apply/bigfcupP/idP => [[] /= s _| xS]. @@ -1699,13 +1867,18 @@ have {}Scast : [fset SAset_cast n s | s in S] = S'. exists [` xS]; first by rewrite mem_index_enum. apply/imfsetP => /=. case: (roots2_on (S'const [` xS])) => /= r [] rsort _. - exists (SAset_cast n.+1 ((nth (SAset0 R _) (partition_of_graphs r) 0) :&: (x :*: SAsetT R 1))). - apply/imfsetP; exists ((nth (SAset0 R _) (partition_of_graphs r) 0) :&: (x :*: SAsetT R 1)) => //=. - apply/imfsetP => /=; exists (nth (SAset0 R _) (partition_of_graphs r) 0) => //. + exists (SAset_cast n.+1 + ((nth (SAset0 R _) (partition_of_graphs r) 0) :&: (x :*: SAsetT R 1))). + apply/imfsetP. + exists ((nth (SAset0 R _) (partition_of_graphs r) 0) + :&: (x :*: SAsetT R 1)) => //=. + apply/imfsetP => /=. + exists (nth (SAset0 R _) (partition_of_graphs r) 0) => //. exact/mem_nth. rewrite SAset_cast_trans; last by rewrite geq_min addn1 leqnn. apply/esym/(SAset_cast_partition_of_graphs_above rsort). - apply/imfsetP => /=; exists (nth (SAset0 R _) (partition_of_graphs r) 0) => //. + apply/imfsetP => /=. + exists (nth (SAset0 R _) (partition_of_graphs r) 0) => //. exact/mem_nth. have ScastE (s : S') y : y \in partition_of_graphs_above (fsval s) (sval (roots2_on (S'const s))) -> @@ -1724,22 +1897,24 @@ have SCD: isCD S. case: rR rE => /= r [] rsort rroot rE. exists (size r), (in_tuple r); split. move=> i. - apply(@subspace_eq_continuous _ _ 'M[R]_(1, 1) (fun x => \row__ (rootsR (evalpmp x (\prod_(p : P' s) val p)))`_i)); last first. - apply/mx_continuous => j k. - apply(@subspace_eq_continuous _ _ R (fun x => (rootsR (evalpmp x (\prod_(p : P' s) val p)))`_i)). - by move=> x; rewrite inE/= => xs; rewrite mxE. - apply/(rootsR_continuous (proj2_sig (pick s))); first last. - - exact/S'const. - - move=> x xs; exact/(size_gcdpm_const s). - move=> x xs; rewrite ![evalpmp _ _]rmorph_prod/= !size_prod/=. - + congr (_.+1 - _)%N; apply/eq_bigr; case=> /= q /imfsetP[/=] p _ -> _. - exact/S'size. - + by case=> /= q /imfsetP[/=] p p0 -> _. - + case=> /= q /imfsetP[/=] p p0 -> _. - by rewrite -size_poly_eq0 (S'size s p x xs) size_poly_eq0. - move=> x; rewrite inE/= => xs. - apply/eqP; rewrite rowPE forall_ord1 mxE (tnth_nth 0)/=. - by rewrite -(rroot x xs) (nth_map 0). + apply/(@subspace_eq_continuous _ _ 'M[R]_(1, 1) + (fun x => \row__ (rootsR (evalpmp x (\prod_(p : P' s) val p)))`_i)). + move=> x; rewrite inE/= => xs. + apply/eqP; rewrite rowPE forall_ord1 mxE (tnth_nth 0)/=. + by rewrite -(rroot x xs) (nth_map 0). + apply/mx_continuous => j k. + apply(@subspace_eq_continuous _ _ R + (fun x => (rootsR (evalpmp x (\prod_(p : P' s) val p)))`_i)). + by move=> x; rewrite inE/= => xs; rewrite mxE. + apply/(rootsR_continuous (proj2_sig (pick s))); first last. + - exact/S'const. + - move=> x xs; exact/(size_gcdpm_const s). + move=> x xs; rewrite ![evalpmp _ _]rmorph_prod/= !size_prod/=. + + congr (_.+1 - _)%N; apply/eq_bigr; case=> /= q /imfsetP[/=] p _ -> _. + exact/S'size. + + by case=> /= q /imfsetP[/=] p p0 -> _. + + case=> /= q /imfsetP[/=] p p0 -> _. + by rewrite -size_poly_eq0 (S'size s p x xs) size_poly_eq0. split=> //. apply/fsetP => /= x; rewrite 2!inE/=. apply/andP/imfsetP. @@ -1761,7 +1936,9 @@ have tS: SAset_cast n.+1 t \in S. move: ts. case: (roots2_on (S'const s)) => /= [] r [] rsort rroot tpart. have mevalE q x : - q.@[tnth (ngraph x)] = (evalpmp (\row_i (x ord0 (widen_ord (leqnSn n) i))) (muni q)).[x ord0 ord_max]. + q.@[tnth (ngraph x)] + = (evalpmp (\row_i (x ord0 (widen_ord (leqnSn n) i))) (muni q)) + .[x ord0 ord_max]. rewrite horner_evalpmp -{1}(muniK q) meval_mmulti (tnth_nth 0) nth_ngraph. by apply/meval_eq => i; rewrite !tnth_mktuple mxE. have ts x: x \in SAset_cast n.+1 t -> @@ -1779,7 +1956,8 @@ have [p0|p0] := eqVneq (evalpmp (\val (pick s)) (muni (\val p))) 0. set q := (evalpmp _ _); suff ->: q = 0 by apply/horner0. apply/eqP; rewrite -size_poly_eq0 /q (S'size s); last exact/ts. by apply/eqP; rewrite p0 size_poly0. -case: (set0Vmem (SAimset (SAmpoly (val p)) (SAset_cast n.+1 t) :&: SAset_seq [:: 0])). +case: (set0Vmem + (SAimset (SAmpoly (val p)) (SAset_cast n.+1 t) :&: SAset_seq [:: 0])). move=> t0. have {}p0 x : x \in SAset_cast n.+1 t -> (val p).@[tnth (ngraph x)] != 0. move=> xt; apply/eqP => {}p0. @@ -1802,10 +1980,12 @@ case: (set0Vmem (SAimset (SAmpoly (val p)) (SAset_cast n.+1 t) :&: SAset_seq [:: have: SAconnected (SAimset (SAmpoly (val p)) (SAset_cast n.+1 t)). apply/SAimset_connected. exact/(@SAconnected_CD_cell _ _ [` tS] SCD). - apply/(@eq_continuous (subspace [set` SAset_cast n.+1 t]) _ (fun x : 'rV[R]_n.+1 => \row_(_ < 1) (val p).@[x ord0])). + apply/(@eq_continuous (subspace [set` SAset_cast n.+1 t]) _ + (fun x : 'rV[R]_n.+1 => \row_(_ < 1) (val p).@[x ord0])). by move=> z; rewrite SAmpolyE. apply/mx_continuous => i j. - apply/(@eq_continuous (subspace [set` SAset_cast n.+1 t]) _ (fun x : 'rV[R]_n.+1 => (val p).@[x ord0])). + apply/(@eq_continuous (subspace [set` SAset_cast n.+1 t]) _ + (fun x : 'rV[R]_n.+1 => (val p).@[x ord0])). by move=> z; rewrite mxE. exact/continuous_subspaceT/meval_continuous. move=> /(_ (SAset_itv `]-oo, 0[%R) (SAset_itv `]0, +oo[)); mp. @@ -1850,14 +2030,15 @@ case: (set0Vmem (SAimset (SAmpoly (val p)) (SAset_cast n.+1 t) :&: SAset_seq [:: move=> /negP; apply; rewrite -subset0; apply/SAset_subP => z. rewrite 2!inSAsetI !inSAset_itv/= !in_itv/= andbT => /andP[]/andP[_] z0. by move=> /(lt_trans z0); rewrite ltxx. -move=> [px]; rewrite inSAsetI inSAset_seq mem_seq1 => /andP[] /SAimsetP[x] xt -> {px}. +move=> [px]. +rewrite inSAsetI inSAset_seq mem_seq1 => /andP[] /SAimsetP[x] xt -> {px}. rewrite SAmpolyE rowPE forall_ord1 !mxE => /eqP px0. case: (roots2_on (S'constR s p)) => rp [] rpsort rproot. have f1_inj : injective (fun x : 'rV[R]_1 => x ord0 ord0). by move=> a b ab; apply/eqP; rewrite rowPE forall_ord1; apply/eqP. -have runiq: forall (r : seq (SAfunltType R n)), - sorted (SAfun_lt (n:=n)) r -> - {in \val s, forall x0 : 'rV_n, uniq [seq (f : {SAfun R^n -> R^1}) x0 | f <- r]}. +have runiq: forall (r : seq (SAfunltType R n)), sorted (SAfun_lt (n:=n)) r -> + {in \val s, forall x0 : 'rV_n, + uniq [seq (f : {SAfun R^n -> R^1}) x0 | f <- r]}. move=> r' r'sort y ys; rewrite -(map_inj_uniq f1_inj). apply/lt_sorted_uniq/(sortedP 0) => i. rewrite !size_map => ilt. @@ -1869,12 +2050,14 @@ mp; first exact/runiq. mp; first exact/runiq. mp. move=> i. - apply/(@subspace_eq_continuous _ _ _ (fun x : 'rV[R]_n => \row_(_ < 1) (rootsR (evalpmp x (muni (\val p))))`_i)). + apply/(@subspace_eq_continuous _ _ _ (fun x : 'rV[R]_n => + \row_(_ < 1) (rootsR (evalpmp x (muni (\val p))))`_i)). move=> y; rewrite mem_setE => ys. apply/eqP; rewrite rowPE forall_ord1 mxE. by rewrite -rproot// (nth_map 0). apply/mx_continuous => j k. - apply/(@eq_continuous (subspace [set` val s]) _ (fun x => (rootsR (evalpmp x (muni (\val p))))`_i)). + apply/(@eq_continuous (subspace [set` val s]) _ + (fun x => (rootsR (evalpmp x (muni (\val p))))`_i)). by move=> y; rewrite mxE. apply/(rootsR_continuous (valP (pick s))). - by move=> y ys; apply/S'size. @@ -1883,13 +2066,15 @@ mp. by move=> y ys; apply/S'constR. mp. move=> i. - apply/(@subspace_eq_continuous _ _ _ (fun x : 'rV[R]_n => \row_(_ < 1) (rootsR (evalpmp x (\prod_(p : P' s) fsval p)))`_i)). + apply/(@subspace_eq_continuous _ _ _ (fun x : 'rV[R]_n => + \row_(_ < 1) (rootsR (evalpmp x (\prod_(p : P' s) fsval p)))`_i)). move=> y; rewrite mem_setE => ys. apply/eqP; rewrite rowPE forall_ord1 mxE. move: (rroot y ys) => /(congr1 (fun x => x`_i)). by rewrite (nth_map 0)// => ->; apply/eqP. apply/mx_continuous => j k. - apply/(@eq_continuous (subspace [set` val s]) _ (fun x => (rootsR (evalpmp x (\prod_(p : P' s) fsval p)))`_i)). + apply/(@eq_continuous (subspace [set` val s]) _ + (fun x => (rootsR (evalpmp x (\prod_(p : P' s) fsval p)))`_i)). by move=> y; rewrite mxE. apply/(rootsR_continuous (valP (pick s))). - move=> y ys. @@ -1912,7 +2097,9 @@ mp. rewrite (nth_map 0)//=. by move: rpsort => /(sortedP 0)/(_ i ilt)/SAfun_ltP/(_ y). - - apply/(@subseq_lt_sorted _ _ _ [seq (x0 : 'rV_1) ord0 ord0 | x0 <- [seq (f : {SAfun R^n -> R^1}) y | f <- r]]). + apply/(@subseq_lt_sorted _ _ _ + [seq (x0 : 'rV_1) ord0 ord0 | x0 <- + [seq (f : {SAfun R^n -> R^1}) y | f <- r]]). exact/map_subseq/filter_subseq. apply/(sortedP 0) => i; rewrite 2!size_map => ilt. rewrite -map_comp (nth_map 0); last by apply/(ltn_trans _ ilt). @@ -1924,10 +2111,12 @@ mp. move=> _ /(nthP 0)[] i irp <- ->. exists (rp`_i y) => //; rewrite mem_filter; apply/andP; split. by apply/mapP; exists (rp`_i) => //; apply/mem_nth. - rewrite -(mem_map f1_inj) -map_comp/= rroot// in_rootsR evalpmp_prod; apply/andP; split. + rewrite -(mem_map f1_inj) -map_comp/= rroot// in_rootsR evalpmp_prod. + apply/andP; split. apply/prodf_neq0; case=> /= _ /imfsetP[/=] q q0 -> _. by rewrite -size_poly_eq0 (S'size s)// size_poly_eq0. - have: rp`_i y ord0 ord0 \in [seq (xi : {SAfun R^n -> R^1}) y ord0 ord0 | xi <- rp]. + have: rp`_i y ord0 ord0 \in + [seq (xi : {SAfun R^n -> R^1}) y ord0 ord0 | xi <- rp]. by apply/mapP; exists (rp`_i) => //; apply/mem_nth. rewrite rproot// in_rootsR => /andP[_] i0. rewrite root_bigmul; apply/hasP. @@ -1935,7 +2124,9 @@ mp. exists [` pP]; first exact/mem_index_enum. exact: i0. move=> [rpr] [/mem_subseq rprr] rpE. -have: \row__ x ord0 ord_max \in [seq (f : {SAfun R^n -> R^1}) (\row_i (x ord0 (widen_ord (leqnSn n) i))) | f <- rp]. +have: \row__ x ord0 ord_max \in + [seq (f : {SAfun R^n -> R^1}) (\row_i (x ord0 (widen_ord (leqnSn n) i))) | + f <- rp]. rewrite -(mem_map f1_inj) -map_comp/= rproot; last exact/ts. rewrite mxE in_rootsR; apply/andP; split. by rewrite -size_poly_eq0 (S'size s) ?size_poly_eq0//; apply/ts. @@ -1975,18 +2166,21 @@ have [_ tE|jr] := eqVneq j (size r).*2. - exact/val_inj. rewrite xE; congr (r`_(rpr`_i)%R _ _ _). by apply/rowP => k; rewrite !mxE castmxE; congr (x _ _); apply/val_inj. -have xlE a: lsubmx (castmx (erefl 1%N, esym (addn1 n)) a) = \row_i a ord0 (widen_ord (leqnSn n) i). +have xlE a: lsubmx (castmx (erefl 1%N, esym (addn1 n)) a) + = \row_i a ord0 (widen_ord (leqnSn n) i). by apply/rowP => k; rewrite !mxE castmxE; congr (a _ _); apply/val_inj. -have xrE (a : 'rV[R]_n.+1): a (cast_ord (esym ((erefl 1%N, esym (addn1 n)).1)%PAIR) ord0) - (cast_ord (esym ((erefl 1%N, esym (addn1 n)).2)%PAIR) (rshift n ord0)) - = a ord0 ord_max. +have xrE (a : 'rV[R]_n.+1): + a (cast_ord (esym ((erefl 1%N, esym (addn1 n)).1)%PAIR) ord0) + (cast_ord (esym ((erefl 1%N, esym (addn1 n)).2)%PAIR) (rshift n ord0)) + = a ord0 ord_max. by congr (a _ _); apply/val_inj => //; apply/addn0. move: jlt; rewrite ltnS leq_eqVlt (negPf jr)/= => jlt. case: ifP => jodd tE; last first. move: xt; rewrite tE (inSAset_cast _ _ (esym (addn1 n))) !inSAsetI. rewrite inSAepigraph inSAhypograph !mxE castmxE xlE xrE. rewrite xE => /andP[]/andP[] rji rij _. - have {}rsort: {in gtn (size r) &, {homo nth 0 r : i j / (i <= j)%N >-> SAfun_le i j}}. + have {}rsort: {in gtn (size r) &, + {homo nth 0 r : i j / (i <= j)%N >-> SAfun_le i j}}. move=> b c br cr bc. have: sorted (SAfun_le (n:=n)) r. move: rsort; apply/sub_sorted => f g /SAfun_ltP fg. @@ -1999,11 +2193,13 @@ case: ifP => jodd tE; last first. rewrite succnK => ij. move: (rsort (rpr`_i)%R j./2.-1); rewrite !inE prednK; last first. by case: j j0 jodd {jlt jr tE rji rij ij} => [//|]; case. - rewrite leq_half_double => /(_ ir (ltnW (leq_trans jlt (leqnSn _))) ij) /SAfun_leP. + rewrite leq_half_double. + move=> /(_ ir (ltnW (leq_trans jlt (leqnSn _))) ij) /SAfun_leP. move=> /(_ (\row_i x ord0 (widen_ord (leqnSn n) i)))/(lt_le_trans rji). by rewrite ltxx. move: (rsort j./2 (rpr`_i)%R); rewrite !inE ltn_half_double. - move=> /(_ jlt ir ji) /SAfun_leP/(_ (\row_i x ord0 (widen_ord (leqnSn n) i)))/(lt_le_trans rij). + move=> /(_ jlt ir ji) /SAfun_leP/(_ (\row_i x ord0 (widen_ord (leqnSn n) i))). + move=> /(lt_le_trans rij). by rewrite ltxx. move: rsort; rewrite sorted_pairwise; last exact/SAfun_lt_trans. move=> /(pairwiseP 0) rsort. @@ -2027,13 +2223,16 @@ suff r0 a: a \in SAset_cast n.+1 (r`_(rpr`_i)%R :&: fsval s :*: SAsetT R 1) -> move=> {y z yr zr}; rewrite (inSAset_cast _ _ (esym (addn1 n))) inSAsetI. rewrite -[castmx _ _]hsubmxK -inSAfun mevalE xlE inSAsetX row_mxKl. move=> /andP[] /eqP aE /andP[] las _. -have: rsubmx (castmx (erefl 1%N, esym (addn1 n)) a) \in [seq r`_i (\row_i a ord0 (widen_ord (leqnSn n) i)) | i <- rpr]. +have: rsubmx (castmx (erefl 1%N, esym (addn1 n)) a) \in + [seq r`_i (\row_i a ord0 (widen_ord (leqnSn n) i)) | i <- rpr]. rewrite -aE; apply/map_f/mem_nth. by move: (rpE _ (ts _ xt)) => /(congr1 size); rewrite !size_map => <-. rewrite -rpE// => /(nthP 0)[] j; rewrite size_map => jlt. rewrite (nth_map 0)// => /(congr1 (fun x : 'rV_1 => x ord0 ord0)). rewrite !mxE castmxE xrE => {}aE. -have: a ord0 ord_max \in rootsR (evalpmp (\row_i0 a ord0 (widen_ord (leqnSn n) i0)) (muni (fsval p))). +have: a ord0 ord_max \in + rootsR (evalpmp (\row_i0 a ord0 (widen_ord (leqnSn n) i0)) + (muni (fsval p))). by rewrite -rproot// -aE; apply/map_f/mem_nth. by rewrite in_rootsR => /andP[_] /eqP. Qed. diff --git a/formula.v b/formula.v index 358c9c5..34c560d 100644 --- a/formula.v +++ b/formula.v @@ -2,11 +2,11 @@ Require Import ZArith Init. From HB Require Import structures. Require Import mathcomp.ssreflect.ssreflect. -From mathcomp Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype div. -From mathcomp Require Import tuple finfun generic_quotient bigop finset perm. -From mathcomp Require Import ssralg poly polydiv ssrnum mxpoly binomial finalg. -From mathcomp Require Import zmodp mxpoly mxtens qe_rcf ordered_qelim realalg. -From mathcomp Require Import matrix finmap order finset mpoly. +From mathcomp Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import div tuple finfun generic_quotient bigop finset. +From mathcomp Require Import perm ssralg poly polydiv ssrnum mxpoly binomial. +From mathcomp Require Import finalg zmodp mxpoly mxtens qe_rcf ordered_qelim. +From mathcomp Require Import realalg matrix finmap order finset mpoly. From SemiAlgebraic Require Import auxresults. @@ -60,9 +60,10 @@ by rewrite in_fset1 mem_seq1. Qed. Lemma mnfsetU (i j k l : nat) : -let a := minn i k in -(i <= k + l -> k <= i + j -> - mnfset i j `|` mnfset k l = mnfset a ((maxn (i + j) (k + l)) - a))%N. + let a := minn i k in + (i <= k + l + -> k <= i + j + -> mnfset i j `|` mnfset k l = mnfset a ((maxn (i + j) (k + l)) - a))%N. Proof. move=> a h1 h2. apply/eqP/fset_eqP => x. @@ -103,20 +104,19 @@ Lemma mnfset_sub (a b c d : nat) : b != O -> (mnfset a b `<=` mnfset c d) = ((c <= a) && (a + b <= c + d))%N. Proof. case: b => // b _; case: d. -- rewrite addn0; apply/idP/idP. - + by move/fsubsetP/(_ a); rewrite !seq_fsetE in_fset0 inE eqxx; move/implyP. - + move=> /andP [leq_ca leq__c]. - by move: (leq_trans leq__c leq_ca); rewrite leqNgt addnS ltnS /= leq_addr. -- move=> d; apply/fsubsetP/idP; last first. - + move/andP=> [leq_ca leq_bd] x; rewrite !mnfsetE; move/andP => [leq_ax lt_xb]. - by rewrite (leq_trans leq_ca) // (leq_trans lt_xb). - + move=> h. - apply/andP; split; - [move/(_ a) : h | move/(_ (a + b)%N) : h]; rewrite !mnfsetE. - - rewrite leqnn addnS ltnS leq_addr; move/implyP. - by rewrite implyTb => /andP []. - - rewrite /= addnS ltnS leq_addr leqnn. - by move/implyP; rewrite andbT => /andP []. + rewrite addn0; apply/idP/idP. + by move/fsubsetP/(_ a); rewrite !seq_fsetE in_fset0 inE eqxx; move/implyP. + move=> /andP [leq_ca leq__c]. + by move: (leq_trans leq__c leq_ca); rewrite leqNgt addnS ltnS /= leq_addr. +move=> d; apply/fsubsetP/idP; last first. + move/andP => [leq_ca leq_bd] x; rewrite !mnfsetE; move/andP => [leq_ax lt_xb]. + rewrite (leq_trans leq_ca) // (leq_trans lt_xb)//. +move=> h. +apply/andP; split; [move/(_ a) : h | move/(_ (a + b)%N) : h]; rewrite !mnfsetE. + rewrite leqnn addnS ltnS leq_addr; move/implyP. + by rewrite implyTb => /andP []. +rewrite /= addnS ltnS leq_addr leqnn. +by move/implyP; rewrite andbT => /andP []. Qed. Lemma m0fset (m : nat) : mnfset m 0 = fset0. @@ -138,12 +138,14 @@ rewrite eqEfsubset !mnfset_sub // andbACA -!eqn_leq eq_sym. by have [->|//] := altP (a =P c); rewrite eqn_add2l. Qed. -Local Lemma set_nth_size (T : Type) (d : T) (n : nat) (x : n.-tuple T) (i : 'I_n) (y : T) : +Local Lemma set_nth_size (T : Type) (d : T) (n : nat) + (x : n.-tuple T) (i : 'I_n) (y : T) : size (set_nth d x i y) == n. Proof. by rewrite size_set_nth size_tuple; apply/eqP/maxn_idPr. Qed. -Canonical set_nth_tuple (T : Type) (d : T) (n : nat) (x : n.-tuple T) (i : 'I_n) (y : T) := - Tuple (set_nth_size d x i y). +Canonical set_nth_tuple (T : Type) (d : T) (n : nat) + (x : n.-tuple T) (i : 'I_n) (y : T) := + Tuple (set_nth_size d x i y). End SeqFset. @@ -155,10 +157,11 @@ Import GRing. Definition term_poly (P : {poly R}) (t : term R) : term R := \big[Add/Const 0]_(i < size P) Mul (Const P`_i) (Exp t i). -Lemma eval_big (op : term R -> term R -> term R) (x : term R) (I : Type) (r : seq I) - (P : pred I) (F : I -> term R) (e : seq R) (eop : R -> R -> R) : +Lemma eval_big (op : term R -> term R -> term R) (x : term R) (I : Type) + (r : seq I) (P : pred I) (F : I -> term R) (e : seq R) (eop : R -> R -> R) : (forall t u, eval e (op t u) = eop (eval e t) (eval e u)) - -> eval e (\big[op/x]_(i <- r | P i) F i) = \big[eop/eval e x]_(i <- r | P i) eval e (F i). + -> eval e (\big[op/x]_(i <- r | P i) F i) + = \big[eop/eval e x]_(i <- r | P i) eval e (F i). Proof. move=> opE. elim: r => [|i r IHr]; first by rewrite !big_nil. @@ -171,9 +174,11 @@ Lemma eval_term_poly (e : seq R) (P : {poly R}) (t : term R) : Proof. by rewrite -{2}[P]coefK horner_poly (eval_big _ _ _ _ (eop:=+%R)). Qed. Definition term_mpoly n (P : {mpoly R[n]}) (t : 'I_n -> term R) : term R := - \big[Add/Const 0]_(m <- msupp P) Mul (Const P@_m) (\big[Mul/Const 1]_i Exp (t i) (m i)). + \big[Add/Const 0]_(m <- msupp P) Mul (Const P@_m) + (\big[Mul/Const 1]_i Exp (t i) (m i)). -Definition eval_term_mpoly n (P : {mpoly R[n]}) (t : 'I_n -> term R) (e : seq R) : +Definition eval_term_mpoly n (P : {mpoly R[n]}) + (t : 'I_n -> term R) (e : seq R) : eval e (term_mpoly P t) = (P.@[fun i => eval e (t i)])%R. Proof. rewrite mevalE (eval_big _ _ _ _ (eop:=+%R))//. @@ -231,7 +236,8 @@ Fixpoint mpoly_rterm (R : unitRingType) (n : nat) (t : term R) : {mpoly R[n]} := | Exp t i => (mpoly_rterm n t) ^+ i end. -Lemma meval_mpoly_rterm (R : realDomainType) (n : nat) (x : 'I_n -> R) (t : term R) : +Lemma meval_mpoly_rterm (R : realDomainType) (n : nat) + (x : 'I_n -> R) (t : term R) : (mpoly_rterm n t).@[x] = eval [seq x i | i <- enum 'I_n] t. Proof. elim: t => /=. @@ -292,7 +298,7 @@ Fixpoint gen_var_seq (s : seq nat) (f : formula T) := match s with end%oT. Definition equiv_formula (f g : formula T) := - gen_var_seq (enum_fset ((formula_fv f) `|` (formula_fv g))) (f <==> g)%oT. + gen_var_seq (enum_fset ((formula_fv f) `|` (formula_fv g))) (f <==> g)%oT. Definition nvar n := fun f : formula T => (formula_fv f `<=` mnfset O n). @@ -333,11 +339,11 @@ Variable (R : Type). Fact nquantify_key : unit. Proof. exact: tt. Qed. Definition nquantify (n k : nat) (Q : nat -> formula R -> formula R) - (f : formula R) := - locked_with nquantify_key (iteri k (fun i f => (Q (n + k - i.+1)%N f)) f). + (f : formula R) := + locked_with nquantify_key (iteri k (fun i f => (Q (n + k - i.+1)%N f)) f). Lemma nquantSout (n k : nat) Q (f : formula R) : - nquantify n k.+1 Q f = Q n (nquantify n.+1 k Q f). + nquantify n k.+1 Q f = Q n (nquantify n.+1 k Q f). Proof. rewrite /nquantify !unlock /= addnK; congr (Q _ _); apply: eq_iteri => i g. by rewrite addnS addSn. @@ -350,7 +356,7 @@ Lemma nquantify1 (n : nat) Q (f : formula R) : nquantify n 1 Q f = Q n f. Proof. by rewrite nquantSout nquantify0. Qed. Lemma nquantify_add (m n k : nat) Q (f : formula R) : - nquantify m (n + k) Q f = nquantify m n Q (nquantify (m + n) k Q f). + nquantify m (n + k) Q f = nquantify m n Q (nquantify (m + n) k Q f). Proof. elim: n => [|n IHn] in k m *; rewrite ?(nquantify0, nquantSout, addn0, addSn) //=. @@ -358,7 +364,7 @@ by rewrite IHn addnS addSn. Qed. Lemma nquantSin (n k : nat) Q (f : formula R) : - nquantify n k.+1 Q f = (nquantify n k Q (Q (n + k)%N f)). + nquantify n k.+1 Q f = (nquantify n k Q (Q (n + k)%N f)). Proof. by rewrite -addn1 nquantify_add nquantify1. Qed. Lemma formula_fv_nforall (n k : nat) (f : formula R) : @@ -522,7 +528,7 @@ Fact fv_tsubst_nil (t : GRing.term F) : term_fv (subst_term [::] t) = fset0. Proof. by elim: t => //= t1 -> t2 ->; rewrite fsetU0. Qed. Fact fv_tsubst (k : unit) (s : seq nat) (t : GRing.term F) : - term_fv (subst_term s t) `<=` seq_fset k s. + term_fv (subst_term s t) `<=` seq_fset k s. Proof. elim: t => //. - move=> i /=. @@ -628,7 +634,7 @@ by rewrite fset1D1 eq_sym neq_ij. Qed. Lemma formula_fv_fsubst (i : nat) (x : T) (f : formula T) : - formula_fv (fsubst f (i, (x%:T)%oT)) = (formula_fv f) `\ i. + formula_fv (fsubst f (i, (x%:T)%oT)) = (formula_fv f) `\ i. Proof. elim: f. + by move=> b; rewrite fset0D. @@ -705,20 +711,21 @@ Qed. Definition is_equiv (f g : formula R) := holds [::] (equiv_formula f g). Lemma holds_rcons_zero (e : seq R) (f : formula R) : - holds (rcons e 0%:R) f <-> holds e f. + holds (rcons e 0%:R) f <-> holds e f. Proof. split; apply: eq_holds=> // i; rewrite nth_rcons; by have [| /ltnW h|->] := ltngtP _ (size _)=> //; rewrite ?nth_default. Qed. Lemma holds_cat_nseq (i : nat) (e : seq R) (f : formula R) : - holds (e ++ (nseq i 0)) f <-> holds e f. + holds (e ++ (nseq i 0)) f <-> holds e f. Proof. rewrite nseq_cat; move: e f; elim: i => // i ih e f. by apply: (iff_trans _ (ih e f)); apply: holds_rcons_zero. Qed. -Lemma holdsAnd (I : eqType) (r : seq I) (P : pred I) (e : seq R) (f : I -> formula R) : +Lemma holdsAnd (I : eqType) (r : seq I) (P : pred I) + (e : seq R) (f : I -> formula R) : holds e (\big[And/True%oT]_(i <- r | P i) f i) <-> forall i, i \in r -> P i -> holds e (f i). Proof. @@ -733,7 +740,8 @@ split; first by apply: hr => //; rewrite in_cons eq_refl. by apply/IHr => j jr; apply: hr; rewrite in_cons jr orbT. Qed. -Lemma holdsOr (I : eqType) (r : seq I) (P : pred I) (e : seq R) (f : I -> formula R) : +Lemma holdsOr (I : eqType) (r : seq I) (P : pred I) + (e : seq R) (f : I -> formula R) : holds e (\big[Or/False%oT]_(i <- r | P i) f i) <-> exists i, i \in r /\ P i /\ holds e (f i). Proof. @@ -776,7 +784,7 @@ by exists 0; rewrite set_nth_over //; apply/holds_rcons_zero/holds_cat_nseq. Qed. Lemma fv0_holds (e : seq R) f : - formula_fv f = fset0 -> (holds e f <-> holds [::] f). + formula_fv f = fset0 -> (holds e f <-> holds [::] f). Proof. move/eqP; move=> h; elim/last_ind: e => //. move=> e x <-; move: h; elim: f => //. @@ -826,8 +834,8 @@ move=> e x <-; move: h; elim: f => //. Qed. Lemma nforallP (k : nat) (e : seq R) (f : formula R) : - (forall v : k.-tuple R, holds (e ++ v) f) - <-> (holds e (nquantify (size e) k Forall f)). + (forall v : k.-tuple R, holds (e ++ v) f) + <-> (holds e (nquantify (size e) k Forall f)). Proof. elim: k => [|k IHk] /= in e *. rewrite nquantify0; split. @@ -844,8 +852,8 @@ by rewrite /e' /= rcons_set_nth cat_rcons. Qed. Lemma nexistsP (k : nat) (e : seq R) (f : formula R) : - (exists v : k.-tuple R, holds (e ++ v) f) - <-> (holds e (nquantify (size e) k Exists f)). + (exists v : k.-tuple R, holds (e ++ v) f) + <-> (holds e (nquantify (size e) k Exists f)). Proof. elim: k => [|k IHk] /= in e *. - rewrite nquantify0; split; first by move=> [v]; rewrite tuple0 cats0. @@ -862,26 +870,26 @@ elim: k => [|k IHk] /= in e *. Qed. Lemma nforall_is_true (f : formula R) : - (forall (e : seq R), holds e f) -> - forall (n i : nat) (e : seq R), holds e (nquantify n i Forall f). + (forall (e : seq R), holds e f) -> + forall (n i : nat) (e : seq R), holds e (nquantify n i Forall f). Proof. move=> h n i; elim: i => [|i IHi] in n *; by rewrite ?(nquantify0, nquantSout) /=. Qed. Lemma monotonic_forall_if (i : nat) (e : seq R) (f g : formula R) : -(forall (e' : seq R), holds e' f -> holds e' g) -> -holds e ('forall 'X_i, f) -> holds e ('forall 'X_i, g). + (forall (e' : seq R), holds e' f -> holds e' g) -> + holds e ('forall 'X_i, f) -> holds e ('forall 'X_i, g). Proof. by move=> h /= fholds x; apply/h. Qed. Lemma monotonic_exists_if (i : nat) (e : seq R) (f g : formula R) : -(forall (e' : seq R), holds e' f -> holds e' g) -> -holds e ('exists 'X_i, f) -> holds e ('exists 'X_i, g). + (forall (e' : seq R), holds e' f -> holds e' g) -> + holds e ('exists 'X_i, f) -> holds e ('exists 'X_i, g). Proof. by move=> h /= [x fx]; exists x; apply/h. Qed. Lemma monotonic_nforall (n k : nat) (e : seq R) (f g : formula R) : - (forall (e' : seq R), holds e' f -> holds e' g) -> - holds e (nquantify n k Forall f) -> holds e (nquantify n k Forall g). + (forall (e' : seq R), holds e' f -> holds e' g) -> + holds e (nquantify n k Forall f) -> holds e (nquantify n k Forall g). Proof. move: n e f g; elim: k => [k e f g | k ih n e f g h]. by rewrite !nquantify0; move/(_ e). @@ -891,8 +899,8 @@ exact/monotonic_forall_if. Qed. Lemma monotonic_nexist (n k : nat) (e : seq R) (f g : formula R) : - (forall (e' : seq R), holds e' f -> holds e' g) -> - holds e (nquantify n k Exists f) -> holds e (nquantify n k Exists g). + (forall (e' : seq R), holds e' f -> holds e' g) -> + holds e (nquantify n k Exists f) -> holds e (nquantify n k Exists g). Proof. move: n e f g; elim: k => [k e f g | k ih n e f g h]. by rewrite !nquantify0; move/(_ e). @@ -902,16 +910,16 @@ exact/monotonic_exists_if. Qed. Fact monotonic_forall_iff (i : nat) (e : seq R) (f g : formula R) : -(forall (e' : seq R), holds e' f <-> holds e' g) -> -holds e ('forall 'X_i, f) <-> holds e ('forall 'X_i, g). + (forall (e' : seq R), holds e' f <-> holds e' g) -> + holds e ('forall 'X_i, f) <-> holds e ('forall 'X_i, g). Proof. by move=> h; split; apply: monotonic_forall_if=> e'; move/(h e'). Qed. Fact holds_forall (i : nat) (e : seq R) (f : formula R) : - holds e ('forall 'X_i, f) -> holds e f. + holds e ('forall 'X_i, f) -> holds e f. Proof. by move=> /= /(_ e`_i); rewrite set_nth_nth; move/holds_cat_nseq. Qed. Lemma holds_nforall (n k : nat) (e : seq R) (f : formula R) : - holds e (nquantify n k Forall f) -> holds e f. + holds e (nquantify n k Forall f) -> holds e f. Proof. move: e f; elim: k => [e f|k iHk e f h]; first by rewrite nquantify0. apply: iHk; move: h; rewrite nquantSin; apply/monotonic_nforall => e'. @@ -919,11 +927,11 @@ exact/holds_forall. Qed. Fact closed_nforall_formulan (n : nat) (f : {formula_n R}) : - formula_fv (nquantify O n Forall f) == fset0. + formula_fv (nquantify O n Forall f) == fset0. Proof. by rewrite formula_fv_nforall fsetD_eq0 fsubset_formulan_fv. Qed. Fact closed_nexists_formulan (n : nat) (f : {formula_n R}) : - formula_fv (nquantify O n Exists f) == fset0. + formula_fv (nquantify O n Exists f) == fset0. Proof. by rewrite formula_fv_nexists fsetD_eq0 fsubset_formulan_fv. Qed. End RealDomainFormula. @@ -944,7 +952,8 @@ by move=> h; move/andP: (quantifier_elim_wf (@wf_QE_wproj _) h) => [_ rform_f]. Qed. Fact elim_rformP (f : formula F) (e : seq F) : -rformula f -> reflect (holds e f) (qf_eval e (@quantifier_elim _ (@wproj _) f)). + rformula f + -> reflect (holds e f) (qf_eval e (@quantifier_elim _ (@wproj _) f)). Proof. move=> rform_f; apply: quantifier_elim_rformP => //. - move=> i bc /= h. @@ -986,10 +995,11 @@ Proof. by []. Qed. Fact rcf_sat_Implies (e : seq F) (f g : formula F) : rcf_sat e (f ==> g) = ((rcf_sat e f) ==> (rcf_sat e g)). Proof. -by apply/rcf_satP/implyP=> /= hfg; move/rcf_satP=> h; apply/rcf_satP; apply: hfg. +by apply/rcf_satP/implyP => /= hfg; move/rcf_satP => h; apply/rcf_satP/hfg. Qed. -Fact rcf_sat_Not (e : seq F) (f : formula F): rcf_sat e (~ f) = ~~ (rcf_sat e f). +Fact rcf_sat_Not (e : seq F) (f : formula F) : + rcf_sat e (~ f) = ~~ (rcf_sat e f). Proof. by []. Qed. (* TODO: generalize to sequences with `has`? *) @@ -1012,7 +1022,7 @@ by move=> /nexistsP[v] /[dup] /rcf_satP/(u_uniq _ (size_tuple v)) ->. Qed. Lemma n_forall_formula (e : seq F) (f : formula F) (i : nat) : - holds e (~ ('forall 'X_i, f)) <-> holds e ('exists 'X_i, ~ f). + holds e (~ ('forall 'X_i, f)) <-> holds e ('exists 'X_i, ~ f). Proof. split; last by move=> [x hx] h2; apply/hx/h2. move=> /nn_formula/rcf_satP Nallf. @@ -1023,7 +1033,7 @@ by apply/rcf_satP => /= Nf_holds; apply: NexNf; exists x. Qed. Lemma n_nforall_formula (e : seq F) (f : formula F) (a b : nat) : - holds e (~ (nquantify a b Forall f)) <-> holds e (nquantify a b Exists (~ f)). + holds e (~ (nquantify a b Forall f)) <-> holds e (nquantify a b Exists (~ f)). Proof. move: f; elim: b => [f|b ih f]; first by rewrite !nquantify0. rewrite !nquantSin; split. @@ -1034,9 +1044,9 @@ rewrite !nquantSin; split. exact: (iffRL (n_forall_formula _ _ _)). Qed. -Definition simp_rcf_sat := (rcf_sat_Bool, rcf_sat_Equal, rcf_sat_Lt, rcf_sat_Le, - rcf_sat_Unit, rcf_sat_And, rcf_sat_Or, - rcf_sat_Implies, rcf_sat_Not). +Definition simp_rcf_sat := + (rcf_sat_Bool, rcf_sat_Equal, rcf_sat_Lt, rcf_sat_Le, rcf_sat_Unit, + rcf_sat_And, rcf_sat_Or, rcf_sat_Implies, rcf_sat_Not). Lemma rcf_sat_cat_nseq (i : nat) (e : seq F) (f : formula F) : rcf_sat (e ++ nseq i 0) f = rcf_sat e f. @@ -1050,7 +1060,7 @@ Lemma rcf_sat_take [n : nat] (f : {formula_n F}) (e : seq F) : Proof. by apply/rcf_satP/rcf_satP => /holds_take. Qed. Lemma rcf_sat_forall k l (E : 'I_k -> formula F) : - rcf_sat l (\big[And/True%oT]_(i < k) E i) = [forall i, rcf_sat l (E i)]. + rcf_sat l (\big[And/True%oT]_(i < k) E i) = [forall i, rcf_sat l (E i)]. Proof. elim: k=> [|k Ihk] in E *. by rewrite big_ord0 simp_rcf_sat; symmetry; apply/forallP => -[]. @@ -1059,7 +1069,7 @@ by rewrite -![qf_eval _ _]/(rcf_sat _ _) Ihk -(big_andE xpredT). Qed. Lemma rcf_sat_forallP k l (E : 'I_k -> formula F) : - rcf_sat l (\big[And/True%oT]_(i < k) E i) = [forall i, rcf_sat l (E i)]. + rcf_sat l (\big[And/True%oT]_(i < k) E i) = [forall i, rcf_sat l (E i)]. Proof. elim: k=> [|k Ihk] in E *. by rewrite big_ord0; apply/rcf_satP/forallP; move=> _ // [[ ]]. @@ -1071,31 +1081,32 @@ apply/rcf_satP; rewrite simp_rcf_sat Eall Ihk. by apply/forallP=> x; apply: Eall. Qed. -Lemma formula_fv_bigAnd (I : Type) (r : seq I) (P : pred I) - (E : I -> formula F) : -formula_fv (\big[And/True%oT]_(i <- r | P i) (E i)) = -\bigcup_(i <- r | P i) (formula_fv (E i)). +Lemma formula_fv_bigAnd (I : Type) (r : seq I) + (P : pred I) (E : I -> formula F) : + formula_fv (\big[And/True%oT]_(i <- r | P i) (E i)) = + \bigcup_(i <- r | P i) (formula_fv (E i)). Proof. exact: big_morph. Qed. -Lemma formula_fv_bigOr (I : Type) (r : seq I) (P : pred I) (E : I -> formula F) : -formula_fv (\big[Or/False%oT]_(i <- r | P i) (E i)) = -\bigcup_(i <- r | P i) (formula_fv (E i)). +Lemma formula_fv_bigOr (I : Type) (r : seq I) + (P : pred I) (E : I -> formula F) : + formula_fv (\big[Or/False%oT]_(i <- r | P i) (E i)) = + \bigcup_(i <- r | P i) (formula_fv (E i)). Proof. exact: big_morph. Qed. Lemma formula_fv_bigU (a : nat) (E : 'I_a -> formula F) : -formula_fv (\big[And/True%oT]_(i < a) (E i)) = -\bigcup_(i in 'I_a) (formula_fv (E i)). + formula_fv (\big[And/True%oT]_(i < a) (E i)) = + \bigcup_(i in 'I_a) (formula_fv (E i)). Proof. exact: big_morph. Qed. Definition is_independent (i : nat) (f : formula F) := -forall (e : seq F), holds e ('forall 'X_i, f) <-> holds e f. + forall (e : seq F), holds e ('forall 'X_i, f) <-> holds e f. Lemma independent (f : formula F) (i : nat) : i \notin (formula_fv f) -> is_independent i f. Proof. by rewrite /is_independent; case: f => *; apply: holds_Nfv_all. Qed. Lemma fsubst_indep (i : nat) (f : formula F) (x : F) (e : seq F) : - is_independent i f -> (holds e f) -> holds e (fsubst f (i, GRing.Const x)). + is_independent i f -> (holds e f) -> holds e (fsubst f (i, GRing.Const x)). Proof. by move=> h1 h2; apply/holds_fsubst; move/h1/(_ x): h2. Qed. Lemma is_independentP (i : nat) (f : formula F) : @@ -1128,7 +1139,7 @@ by apply: (iff_trans (ha _ _ e`_a)); rewrite set_nth_nth; apply/holds_cat_nseq. Qed. Lemma indep_to_rform (f : formula F) (i : nat) : - is_independent i (to_rform f) <-> is_independent i f. + is_independent i (to_rform f) <-> is_independent i f. Proof. split=> h e. + apply: (iff_trans _ (to_rformP _ _)). @@ -1148,14 +1159,14 @@ Variable (F : Type) (n : nat). Lemma and_formulan (f1 f2 : {formula_n F}) : nvar n (f1 /\ f2)%oT. Proof. by rewrite /nvar fsubUset !fsubset_formulan_fv. Qed. -Canonical Structure formulan_and (f1 f2 : {formula_n F}) - := MkFormulan (and_formulan f1 f2). +Canonical Structure formulan_and (f1 f2 : {formula_n F}) := + MkFormulan (and_formulan f1 f2). Lemma implies_formulan (f1 f2 : {formula_n F}) : nvar n (f1 ==> f2)%oT. Proof. by rewrite /nvar fsubUset !fsubset_formulan_fv. Qed. Canonical Structure formulan_implies (f1 f2 : {formula_n F}) := - MkFormulan (implies_formulan f1 f2). + MkFormulan (implies_formulan f1 f2). Lemma bool_formulan (b : bool) : @nvar F n (Bool b). Proof. by rewrite /nvar fsub0set. Qed. @@ -1166,7 +1177,7 @@ Lemma or_formulan (f1 f2 : {formula_n F}) : nvar n (f1 \/ f2)%oT. Proof. by rewrite /nvar fsubUset !fsubset_formulan_fv. Qed. Canonical Structure formulan_or (f1 f2 : {formula_n F}) := - MkFormulan (or_formulan f1 f2). + MkFormulan (or_formulan f1 f2). Lemma not_formulan (f : {formula_n F}) : nvar n (~ f)%oT. Proof. by rewrite /nvar fsubset_formulan_fv. Qed. @@ -1180,16 +1191,17 @@ Proof. by rewrite /nvar (fsubset_trans (@fsubD1set _ _ _)) // fsubset_formulan_fv. Qed. -Canonical Structure formulan_exists (i : nat) (f : {formula_n F}) - := MkFormulan (exists_formulan i f). +Canonical Structure formulan_exists (i : nat) (f : {formula_n F}) := + MkFormulan (exists_formulan i f). -Lemma forall_formulan (i : nat) (f : {formula_n F}) : nvar n ('forall 'X_i, f)%oT. +Lemma forall_formulan (i : nat) (f : {formula_n F}) : + nvar n ('forall 'X_i, f)%oT. Proof. by rewrite /nvar (fsubset_trans (@fsubD1set _ _ _)) // fsubset_formulan_fv. Qed. -Canonical Structure formulan_forall (i : nat) (f : {formula_n F}) - := MkFormulan (forall_formulan i f). +Canonical Structure formulan_forall (i : nat) (f : {formula_n F}) := + MkFormulan (forall_formulan i f). Lemma fsubst_formulan (i : nat) (x : F) (f : {formula_n F}) : nvar n (fsubst f (i, (x%:T)%oT))%oT. @@ -1199,9 +1211,10 @@ by rewrite (fsubset_trans (@fsubD1set _ _ _)) // fsubset_formulan_fv. Qed. Canonical Structure formulan_fsubst (i : nat) (x : F) (f : {formula_n F}) := - MkFormulan (fsubst_formulan i x f). + MkFormulan (fsubst_formulan i x f). -Lemma And_formulan (I : finType) (x : {formula_n F}) (r : seq I) (P : pred I) (f : I -> {formula_n F}) : +Lemma And_formulan (I : finType) (x : {formula_n F}) (r : seq I) + (P : pred I) (f : I -> {formula_n F}) : nvar n (\big[And/x : formula F]_(i <- r | P i) f i). Proof. elim: r => [|i r IHr]. @@ -1210,10 +1223,12 @@ rewrite big_cons; case: (P i) => //. exact: and_formulan (f i) (MkFormulan IHr). Qed. -Canonical Structure formulan_And (I : finType) (x : {formula_n F}) (r : seq I) (P : pred I) (f : I -> {formula_n F}) := +Canonical Structure formulan_And (I : finType) (x : {formula_n F}) + (r : seq I) (P : pred I) (f : I -> {formula_n F}) := MkFormulan (And_formulan x r P f). -Lemma Or_formulan (I : finType) (x : {formula_n F}) (r : seq I) (P : pred I) (f : I -> {formula_n F}) : +Lemma Or_formulan (I : finType) (x : {formula_n F}) (r : seq I) + (P : pred I) (f : I -> {formula_n F}) : nvar n (\big[Or/x : formula F]_(i <- r | P i) f i). Proof. elim: r => [|i r IHr]. @@ -1222,7 +1237,8 @@ rewrite big_cons; case: (P i) => //. exact: or_formulan (f i) (MkFormulan IHr). Qed. -Canonical Structure formulan_Or (I : finType) (x : {formula_n F}) (r : seq I) (P : pred I) (f : I -> {formula_n F}) := +Canonical Structure formulan_Or (I : finType) (x : {formula_n F}) + (r : seq I) (P : pred I) (f : I -> {formula_n F}) := MkFormulan (Or_formulan x r P f). Lemma existsn_formulaSn (m : nat) (f : {formula_(m.+1) F}) : @@ -1252,7 +1268,7 @@ by rewrite leq0n (leq_trans lt_xm) // (leq_trans lt_mn) // leq_addr. Qed. Canonical Structure formulan_nexists m (f : {formula_m F}) := - MkFormulan (nexists_formulan f). + MkFormulan (nexists_formulan f). Lemma formulaSn_proof (f : {formula_n F}) : nvar n.+1 f. Proof. @@ -1273,7 +1289,7 @@ exact: ltn_addr. Qed. Canonical Structure formulan_add (m p : nat) (f : {formula_m F}) := - MkFormulan (formuladd p f). + MkFormulan (formuladd p f). End Closure. @@ -1295,11 +1311,11 @@ by rewrite formula_fv_fsubst ih seq_fset_cons fsetDDl fsetUC. Qed. Fact qf_form_fsubst (f : formula F) (i : nat) (t : GRing.term F) : - qf_form (fsubst f (i, t)) = (qf_form f). + qf_form (fsubst f (i, t)) = (qf_form f). Proof. by elim: f=> //=; move=> f1 -> f2 ->. Qed. Fact qf_form_fsubstn (f : formula F) (s : seq nat) (t : GRing.term F) : - qf_form (foldr (fun i h => fsubst h (i, t)) f s) = (qf_form f). + qf_form (foldr (fun i h => fsubst h (i, t)) f s) = (qf_form f). Proof. by elim: s => // x s ih; rewrite qf_form_fsubst ih. Qed. Lemma qf_elim_qf (f : formula F) : qf_form (qf_elim f). @@ -1318,18 +1334,16 @@ Lemma indep_elim (i : nat) (f : formula F) : Proof. move=> rform_f; rewrite /is_independent. split => h e; (split; first exact: holds_forall). -- move/(rwP (elim_rformP _ rform_f))/(rwP (qf_evalP _ (qf_form_elim rform_f))). - move/h; apply: monotonic_forall_if=> e2 h2. - apply/(rwP (elim_rformP _ rform_f)). - by apply/(rwP (qf_evalP _ (qf_form_elim rform_f))). -- move/(rwP (qf_evalP _ (qf_form_elim rform_f)))/(rwP (elim_rformP _ rform_f)). - move/h; apply: monotonic_forall_if=> e2 h2. - apply/(rwP (qf_evalP _ (qf_form_elim rform_f))). - by apply/(rwP (elim_rformP _ rform_f)). + move/(rwP (elim_rformP _ rform_f))/(rwP (qf_evalP _ (qf_form_elim rform_f))). + move/h; apply: monotonic_forall_if => e2 h2. + exact/(rwP (elim_rformP _ rform_f))/(rwP (qf_evalP _ (qf_form_elim rform_f))). +move/(rwP (qf_evalP _ (qf_form_elim rform_f)))/(rwP (elim_rformP _ rform_f)). +move/h; apply: monotonic_forall_if => e2 h2. +exact/(rwP (qf_evalP _ (qf_form_elim rform_f)))/(rwP (elim_rformP _ rform_f)). Qed. Lemma qf_elim_holdsP (f : formula F) (e : seq F) : - reflect (holds e f) (rcf_sat e (qf_elim f)). + reflect (holds e f) (rcf_sat e (qf_elim f)). Proof. apply: (equivP _ (to_rformP _ _)); apply: (equivP (rcf_satP _ _)). apply: (iff_trans (foldr_fsubst_indep _ _ _)) => [i | ]; last first. @@ -1344,14 +1358,14 @@ apply: (iff_trans _ (rwP (elim_rformP _ (to_rform_rformula _)))). move/(_ e2) : (independent not_fv) => h. move: (independent not_fv) => /(indep_to_rform _ _) /(_ e2) indep. apply: (iff_trans _ indep). -apply: monotonic_forall_iff=> e3. +apply: monotonic_forall_iff => e3. apply: (iff_trans (rwP (qf_evalP _ (qf_form_elim (to_rform_rformula _))))). apply: iff_sym. by apply: (iff_trans _ (rwP (elim_rformP _ (to_rform_rformula _)))). Qed. Fixpoint qf_subst_formula s (f : formula F) := let sterm := subst_term s in -match f with + match f with | (t1 == t2) => (sterm t1) == (sterm t2) | t1 <% t2 => (sterm t1) <% (sterm t2) | t1 <=% t2 => (sterm t1) <=% (sterm t2) @@ -1362,7 +1376,7 @@ match f with | ~ f => ~ (qf_subst_formula s f) | ('forall 'X_i, _) | ('exists 'X_i, _) => False | _ => f -end%oT. + end%oT. Definition subst_formula s (f : formula F) := qf_subst_formula s (qf_elim f). @@ -1386,44 +1400,19 @@ Fact fv_qf_subst_formula (k : unit) (s : seq nat) f : Proof. move: s; elim: f => //. - move=> t1 t2 s; rewrite fsubUset /=. - apply/andP; split. - + rewrite (fsubset_trans (fv_tsubst_map k _ _)) //. - apply/seq_fset_sub. - apply: sub_map_filter. - move=> i. - by rewrite in_fsetU => ->. - + rewrite (fsubset_trans (fv_tsubst_map k _ _)) //. - apply/seq_fset_sub. - apply: sub_map_filter. - move=> i. - by rewrite in_fsetU => ->; rewrite orbT. + apply/andP; split; rewrite (fsubset_trans (fv_tsubst_map k _ _)) //; + apply/seq_fset_sub/sub_map_filter => i; rewrite in_fsetU => -> //. + by rewrite orbT. - move=> t1 t2 s; rewrite fsubUset /=. - apply/andP; split. - + rewrite (fsubset_trans (fv_tsubst_map k _ _)) //. - apply/seq_fset_sub. - apply: sub_map_filter. - move=> i. - by rewrite in_fsetU => ->. - + rewrite (fsubset_trans (fv_tsubst_map k _ _)) //. - apply/seq_fset_sub. - apply: sub_map_filter. - move=> i. - by rewrite in_fsetU => ->; rewrite orbT. + apply/andP; split; rewrite (fsubset_trans (fv_tsubst_map k _ _)) //; + apply/seq_fset_sub/sub_map_filter => i; rewrite in_fsetU => -> //. + by rewrite orbT. - move=> t1 t2 s; rewrite fsubUset /=. - apply/andP; split. - + rewrite (fsubset_trans (fv_tsubst_map k _ _)) //. - apply/seq_fset_sub. - apply: sub_map_filter. - move=> i. - by rewrite in_fsetU => ->. - + rewrite (fsubset_trans (fv_tsubst_map k _ _)) //. - apply/seq_fset_sub. - apply: sub_map_filter. - move=> i. - by rewrite in_fsetU => ->; rewrite orbT. + apply/andP; split; rewrite (fsubset_trans (fv_tsubst_map k _ _)) //; + apply/seq_fset_sub/sub_map_filter => i; rewrite in_fsetU => -> //. + by rewrite orbT. - by move=> t s; apply: fv_tsubst_map. -- move=> f1 h1 f2 h2 s /=. - rewrite fsubUset. +- move=> f1 h1 f2 h2 s /=; rewrite fsubUset. apply/andP; split. + rewrite (fsubset_trans (h1 _)) //. apply/seq_fset_sub. @@ -1465,13 +1454,10 @@ Qed. Fact fv_subst_formula_map (k : unit) (s : seq nat) f : formula_fv (subst_formula s f) `<=` - seq_fset k [seq nth O s i | i <- (iota O (size s)) & (i \in formula_fv f)]. + seq_fset k [seq nth O s i | i <- (iota O (size s)) & (i \in formula_fv f)]. Proof. -rewrite /subst_formula. -rewrite (fsubset_trans (fv_qf_subst_formula k _ _)) //. -apply/seq_fset_sub. -apply: sub_map_filter. -move=> i. +rewrite /subst_formula (fsubset_trans (fv_qf_subst_formula k _ _)) //. +apply/seq_fset_sub/sub_map_filter => i. by move/fsubsetP/(_ i): (qf_elim_fv f). Qed. @@ -1480,8 +1466,7 @@ Proof. by apply/eqP; rewrite -fsubset0 -(seq_fset_nil _ tt) fv_subst_formula. Qed. -Definition cut(n : nat) (f : formula F) := - subst_formula (iota 0 n) f. +Definition cut (n : nat) (f : formula F) := subst_formula (iota 0 n) f. Fact nvar_cut n f : nvar n (cut n f). Proof. @@ -1492,7 +1477,8 @@ Qed. Canonical Structure cut_formulan n f := MkFormulan (nvar_cut n f). -Lemma rterm_tsubst (R : unitRingType) (t : GRing.term R) (s : nat * GRing.term R) : +Lemma rterm_tsubst (R : unitRingType) (t : GRing.term R) + (s : nat * GRing.term R) : GRing.rterm t -> GRing.rterm (snd s) -> GRing.rterm (GRing.tsubst t s). Proof. move=> + sr; elim: t => //=. @@ -1501,7 +1487,8 @@ move=> + sr; elim: t => //=. - by move=> t IHt u IHu /andP[] /IHt {}IHt /IHu {}IHu; apply/andP; split. Qed. -Lemma rform_fsubst (R : realDomainType) (f : formula R) (s : nat * GRing.term R) : +Lemma rform_fsubst (R : realDomainType) (f : formula R) + (s : nat * GRing.term R) : rformula f -> GRing.rterm (snd s) -> rformula (fsubst f s). Proof. move=> + sr; elim: f => //=. @@ -1515,8 +1502,7 @@ move=> + sr; elim: f => //=. - by move=> n f IHf fr; case: (n == fst s); last apply/IHf. Qed. -Lemma rform_qf_elim (f : formula F) : - rformula (qf_elim f). +Lemma rform_qf_elim (f : formula F) : rformula (qf_elim f). Proof. rewrite /qf_elim. elim: (enum_fset _) => /= [|x r IHr]; last exact/rform_fsubst. @@ -1531,7 +1517,7 @@ Variable (F : rcfType). Definition subst_env (s : seq nat) (e : seq F) := [seq nth 0 e i | i <- s]. Lemma subst_env_cat s1 s2 e : - subst_env (s1 ++ s2) e = subst_env s1 e ++ subst_env s2 e. + subst_env (s1 ++ s2) e = subst_env s1 e ++ subst_env s2 e. Proof. by rewrite /subst_env map_cat. Qed. Lemma subst_env_iota k1 k2 e1 e2 e3 : size e1 = k1 -> size e2 = k2 -> @@ -1545,11 +1531,11 @@ by rewrite !nth_cat nth_iota // ltnNge h1 leq_addr addnC addnK h2 lt_ik2. Qed. Lemma subst_env_iota_catl k e1 e2 : size e1 = k -> - subst_env (iota 0 k) (e1 ++ e2) = e1. + subst_env (iota 0 k) (e1 ++ e2) = e1. Proof. by move=> ?; rewrite -[e1 ++ e2]cat0s (@subst_env_iota 0). Qed. Lemma subst_env_iota_catr k1 k2 e1 e2 : size e1 = k1 -> size e2 = k2 -> - subst_env (iota k1 k2) (e1 ++ e2) = e2. + subst_env (iota k1 k2) (e1 ++ e2) = e2. Proof. by move=> h1 h2; rewrite -[e1 ++ e2]cats0 -catA subst_env_iota. Qed. Lemma subst_env_nil s : subst_env s [::] = nseq (size s) 0. @@ -1559,7 +1545,7 @@ by rewrite (nth_map O) // nth_nil nth_nseq if_same. Qed. Lemma eval_subst (e : seq F) (s : seq nat) (t : GRing.term F) : - GRing.eval e (subst_term s t) = GRing.eval (subst_env s e) t. + GRing.eval e (subst_term s t) = GRing.eval (subst_env s e) t. Proof. elim: t. - move=> i //=. @@ -1577,7 +1563,7 @@ elim: t. Qed. Lemma holds_subst e s f : - holds e (subst_formula s f) <-> holds (subst_env s e) f. + holds e (subst_formula s f) <-> holds (subst_env s e) f. Proof. rewrite (rwP (@qf_elim_holdsP F f _)) -(rwP (@rcf_satP _ _ _)) /subst_formula. move: e s; elim: (qf_elim f) (qf_elim_qf f) => // {f}. @@ -1592,7 +1578,7 @@ move: e s; elim: (qf_elim f) (qf_elim_qf f) => // {f}. Qed. Lemma holds_eq_vec e v1 v2 : - holds e (eq_vec F v1 v2) <-> (subst_env v1 e) = (subst_env v2 e). + holds e (eq_vec F v1 v2) <-> (subst_env v1 e) = (subst_env v2 e). Proof. move: v2; elim: v1 => [v2|] /=. by case: v2 => /=; rewrite /eq_vec ?big_ord0. @@ -1620,7 +1606,7 @@ Fact subst_env_tuple_subproof (i : nat) (t : i.-tuple nat) (e : seq F) : Proof. by apply/eqP/size_subst_env. Qed. Canonical subst_env_tuple (i : nat) (t : i.-tuple nat) (e : seq F) := - Tuple (subst_env_tuple_subproof t e). + Tuple (subst_env_tuple_subproof t e). End SubstEnv. diff --git a/semialgebraic.v b/semialgebraic.v index 88b50ed..70d2709 100644 --- a/semialgebraic.v +++ b/semialgebraic.v @@ -29,10 +29,11 @@ Require Import ZArith Init. From HB Require Import structures. Require Import mathcomp.ssreflect.ssreflect. -From mathcomp Require Import ssrfun ssrbool eqtype ssrnat seq choice path fintype div. -From mathcomp Require Import tuple finfun generic_quotient bigop finset perm. -From mathcomp Require Import ssralg ssrint poly polydiv ssrnum mxpoly binomial interval finalg complex. -From mathcomp Require Import zmodp mxpoly mpoly mxtens qe_rcf ordered_qelim realalg. +From mathcomp Require Import ssrfun ssrbool eqtype ssrnat seq choice path. +From mathcomp Require Import fintype div tuple finfun generic_quotient bigop. +From mathcomp Require Import finset perm ssralg ssrint poly polydiv ssrnum. +From mathcomp Require Import mxpoly binomial interval finalg complex zmodp. +From mathcomp Require Import mxpoly mpoly mxtens qe_rcf ordered_qelim realalg. From mathcomp Require Import matrix finmap order finset classical_sets topology. From mathcomp Require Import normedtype polyrcf polyorder. @@ -70,14 +71,14 @@ Variable (n : nat) (F : Type). Definition ngraph (m : nat) (x : 'rV[F]_m) := [tuple x ord0 i | i < m]. Definition ngraph_tnth k (t : k.-tuple F) : - ngraph (\row_(i < k) (tnth t i)) = t. + ngraph (\row_(i < k) (tnth t i)) = t. Proof. apply/val_inj; rewrite /= -map_tnth_enum; apply/eq_map => i. by rewrite mxE. Qed. Definition ngraph_nth k (x : F) (t : k.-tuple F) : - ngraph (\row_(i < k) (nth x t i)) = t. + ngraph (\row_(i < k) (nth x t i)) = t. Proof. rewrite -{2}[t]ngraph_tnth; congr ngraph; apply/rowP => i. by rewrite !mxE -tnth_nth. @@ -94,18 +95,20 @@ Fact size_ngraph (m : nat) (t : 'rV[F]_m) : size (ngraph t) = m. Proof. by rewrite size_tuple. Qed. Fact cat_ffunE (x0 : F) (m : nat) (t : 'rV[F]_m) (p : nat) - (u : 'rV[F]_p) (i : 'I_(m + p)) : -(row_mx t u) ord0 i = if (i < m)%N then nth x0 (ngraph t) i else nth x0 (ngraph u) (i - m). + (u : 'rV[F]_p) (i : 'I_(m + p)) : + (row_mx t u) ord0 i + = if (i < m)%N then nth x0 (ngraph t) i else nth x0 (ngraph u) (i - m). Proof. by rewrite mxE; case: splitP => j ->; rewrite ?(addnC, addnK) nth_ngraph. Qed. Fact ngraph_cat (m : nat) (t : 'rV[F]_m) (p : nat) (u : 'rV[F]_p) : - ngraph (row_mx t u) = ngraph t ++ ngraph u :> seq F. + ngraph (row_mx t u) = ngraph t ++ ngraph u :> seq F. Proof. case: m t => [|m] t. by rewrite row_thin_mx ngraph_nil. -apply: (@eq_from_nth _ (t ord0 ord0)) => [|i]; first by rewrite size_cat ?size_ngraph. +apply: (@eq_from_nth _ (t ord0 ord0)) => [|i]. + by rewrite size_cat ?size_ngraph. rewrite size_ngraph=> lt_i_mp; rewrite nth_cat. have -> : i = nat_of_ord (Ordinal lt_i_mp) by []. by rewrite nth_ngraph (cat_ffunE (t ord0 ord0)) size_ngraph. @@ -141,7 +144,7 @@ Variable n : nat. (* We define a relation in formulae *) Definition equivf (f g : formula F) := -rcf_sat [::] (nquantify O n Forall ((f ==> g) /\ (g ==> f))). + rcf_sat [::] (nquantify O n Forall ((f ==> g) /\ (g ==> f))). Lemma equivf_refl : reflexive equivf. Proof. by move=> f; apply/rcf_satP; apply: nforall_is_true => e /=. Qed. @@ -154,7 +157,7 @@ move: i e; elim: n=> [i e | n' iHn' i e]. by rewrite !nquantify0; apply/rcf_satP/rcf_satP => [[fg gf] | [gf fg]]; split. rewrite !nquantSout /=. apply/rcf_satP/rcf_satP => /= [h x | h x]; - move/(_ x)/rcf_satP : h => h; apply/rcf_satP. + move/(_ x)/rcf_satP : h => h; apply/rcf_satP. + by rewrite -iHn'. + by rewrite iHn'. Qed. @@ -210,10 +213,10 @@ Variable F : rcfType. (* is also a realDomainType *) Variable n : nat. Definition interp := fun (f : {formula_n F}) => - [pred v : 'rV[F]_n | rcf_sat (ngraph v) f]. + [pred v : 'rV[F]_n | rcf_sat (ngraph v) f]. Definition pred_of_SAset (s : {SAset F^n}) : - pred ('rV[F]_n) := interp (repr s). + pred ('rV[F]_n) := interp (repr s). Canonical SAsetPredType := PredType pred_of_SAset. End Interpretation. @@ -229,7 +232,7 @@ Fact rcf_sat_tuple (t : n.-tuple F) (f : {formula_n F}) : Proof. by rewrite inE ngraph_tnth. Qed. Fact holds_tuple (t : n.-tuple F) (s : {SAset F^n}) : - reflect (holds t s) ((\row_(i < n) (tnth t i)) \in s). + reflect (holds t s) ((\row_(i < n) (tnth t i)) \in s). Proof. apply: (iffP idP) => [h | ]. by apply/rcf_satP; rewrite rcf_sat_tuple. @@ -237,53 +240,49 @@ by move/rcf_satP; rewrite rcf_sat_tuple. Qed. Lemma holds_equivf (t : n.-tuple F) (f g : {formula_n F}) : - sub_equivf f g -> holds t f -> holds t g. + sub_equivf f g -> holds t f -> holds t g. Proof. by move/rcf_satP/nforallP; move/(_ t) => [h _]. Qed. Lemma rcf_sat_equivf (t : n.-tuple F) (f g : {formula_n F}) : - sub_equivf f g -> rcf_sat t f = rcf_sat t g. + sub_equivf f g -> rcf_sat t f = rcf_sat t g. Proof. move=> h; apply/rcf_satP/rcf_satP; apply: holds_equivf => //. by rewrite /sub_equivf /= equivf_sym. Qed. Fact rcf_sat_repr_pi (t : n.-tuple F) (f : {formula_n F}) : - rcf_sat t (repr (\pi_{SAset F^n} f)) = rcf_sat t f. + rcf_sat t (repr (\pi_{SAset F^n} f)) = rcf_sat t f. Proof. by case: piP => ? /eqmodP /rcf_sat_equivf ->. Qed. Fact holds_repr_pi (t : n.-tuple F) (f : {formula_n F}) : - holds t (repr (\pi_{SAset F^n} f)) <-> holds t f. + holds t (repr (\pi_{SAset F^n} f)) <-> holds t f. Proof. by split; apply: holds_equivf; rewrite /sub_equivf -eqmodE reprK. Qed. Lemma equivf_exists (f g : {formula_n F}) (i : nat) : - equivf n f g -> (equivf n ('exists 'X_i, f) ('exists 'X_i, g))%oT. + equivf n f g -> (equivf n ('exists 'X_i, f) ('exists 'X_i, g))%oT. Proof. rewrite /equivf; move/rcf_satP/nforallP => h. apply/rcf_satP/nforallP => u /=. -have [lt_in|leq_ni] := ltnP i n; last first. -+ split=> [[x]|]. - - move/holds_fsubst. - rewrite fsubst_id; last - by rewrite (contra (@in_fv_formulan _ _ _ _)) // -leqNgt. - move=> holds_uf; exists x; apply/holds_fsubst. - rewrite fsubst_id; last - by rewrite (contra (@in_fv_formulan _ _ _ _)) // -leqNgt. - by move: holds_uf; move/(_ u) : h; rewrite cat0s /=; move => []. - - move=> [x]; rewrite set_nth_over ?size_tuple // rcons_cat /=. - move/holds_take; rewrite take_size_cat ?size_tuple // => holds_ug. - exists 0; rewrite set_nth_nrcons ?size_tuple // rcons_nrcons -nseq_cat. - by apply/holds_cat_nseq; move: holds_ug; move/(_ u) : h => []. -+ split. - - move=> [x hxf]; exists x; move: hxf. - move/(_ [tuple of (set_nth 0 u (Ordinal lt_in) x)]) : h. - by rewrite cat0s /=; move=> []. - - move=> [x hxf]; exists x; move: hxf. - move/(_ [tuple of (set_nth 0 u (Ordinal lt_in) x)]) : h. - by rewrite cat0s /=; move=> []. +have [lt_in|leq_ni] := ltnP i n. + by split=> -[x hxf]; exists x; move: hxf; + move/(_ [tuple of (set_nth 0 u (Ordinal lt_in) x)]) : h; + rewrite cat0s /= => -[]. +split=> -[x]. + rewrite set_nth_over ?size_tuple // rcons_cat /=. + move/holds_take; rewrite take_size_cat ?size_tuple // => holds_ug. + exists 0; rewrite set_nth_nrcons ?size_tuple // rcons_nrcons -nseq_cat. + by apply/holds_cat_nseq; move: holds_ug; move/(_ u) : h => []. +move/holds_fsubst. +rewrite fsubst_id; last + by rewrite (contra (@in_fv_formulan _ _ _ _)) // -leqNgt. +move=> holds_uf; exists x; apply/holds_fsubst. +rewrite fsubst_id; last + by rewrite (contra (@in_fv_formulan _ _ _ _)) // -leqNgt. +by move: holds_uf; move/(_ u) : h; rewrite cat0s /=; move => []. Qed. Lemma pi_form (f : {formula_n F}) (x : 'rV[F]_n) : - (x \in \pi_{SAset F^n} f) = rcf_sat (ngraph x) f. + (x \in \pi_{SAset F^n} f) = rcf_sat (ngraph x) f. Proof. by rewrite inE; apply/rcf_satP/rcf_satP => ?; apply/holds_repr_pi. Qed. Lemma SAsetP (s1 s2 : {SAset F^n}) : reflect (s1 =i s2) (s1 == s2). @@ -577,7 +576,9 @@ apply/eqP/SAsetP => x; rewrite inSAsetU; apply/orP/SAin_setP => /=. by case=> xfg; [left|right]; apply/SAin_setP. Qed. -HB.instance Definition _ := Monoid.isComLaw.Build {SAset F^n} (SAset0 F n) (@SAsetU F n) SAsetUA SAsetUC SAset0U. +HB.instance Definition _ := + Monoid.isComLaw.Build {SAset F^n} + (SAset0 F n) (@SAsetU F n) SAsetUA SAsetUC SAset0U. Lemma SAsetIC A B : A :&: B = B :&: A. Proof. by apply/eqP/SAsetP => x; rewrite !inSAsetI andbC. Qed. @@ -603,7 +604,9 @@ apply/eqP/SAsetP => x; rewrite inSAsetI; apply/andP/SAin_setP. by move=> /= [] xf yf; split; apply/SAin_setP. Qed. -HB.instance Definition _ := Monoid.isComLaw.Build {SAset F^n} (SAsetT F n) (@SAsetI F n) SAsetIA SAsetIC SAsetTI. +HB.instance Definition _ := + Monoid.isComLaw.Build {SAset F^n} + (SAsetT F n) (@SAsetI F n) SAsetIA SAsetIC SAsetTI. Lemma SAsetCK A : ~: ~: A = A. @@ -763,8 +766,8 @@ Qed. Lemma SAsetCbigcap (I : Type) (r : seq I) (P : pred I) (f : I -> {SAset F^n}) : (~: \bigcap_(i <- r | P i) f i) = \bigcup_(i <- r | P i) ~: f i. Proof. -apply/eqP/SAsetP => x; rewrite inSAsetC inSAset_bigcap inSAset_bigcup -has_predC. -elim: r => [//|] i r IHr /=. +apply/eqP/SAsetP => x; rewrite inSAsetC inSAset_bigcap inSAset_bigcup. +rewrite -has_predC; elim: r => [//|] i r IHr /=. by rewrite negb_imply IHr inSAsetC. Qed. @@ -809,7 +812,8 @@ by split=> // -[]. Qed. Lemma SAset_cast_ge m k (A : {SAset F^m}) : (m <= k)%N -> - SAset_cast k A = [set | A /\ \big[And/True]_(i <- iota m (k - m)) ('X_i == 0)]. + SAset_cast k A + = [set | A /\ \big[And/True]_(i <- iota m (k - m)) ('X_i == 0)]. Proof. rewrite -subn_eq0 => /eqP km; apply/eqP/SAsetP => x. apply/Bool.eq_iff_eq_true. @@ -826,9 +830,11 @@ by move: x (mn); rewrite mn => x nn; rewrite SAset_cast_id castmx_id. Qed. Lemma inSAset_castDn m k (A : {SAset F^(m+k)}) (x : 'rV[F]_m) : - reflect (exists y : 'rV[F]_(m+k), y \in A /\ x = lsubmx y) (x \in SAset_cast m A). + reflect (exists y : 'rV[F]_(m+k), y \in A /\ x = lsubmx y) + (x \in SAset_cast m A). Proof. -rewrite SAset_cast_le ?leq_addr// subDnCA// subnn addn0 -[X in nquantify X](size_ngraph x). +rewrite SAset_cast_le ?leq_addr// subDnCA// subnn addn0. +rewrite -[X in nquantify X](size_ngraph x). apply/(iffP (SAin_setP _ _)) => [/nexistsP [y] hxy|[y][yA]->]. exists (row_mx x (\row_i tnth y i)); rewrite row_mxKl; split=> //. by apply/rcf_satP; rewrite ngraph_cat ngraph_tnth. @@ -840,11 +846,15 @@ Lemma inSAset_castnD m k (A : {SAset F^m}) (x : 'rV[F]_(m+k)) : x \in SAset_cast (m+k) A = (lsubmx x \in A) && (rsubmx x == 0). Proof. rewrite SAset_cast_ge ?leq_addr//. -apply/SAin_setP/andP => /=; rewrite -holds_take take_ngraph holdsAnd /= => -[/rcf_satP hx]. +apply/SAin_setP/andP => /=; + rewrite -holds_take take_ngraph holdsAnd /= => -[/rcf_satP hx]. move=> h0; split=> //; apply/eqP/rowP => i. - move/(_ (@unsplit m k (inr i))): h0; rewrite nth_ngraph mem_iota subnKC ?leq_addr//= -addnS leq_add// => /(_ Logic.eq_refl Logic.eq_refl). + move/(_ (@unsplit m k (inr i))): h0. + rewrite nth_ngraph mem_iota subnKC ?leq_addr//= -addnS leq_add//. + move=> /(_ Logic.eq_refl Logic.eq_refl). by rewrite !mxE. -move=> /eqP /rowP x0; split=> // => i; rewrite mem_iota subnKC ?leq_addr// => /andP[mi im] _. +move=> /eqP /rowP x0; split=> // => i. +rewrite mem_iota subnKC ?leq_addr// => /andP[mi im] _. rewrite (nth_ngraph _ _ (Ordinal im)) -(splitK (Ordinal im)). move: mi; rewrite leqNgt -{1}[i%N]/(Ordinal im : nat). case: splitP => // j _ _. @@ -862,7 +872,10 @@ case: (ltnP m n) => [mn|nm _]; last first. move: (lsubmx x) (rsubmx x) (hsubmxK x) => l r <- {x}. move: (lsubmx l) (rsubmx l) (hsubmxK l) => ll lr <- {l}. rewrite SAset_cast_ge; last by rewrite subnKC// subnKC// (leq_trans nm mk). - apply/andP/SAin_setP => /=; rewrite holdsAnd -holds_take -(take_takel _ (@leq_addr (m-n) n)%N) !take_ngraph !row_mxKl (rwP (rcf_satP _ _)) subDnCA ?leq_addr// subDnCA// subnn addn0 addnC. + apply/andP/SAin_setP => /=; + rewrite holdsAnd -holds_take -(take_takel _ (@leq_addr (m-n) n)%N); + rewrite !take_ngraph !row_mxKl (rwP (rcf_satP _ _)); + rewrite subDnCA ?leq_addr// subDnCA// subnn addn0 addnC. move=> [] /andP[] llA /eqP -> /eqP ->; split=> //= i. rewrite mem_iota addnA => /andP[+ ilt] _. rewrite -[i%N]/(Ordinal ilt : nat) nth_ngraph mxE. @@ -874,8 +887,9 @@ case: (ltnP m n) => [mn|nm _]; last first. rewrite nth_ngraph !mxE unsplitK. by rewrite mem_iota addnA ltn_ord/= -addnA leq_addr; apply. apply/andP; split=> //. - apply/eqP/rowP => i. - move/(_ (unsplit (inl (unsplit (inr i))) : 'I_(n + (m - n) + (k - m))%N)): h0. + apply/eqP/rowP => i; move: h0. + move=> /(_ (unsplit (inl (unsplit (inr i))) : + 'I_(n + (m - n) + (k - m))%N)). rewrite nth_ngraph !mxE unsplitK mxE unsplitK. by rewrite mem_iota addnA ltn_ord/= leq_addr; apply. case/orP: (leq_total n k) => [nk|kn]. @@ -885,45 +899,59 @@ case: (ltnP m n) => [mn|nm _]; last first. move: (lsubmx x) (rsubmx x) (hsubmxK x) => l r <- {x}. apply/inSAset_castDn/andP => [[y]|[lA] /eqP ->]; rewrite SAset_cast_ge -?addnA ?leq_addr//. - move: (lsubmx y) (rsubmx y) (hsubmxK y) => yl yr <- {y} [] /[swap] <- {yl}. - move=> /SAin_setP/= [] /holds_take. - rewrite -(take_takel _ (@leq_addr (k - n)%N n)) !take_ngraph !row_mxKl => /rcf_satP lA /holdsAnd. + move: (lsubmx y) (rsubmx y) (hsubmxK y). + move=> yl yr <- {y} [] /[swap] <- {yl} /SAin_setP/= [] /holds_take. + rewrite -(take_takel _ (@leq_addr (k - n)%N n)) !take_ngraph !row_mxKl. + move=> /rcf_satP lA /holdsAnd. rewrite subDnCA ?leq_addr// subDnCA// subnn addn0 addnC /= => h0. - split=> //; apply/eqP/rowP => i. - move/(_ (unsplit (inl (unsplit (inr i))) : 'I_(n + (k - n) + (m - k))%N)): h0. - by rewrite nth_ngraph !mxE unsplitK mxE unsplitK mem_iota addnA ltn_ord/= leq_addr; apply. + split=> //; apply/eqP/rowP => i; move: h0. + move=> /(_ (unsplit (inl (unsplit (inr i))) : + 'I_(n + (k - n) + (m - k))%N)). + rewrite nth_ngraph !mxE unsplitK mxE unsplitK mem_iota addnA ltn_ord/=. + by rewrite leq_addr; apply. exists (row_mx (row_mx l 0) 0); rewrite row_mxKl; split=> //. apply/SAin_setP => /=; split. apply/holds_take. rewrite -(take_takel _ (@leq_addr (k - n)%N n)) !take_ngraph !row_mxKl. exact/rcf_satP. - apply/holdsAnd => i; rewrite mem_iota subDnCA ?leq_addr// subDnCA// subnn addn0 [X in (n + X)%N]addnC /= addnA => /andP[+ ilt] _. + apply/holdsAnd => i; rewrite mem_iota subDnCA ?leq_addr// subDnCA// subnn. + rewrite addn0 [X in (n + X)%N]addnC /= addnA => /andP[+ ilt] _. rewrite -[i%N]/(Ordinal ilt : nat) nth_ngraph mxE. case: (splitP (Ordinal ilt)) => j ->; rewrite mxE//. by case: (splitP j) => j' ->; rewrite leqNgt ?ltn_ord// mxE. - move: A; rewrite -(subnKC nm) -(subnKC kn) [X in (m - X)%N]subnKC// -addnA => A. + move: A; rewrite -(subnKC nm) -(subnKC kn) [X in (m - X)%N]subnKC// -addnA. + move=> A. apply/eqP/SAsetP => x; apply/inSAset_castDn/inSAset_castDn => -[y]. - rewrite [_ _ A]SAset_cast_ge ?addnA ?leq_addr// => -[] /SAin_setP /= [] /holds_take + _. + rewrite [_ _ A]SAset_cast_ge ?addnA ?leq_addr//. + move=> -[] /SAin_setP /= [] /holds_take + _. rewrite takeD take_ngraph drop_ngraph take_ngraph -ngraph_cat => yA -> {x}. exists (row_mx (lsubmx y) (lsubmx (rsubmx y))); split; first exact/rcf_satP. by rewrite row_mxKl. - move=> [] /rcf_satP yA -> {x}; exists (row_mx (lsubmx y) (row_mx (rsubmx y) 0)); split; last by rewrite row_mxKl. + move=> [] /rcf_satP yA -> {x}. + exists (row_mx (lsubmx y) (row_mx (rsubmx y) 0)). + split; last by rewrite row_mxKl. rewrite [_ _ A]SAset_cast_ge ?addnA ?leq_addr//; apply/SAin_setP => /=; split. apply/holds_take. - by rewrite takeD take_ngraph drop_ngraph take_ngraph -ngraph_cat row_mxKr !row_mxKl hsubmxK. - apply/holdsAnd => i; rewrite {1}addnA subnKC// subnKC// mem_iota -{1 2}(subnKC kn) -addnA => /andP[] + ilt _ /=. - rewrite -[i%N]/(Ordinal ilt : nat) nth_ngraph. - rewrite mxE; case: splitP => j ->. - by rewrite leqNgt (leq_trans (ltn_ord j) (leq_addr _ _)). - rewrite leq_add2l mxE; case: splitP => j' ->; last by rewrite mxE. - by rewrite leqNgt ltn_ord. + rewrite takeD take_ngraph drop_ngraph take_ngraph -ngraph_cat row_mxKr. + by rewrite !row_mxKl hsubmxK. + apply/holdsAnd => i; rewrite {1}addnA subnKC// subnKC// mem_iota. + rewrite -{1 2}(subnKC kn) -addnA => /andP[] + ilt _ /=. + rewrite -[i%N]/(Ordinal ilt : nat) nth_ngraph. + rewrite mxE; case: splitP => j ->. + by rewrite leqNgt (leq_trans (ltn_ord j) (leq_addr _ _)). + rewrite leq_add2l mxE; case: splitP => j' ->; last by rewrite mxE. + by rewrite leqNgt ltn_ord. rewrite geq_min leqNgt mn/= => km. -rewrite SAset_cast_le// SAset_cast_le ?(ltnW mn)// SAset_cast_le ?(ltnW (leq_ltn_trans km mn))//. -apply/eqP/SAsetP => x; rewrite -[X in nquantify X](size_ngraph x); apply/SAin_setP/SAin_setP => /nexistsP [y] => /rcf_satP. +rewrite SAset_cast_le// SAset_cast_le ?(ltnW mn)//. +rewrite SAset_cast_le ?(ltnW (leq_ltn_trans km mn))//. +apply/eqP/SAsetP => x; rewrite -[X in nquantify X](size_ngraph x). +apply/SAin_setP/SAin_setP => /nexistsP [y] => /rcf_satP. rewrite -[in X in rcf_sat _ X](subnKC km). rewrite -[y]ngraph_tnth -ngraph_cat => /SAin_setP. - have mE: (k + (m - k))%N = size (ngraph x ++ y) by rewrite size_cat size_ngraph size_tuple subnKC. - rewrite [X in nquantify X]mE -{2}[y]ngraph_tnth -ngraph_cat => /nexistsP [] {mE}. + have mE: (k + (m - k))%N = size (ngraph x ++ y). + by rewrite size_cat size_ngraph size_tuple subnKC. + rewrite [X in nquantify X]mE -{2}[y]ngraph_tnth -ngraph_cat. + move=> /nexistsP [] {mE}. rewrite ngraph_cat (subnKC km) ngraph_tnth => z hA. apply/nexistsP. have ->: (n - k)%N = (n - m + m - k)%N by rewrite subnK// (ltnW mn). @@ -936,12 +964,14 @@ have /eqP ts: size (take (m - k)%N y) = (m - k)%N. exists (Tuple ts); rewrite -[in X in holds _ X](subnKC km). rewrite -[Tuple ts]ngraph_tnth -ngraph_cat. apply/rcf_satP/SAin_setP. -have mE: (k + (m - k))%N = size (ngraph x ++ Tuple ts) by rewrite size_cat size_ngraph size_tuple subnKC. +have mE: (k + (m - k))%N = size (ngraph x ++ Tuple ts). + by rewrite size_cat size_ngraph size_tuple subnKC. rewrite [X in nquantify X]mE -{2}[Tuple ts]ngraph_tnth -ngraph_cat. apply/nexistsP. rewrite ngraph_cat subnKC//. have /eqP ds: size (drop (m - k)%N y) = (n - m)%N. - by rewrite size_drop size_tuple subnBA// addnC subnKC// (ltnW (leq_ltn_trans km mn)). + rewrite size_drop size_tuple subnBA// addnC subnKC//. + exact/(ltnW (leq_ltn_trans km mn)). by exists (Tuple ds); rewrite -catA ngraph_tnth/= cat_take_drop. Qed. @@ -968,10 +998,13 @@ Lemma SAset_disjointC (s1 s2 : {SAset F^n}) : Proof. by rewrite /SAset_disjoint SAsetIC. Qed. Definition SAset_trivI (I : {fset {SAset F^n}}) := - [forall s1 : I, [forall s2 : I, (val s1 != val s2) ==> SAset_disjoint (val s1) (val s2)]]. + [forall s1 : I, + [forall s2 : I, (val s1 != val s2) ==> SAset_disjoint (val s1) (val s2)]]. Definition SAset_partition (I : {fset {SAset F^n}}) := - (SAset0 F n \notin I) && SAset_trivI I && (\big[@SAsetU F n/SAset0 F n]_(s : I) val s == SAsetT F n). + (SAset0 F n \notin I) + && SAset_trivI I + && (\big[@SAsetU F n/SAset0 F n]_(s : I) val s == SAsetT F n). End SAsetTheory. @@ -985,8 +1018,9 @@ Fact SAset_disp : unit. Proof. exact tt. Qed. HB.instance Definition SAset_latticeType := Order.isMeetJoinDistrLattice.Build SAset_disp {SAset _} - (@SAsubsetEI F n) (@properEneq' F n) (@SAsetIC F n) (@SAsetUC F n) - (@SAsetIA F n) (@SAsetUA F n) (@SAsetKU' F n) (@SAsetKI' F n) (@SAsetIUl F n) (@SAsetI_idem F n). + (@SAsubsetEI F n) (@properEneq' F n) (@SAsetIC F n) (@SAsetUC F n) + (@SAsetIA F n) (@SAsetUA F n) (@SAsetKU' F n) (@SAsetKI' F n) + (@SAsetIUl F n) (@SAsetI_idem F n). HB.instance Definition _ := Order.hasBottom.Build SAset_disp {SAset F^n} (@sub0set F n). @@ -994,7 +1028,9 @@ HB.instance Definition _ := HB.instance Definition SAset_tblatticeType := Order.hasTop.Build SAset_disp {SAset F^n} (@subsetT F n). -HB.instance Definition _ := Order.hasRelativeComplement.Build SAset_disp {SAset F^n} (@SAsetID0 F n) (@SAsetUID F n). +HB.instance Definition _ := + Order.hasRelativeComplement.Build SAset_disp {SAset F^n} + (@SAsetID0 F n) (@SAsetUID F n). End POrder. @@ -1005,43 +1041,43 @@ Variable F : rcfType. Variables (n m : nat). Definition ftotal (f : {formula_(n + m) F}) := - nquantify O n Forall (nquantify n m Exists f). + nquantify O n Forall (nquantify n m Exists f). Definition ex_y (f : {formula_(n + m) F}) (x : 'rV[F]_n) := - rcf_sat (ngraph x) (nquantify n m Exists f). + rcf_sat (ngraph x) (nquantify n m Exists f). Definition SAtot := - [pred s : {SAset F ^ _} | rcf_sat [::] (ftotal s)]. + [pred s : {SAset F ^ _} | rcf_sat [::] (ftotal s)]. Definition ext (p : nat) (f : {formula_p F}) : {formula_(p+m) F} := MkFormulan (formuladd m f). Lemma ftotalP (f : {formula_(n + m) F}) : - reflect + reflect (forall (t : n.-tuple F), exists (u : m.-tuple F), rcf_sat (t ++ u) f) (rcf_sat [::] (ftotal f)). Proof. apply: (iffP idP) => [h x | h]. -+ move/rcf_satP/nforallP/(_ x) : h. + move/rcf_satP/nforallP/(_ x) : h. case: x => s /= /eqP -{1}<-. by move/nexistsP => [t h]; exists t; apply/rcf_satP. -+ apply/rcf_satP/nforallP => x /=. - move/(_ x) : h => [t]. - case: x => s /= /eqP -{2}<-. - by move/rcf_satP => h; apply/nexistsP; exists t. +apply/rcf_satP/nforallP => x /=. +move/(_ x) : h => [t]. +case: x => s /= /eqP -{2}<-. +by move/rcf_satP => h; apply/nexistsP; exists t. Qed. Definition functional (f : {formula_(n+m) F}) := (nquantify O (n + m + m) Forall ( - ((subst_formula (iota 0 n ++ iota n m) f) - /\ (subst_formula (iota 0 n ++ iota (n + m) m) f)) - ==> (eq_vec F (iota n m) (iota (n + m) m)))). + ((subst_formula (iota 0 n ++ iota n m) f) + /\ (subst_formula (iota 0 n ++ iota (n + m) m) f)) + ==> (eq_vec F (iota n m) (iota (n + m) m)))). Definition SAfunc := - [pred s : {SAset F ^ _} | rcf_sat [::] (functional s)]. + [pred s : {SAset F ^ _} | rcf_sat [::] (functional s)]. Lemma functionalP (f : {formula_(n + m) F}) : - reflect + reflect (forall (t : n.-tuple F) (u1 u2 : m.-tuple F), rcf_sat (t ++ u1) f -> rcf_sat (t ++ u2) f -> u1 = u2) (rcf_sat [::] (functional f)). @@ -1078,9 +1114,7 @@ apply: (iffP idP). Qed. Lemma inSAtot (s : {SAset F ^ (n + m)}) : - reflect - (forall (x : 'rV[F]_n), exists (y : 'rV[F]_m), (row_mx x y) \in s) - (s \in SAtot). + reflect (forall x, exists y, (row_mx x y) \in s) (s \in SAtot). Proof. rewrite inE; apply: (iffP (ftotalP _)) => s_sat x. have [y sat_s_xy] := s_sat (ngraph x). @@ -1092,9 +1126,8 @@ by move: xy_in_s; rewrite inE ngraph_cat ngraph_nth. Qed. Lemma inSAfunc (s : {SAset F ^ (n + m)}) : - reflect - (forall (x : 'rV[F]_n), forall (y1 y2 : 'rV[F]_m), - (row_mx x y1) \in s -> (row_mx x y2) \in s -> y1 = y2) + reflect + (forall x y1 y2, (row_mx x y1) \in s -> (row_mx x y2) \in s -> y1 = y2) (s \in SAfunc). Proof. rewrite inE; apply: (iffP (functionalP _)) => fun_s x y1 y2. @@ -1110,7 +1143,7 @@ rewrite inE !ngraph_cat !ngraph_nth. Qed. Lemma ex_yE (f : {formula_(n + m) F}) (t : 'rV[F]_n) : - reflect (exists (u : 'rV[F]_m), rcf_sat (ngraph (row_mx t u)) f) (ex_y f t). + reflect (exists (u : 'rV[F]_m), rcf_sat (ngraph (row_mx t u)) f) (ex_y f t). Proof. apply: (iffP idP); rewrite /ex_y. rewrite -{1}[X in nquantify X _ _](size_ngraph t). @@ -1122,15 +1155,12 @@ apply/rcf_satP; rewrite -{1}[X in nquantify X _ _](size_ngraph t). by apply/nexistsP; exists (ngraph u); apply/rcf_satP. Qed. -Definition get_y (f : {formula_(n + m) F}) (x : 'rV[F]_n) : ('rV[F]_m):= +Definition form_to_fun (f : {formula_(n + m) F}) (x : 'rV[F]_n) : 'rV[F]_m := match boolP (ex_y f x) with | AltTrue p => proj1_sig (sigW (ex_yE f x p)) | AltFalse _ => (\row_(i < m) 0) end. -Definition form_to_fun (f : {formula_(n + m) F}) : 'rV[F]_n -> 'rV[F]_m := - fun (x : 'rV[F]_n) => get_y f x. - Record SAfun := MkSAfun { SAgraph :> {SAset F ^ (n + m)}; @@ -1149,16 +1179,15 @@ HB.instance Definition SAfun_of_subType := SubType.copy {SAfun} SAfun. HB.instance Definition SAfun_of_eqType := Equality.copy {SAfun} SAfun. HB.instance Definition SAfun_of_choiceType := Choice.copy {SAfun} SAfun. -Lemma SAfun_func (f : {SAfun}) (x : 'rV[F]_n) (y1 y2 : 'rV[F]_m) : - row_mx x y1 \in SAgraph f -> row_mx x y2 \in SAgraph f -> y1 = y2. +Lemma SAfun_func (f : {SAfun}) x y1 y2 : + row_mx x y1 \in SAgraph f -> row_mx x y2 \in SAgraph f -> y1 = y2. Proof. by apply: inSAfunc; case: f; move => /= [f h /andP [h1 h2]]. Qed. Lemma SAfun_tot (f : {SAfun}) (x : 'rV[F]_n) : - exists (y : 'rV[F]_m), row_mx x y \in SAgraph f. + exists y, row_mx x y \in SAgraph f. Proof. by apply: inSAtot; case: f; move => /= [f h /andP [h1 h2]]. Qed. -Definition SAfun_to_fun (f : SAfun) : 'rV[F]_n -> 'rV[F]_m := - fun x => proj1_sig (sigW (SAfun_tot f x)). +Definition SAfun_to_fun (f : SAfun) x := proj1_sig (sigW (SAfun_tot f x)). Coercion SAfun_to_fun : SAfun >-> Funclass. @@ -1177,14 +1206,13 @@ Section SAfunTheory. Variable (F : rcfType) (n m : nat). -Lemma inSAgraph (f : {SAfun F^n -> F^m}) (x : 'rV[F]_n) : - row_mx x (f x) \in SAgraph f. +Lemma inSAgraph (f : {SAfun F^n -> F^m}) x : + row_mx x (f x) \in SAgraph f. Proof. by rewrite /SAfun_to_fun; case: ((sigW (SAfun_tot f x))) => y h. Qed. -Lemma inSAfun (f : {SAfun F^n -> F^m}) - (x : 'rV[F]_n) (y : 'rV[F]_m): +Lemma inSAfun (f : {SAfun F^n -> F^m}) x y : (f x == y) = (row_mx x y \in SAgraph f). Proof. apply/eqP/idP => [<- | h]; first by rewrite inSAgraph. @@ -1203,8 +1231,7 @@ Definition SAimset (f : {SAset F ^ (n + m)}) (s : {SAset F^n}) : {SAset F^m} := [set | nquantify m n Exists ((subst_formula ((iota m n) ++ (iota O m)) f) /\ (subst_formula (iota m n) s)) ]. -Lemma inSAimset (x : 'rV[F]_n) - (s : {SAset F^n}) (f : {SAfun F^n -> F^m}) : +Lemma inSAimset (f : {SAfun F^n -> F^m}) s x : x \in s -> f x \in SAimset f s. Proof. rewrite pi_form /= => h. @@ -1221,10 +1248,8 @@ split; apply/holds_subst; move: h; rewrite inE => /rcf_satP hs; last first. by move: (inSAgraph f x); rewrite inE => /rcf_satP; rewrite ngraph_cat. Qed. -Lemma SAimsetP (f : {SAfun F^n -> F^m}) - (s : {SAset F^n}) (y : 'rV[F]_m) : - reflect (exists2 x : 'rV[F]_n, x \in s & y = f x) - (y \in (SAimset f s)). +Lemma SAimsetP (f : {SAfun F^n -> F^m}) s y : + reflect (exists2 x, x \in s & y = f x) (y \in (SAimset f s)). Proof. apply: (iffP idP) => [/SAin_setP|[x h]->]; last exact: inSAimset. rewrite -[X in nquantify X _ _ _](size_ngraph y) => /nexistsP [t] /=. @@ -1243,11 +1268,10 @@ exists (\row_(i < n) t`_i). by rewrite inE ngraph_cat ngraph_nth; apply/rcf_satP. Qed. -Definition SApreimset (f : {SAfun F^n -> F^m}) (s : {SAset F^m}) : {SAset F^n} := - [set | nquantify n m Exists (f /\ (subst_formula (iota n m) s)) ]. +Definition SApreimset (f : {SAfun F^n -> F^m}) (s : {SAset F^m}) : {SAset F^n} + := [set | nquantify n m Exists (f /\ (subst_formula (iota n m) s)) ]. -Lemma inSApreimset (x : 'rV[F]_n) - (s : {SAset F^m}) (f : {SAfun F^n -> F^m}) : +Lemma inSApreimset (f : {SAfun F^n -> F^m}) s x : x \in SApreimset f s = (f x \in s). Proof. apply/SAin_setP/rcf_satP => [|fxs]; @@ -1273,18 +1297,20 @@ Definition SAhypograph (f : {SAfun F^n -> F^1}) : {SAset F^(n + 1)} := End SAfunTheory. -Lemma inSAepigraph (F : rcfType) (n : nat) (f : {SAfun F^n -> F^1}) (x : 'rV[F]_(n + 1)) : +Lemma inSAepigraph (F : rcfType) (n : nat) (f : {SAfun F^n -> F^1}) x : (x \in SAepigraph f) = (f (lsubmx x) ord0 ord0 < rsubmx x ord0 ord0). Proof. move: (lsubmx x) (rsubmx x) (hsubmxK x) => l r <- {x}. apply/SAin_setP/idP; rewrite -[X in nquantify X _ _](size_ngraph (row_mx l r)). move=> /nexistsP [y] /= [] /holds_subst. rewrite nth_cat size_map size_enum_ord {11 20}addn1 ltnn subnn. - rewrite nth_cat size_map size_enum_ord {11}addn1 leqnn (nth_map (unsplit (inr ord0))) ?size_enum_ord ?addn1//. + rewrite nth_cat size_map size_enum_ord {11}addn1 leqnn. + rewrite (nth_map (unsplit (inr ord0))) ?size_enum_ord ?addn1//. have {26}->: n = @unsplit n 1 (inr ord0) by rewrite /= addn0. rewrite nth_ord_enum mxE unsplitK ngraph_cat -catA subst_env_cat. rewrite subst_env_iota_catl ?size_ngraph//= !nth_cat size_map size_enum_ord. - rewrite ltnNge leqnSn/= subSnn size_ngraph/= ltnn !subnn/= (nth_map ord0) ?size_enum_ord//. + rewrite ltnNge leqnSn/= subSnn size_ngraph/= ltnn !subnn/=. + rewrite (nth_map ord0) ?size_enum_ord//. rewrite -[X in nth ord0 _ X]/(@ord0 1 : nat) (@nth_ord_enum 1 ord0 ord0). move=> /holds_take; rewrite take_cat size_ngraph ltnNge {1}addn1 leqnSn/=. rewrite subDnCA// subnn/= => hf; congr (_ < _). @@ -1294,8 +1320,8 @@ apply/SAin_setP/idP; rewrite -[X in nquantify X _ _](size_ngraph (row_mx l r)). by rewrite size_cat size_map size_enum_ord size_ngraph. rewrite size_cat size_map size_enum_ord /= => ilt. have i0: 'I_(n+1) by rewrite addn1; exact: ord0. - rewrite (nth_map (Ordinal ilt)) ?size_enum_ord// -[i%N]/(Ordinal ilt : nat) nth_ord_enum. - rewrite mxE -{1}(splitK (Ordinal ilt)); case: (split _) => j. + rewrite (nth_map (Ordinal ilt)) ?size_enum_ord// -[i%N]/(Ordinal ilt : nat). + rewrite nth_ord_enum mxE -{1}(splitK (Ordinal ilt)); case: (split _) => j. rewrite nth_cat size_map size_enum_ord ltn_unsplit/=. by rewrite (nth_map j) ?size_enum_ord// nth_ord_enum /=. rewrite nth_cat/= size_ngraph ltnNge leq_addr/= subDnCA// subnn addn0. @@ -1319,7 +1345,7 @@ have ->: (ngraph l) ++ [:: f l ord0 ord0] = ngraph (row_mx l (f l)). by move: (inSAfun f l (f l)); rewrite eqxx => /esym/rcf_satP. Qed. -Lemma inSAhypograph (F : rcfType) (n : nat) (f : {SAfun F^n -> F^1}) (x : 'rV[F]_(n + 1)) : +Lemma inSAhypograph (F : rcfType) (n : nat) (f : {SAfun F^n -> F^1}) x : (x \in SAhypograph f) = (rsubmx x ord0 ord0 < f (lsubmx x) ord0 ord0). Proof. move: (lsubmx x) (rsubmx x) (hsubmxK x) => l r <- {x}. @@ -1330,7 +1356,8 @@ apply/SAin_setP/idP; rewrite -[X in nquantify X _ _](size_ngraph (row_mx l r)). have {26}->: n = @unsplit n 1 (inr ord0) by rewrite /= addn0. rewrite nth_ord_enum mxE unsplitK ngraph_cat -catA subst_env_cat. rewrite subst_env_iota_catl ?size_ngraph//= !nth_cat size_map size_enum_ord. - rewrite ltnNge leqnSn/= subSnn size_ngraph/= ltnn !subnn/= (nth_map ord0) ?size_enum_ord//. + rewrite ltnNge leqnSn/= subSnn size_ngraph/= ltnn !subnn/=. + rewrite (nth_map ord0) ?size_enum_ord//. rewrite -[X in nth ord0 _ X]/(@ord0 1 : nat) (@nth_ord_enum 1 ord0 ord0). move=> /holds_take; rewrite take_cat size_ngraph ltnNge {1}addn1 leqnSn/=. rewrite subDnCA// subnn/= => hf; congr (_ < _). @@ -1340,8 +1367,8 @@ apply/SAin_setP/idP; rewrite -[X in nquantify X _ _](size_ngraph (row_mx l r)). by rewrite size_cat size_map size_enum_ord size_ngraph. rewrite size_cat size_map size_enum_ord /= => ilt. have i0: 'I_(n+1) by rewrite addn1; exact: ord0. - rewrite (nth_map (Ordinal ilt)) ?size_enum_ord// -[i%N]/(Ordinal ilt : nat) nth_ord_enum. - rewrite mxE -{1}(splitK (Ordinal ilt)); case: (split _) => j. + rewrite (nth_map (Ordinal ilt)) ?size_enum_ord// -[i%N]/(Ordinal ilt : nat). + rewrite nth_ord_enum mxE -{1}(splitK (Ordinal ilt)); case: (split _) => j. rewrite nth_cat size_map size_enum_ord ltn_unsplit/=. by rewrite (nth_map j) ?size_enum_ord// nth_ord_enum /=. rewrite nth_cat/= size_ngraph ltnNge leq_addr/= subDnCA// subnn addn0. @@ -1370,24 +1397,20 @@ Section SAfunOps. Variable (F : rcfType). Definition abs (i j : nat) : formula F := - ((('X_j == 'X_i) \/ ('X_j == - 'X_i)) /\ (0 <=% 'X_j))%oT. + ((('X_j == 'X_i) \/ ('X_j == - 'X_i)) /\ (0 <=% 'X_j))%oT. Lemma absP (e : seq F) (i j : nat) : holds e (abs i j) <-> e`_j = `|e`_i|. Proof. -rewrite /abs /=; split. -+ move=> [[->|-> h]]; first by move=> h; rewrite ger0_norm. +rewrite /abs /=; split=> [|->]. + move=> [[->|-> h]]; first by move=> h; rewrite ger0_norm. by rewrite ler0_norm // -oppr_gte0. -+ move->. - rewrite normr_ge0; split => //. - have [le_e0|lt_0e] := ler0P e`_i; first by right. - by left. +rewrite normr_ge0; split => //. +have [le_e0|lt_0e] := ler0P e`_i; first by right. +by left. Qed. Lemma absP2 (e : seq F) (i j : nat) : rcf_sat e (abs i j) = (e`_j == `|e`_i|). -Proof. -apply/rcf_satP/eqP; first by move/absP. -by move=> h; apply/absP. -Qed. +Proof. by apply/rcf_satP/eqP => [/absP //|h]; apply/absP. Qed. Fact nvar_abs (i j : nat) : @nvar F (maxn i j).+1 (abs i j). Proof. @@ -1404,10 +1427,8 @@ Lemma functional_absset : absset \in SAfunc. Proof. apply/rcf_satP/nforallP => t. move=> [/holds_subst/holds_repr_pi/absP h1 /holds_subst/holds_repr_pi/absP h2]. -apply/holds_eq_vec; move: h1 h2; case: t => s sz //= {sz}. -case: s => // a s; case: s => // b s -> /= {b}; case: s => //. - by move <-. -by move=> b // _ ->. +apply/holds_eq_vec; move: h1 h2; case: t => s _ /=. +by case: s => // a; case=> // b + -> /= {b}; case=> [<-|b _ ->]. Qed. Lemma total_absset : absset \in SAtot. @@ -1421,8 +1442,7 @@ move: size_abs_t; case: t => s; case: s => // x s /=. rewrite eqSS size_eq0 => /eqP -> _. apply/holds_repr_pi => /=. split; last by rewrite normr_ge0. -have [le_e0|lt_0e] := ler0P x; first by right. -by left. +by case: (ler0P x); [right|left]. Qed. Fact SAfun_SAabs : (absset \in SAfunc) && (absset \in SAtot). @@ -1439,7 +1459,8 @@ Proof. apply/SAin_setP/eqP => [/holdsAnd xy|->]; [apply/rowP => i; move: xy => /(_ i (mem_index_enum _) isT) /= | apply/holdsAnd => i _ _ /=]; - (rewrite enum_ordD map_cat nth_catr; last by rewrite 2!size_map size_enum_ord); + (rewrite enum_ordD map_cat nth_catr; + last by rewrite 2!size_map size_enum_ord); rewrite -map_comp (nth_map i) ?size_enum_ord// nth_ord_enum/=; rewrite nth_cat 2!size_map size_enum_ord ltn_ord -map_comp; rewrite (nth_map i) ?size_enum_ord// nth_ord_enum/= !mxE; @@ -1501,58 +1522,55 @@ have h : size ([seq u ord0 i | i <- enum 'I_m] ++ [seq v ord0 i | i <- enum 'I_p]) = (m + p)%N. by rewrite size_cat size_map size_enum_ord size_map size_enum_ord. rewrite /= -[X in nquantify X _ _ _]h. -apply: (sameP (rcf_satP _ _)). -apply: (equivP _ (nexistsP _ _ _)). -apply: (iffP idP); last first. -+ move=> [t] /=. - move=> [ /holds_subst hf /holds_subst hg]. - move: hf hg. - rewrite subst_env_cat -catA. - rewrite subst_env_iota_catl; last by rewrite size_map size_enum_ord. - rewrite catA subst_env_iota_catr ?size_tuple ?card_ord //. - rewrite subst_env_cat subst_env_iota_catr ?size_tuple ?card_ord //. - rewrite -catA subst_env_iota; last 2 first. - - by rewrite size_map size_enum_ord. - - by rewrite size_map size_enum_ord. - rewrite -[t]ngraph_tnth -!ngraph_cat. - move/holds_ngraph; rewrite -inSAfun; move/eqP ->. - by move/holds_ngraph; rewrite -inSAfun; move/eqP ->. -+ move/eqP => eq_gfu_v. +apply/(sameP (rcf_satP _ _))/(equivP _ (nexistsP _ _ _))/(iffP idP). + move/eqP => eq_gfu_v. exists (ngraph (f u)). split; apply/holds_subst; rewrite subst_env_cat. - - rewrite -catA subst_env_iota_catl; last by rewrite size_map size_enum_ord. + rewrite -catA subst_env_iota_catl; last by rewrite size_map size_enum_ord. rewrite catA subst_env_iota_catr ?size_tuple ?card_ord // -ngraph_cat. by apply/holds_ngraph; apply: inSAgraph. - - rewrite subst_env_iota_catr ?size_tuple ?card_ord //. - rewrite -catA subst_env_iota; last 2 first. - by rewrite size_map size_enum_ord. + rewrite subst_env_iota_catr ?size_tuple ?card_ord //. + rewrite -catA subst_env_iota; last 2 first. by rewrite size_map size_enum_ord. - rewrite -ngraph_cat; apply/holds_ngraph; rewrite -eq_gfu_v. - exact: inSAgraph. + by rewrite size_map size_enum_ord. + rewrite -ngraph_cat; apply/holds_ngraph; rewrite -eq_gfu_v. + exact: inSAgraph. +move=> [t] /= [ /holds_subst + /holds_subst]. +rewrite subst_env_cat -catA. +rewrite subst_env_iota_catl; last by rewrite size_map size_enum_ord. +rewrite catA subst_env_iota_catr ?size_tuple ?card_ord //. +rewrite subst_env_cat subst_env_iota_catr ?size_tuple ?card_ord //. +rewrite -catA subst_env_iota; first last. +- by rewrite size_map size_enum_ord. +- by rewrite size_map size_enum_ord. +rewrite -[t]ngraph_tnth -!ngraph_cat. +move/holds_ngraph; rewrite -inSAfun; move/eqP ->. +by move/holds_ngraph; rewrite -inSAfun; move/eqP ->. Qed. Fact SAfun_SAcomp (m n p : nat) (f : SAfun F m n) (g : SAfun F n p) : - (SAcomp_graph f g \in SAfunc) && (SAcomp_graph f g \in SAtot). + (SAcomp_graph f g \in SAfunc) && (SAcomp_graph f g \in SAtot). Proof. apply/andP; split. - by apply/inSAfunc => x y1 y2; rewrite !SAcomp_graphP; move=> /eqP-> /eqP->. + by apply/inSAfunc => x y1 y2; rewrite !SAcomp_graphP => /eqP -> /eqP ->. by apply/inSAtot => x; exists (g (f x)); rewrite SAcomp_graphP. Qed. Definition SAcomp (m n p : nat) (f : SAfun F n p) (g : SAfun F m n) := - MkSAfun (SAfun_SAcomp g f). + MkSAfun (SAfun_SAcomp g f). Lemma SAcompE (m n p : nat) (f : SAfun F n p) (g : SAfun F m n) : - SAcomp f g =1 f \o g. + SAcomp f g =1 f \o g. Proof. move=> x; apply/eqP; rewrite eq_sym -SAcomp_graphP. by move: (inSAgraph (SAcomp f g) x). Qed. Definition SAfun_const_graph n m (x : 'rV[F]_m) : {SAset F^(n + m)%N} := - [set | \big[And/True]_(i : 'I_m) ('X_(@unsplit n m (inr i)) == GRing.Const (x ord0 i))]. + [set | \big[And/True]_(i : 'I_m) + ('X_(@unsplit n m (inr i)) == GRing.Const (x ord0 i))]. -Lemma SAfun_constP n m (x : 'rV[F]_m) (y : 'rV[F]_n) (z : 'rV[F]_m) : +Lemma SAfun_constP n m (x : 'rV[F]_m) y z : row_mx y z \in SAfun_const_graph n x = (z == x). Proof. apply/SAin_setP/eqP => [/holdsAnd zx|->]. @@ -1568,21 +1586,22 @@ Qed. Lemma SAfun_SAfun_const n m (x : 'rV[F]_m) : (SAfun_const_graph n x \in SAfunc) && (SAfun_const_graph n x \in SAtot). Proof. -apply/andP; split; last by apply/inSAtot => y; exists x; rewrite SAfun_constP. -by apply/inSAfunc => x0 y1 y2; rewrite !SAfun_constP => /eqP -> /eqP. +apply/andP; split. + by apply/inSAfunc => x0 y1 y2; rewrite !SAfun_constP => /eqP -> /eqP. +by apply/inSAtot => y; exists x; rewrite SAfun_constP. Qed. Definition SAfun_const n m (x : 'rV[F]_m) := MkSAfun (SAfun_SAfun_const n x). -Lemma SAfun_constE n m (x : 'rV[F]_m) (y : 'rV[F]_n) : - SAfun_const n x y = x. +Lemma SAfun_constE n m (x : 'rV[F]_m) (y : 'rV[F]_n) : SAfun_const n x y = x. Proof. by apply/eqP; rewrite inSAfun /SAfun_const SAfun_constP. Qed. -Definition join_formula (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) : formula F := - (repr (val f)) /\ - (subst_formula (iota 0 m ++ iota (m+n) p) (repr (val g))). +Definition join_formula (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) : formula F := + (repr (val f)) /\ (subst_formula (iota 0 m ++ iota (m+n) p) (repr (val g))). -Lemma nvar_join_formula (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) : +Lemma nvar_join_formula (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) : @nvar F (m + (n + p)) (join_formula f g). Proof. rewrite /nvar /join_formula /=; apply/fsubUsetP; split. @@ -1595,10 +1614,12 @@ case: {f g} p => [|p]; first by rewrite m0fset fsub0set. by rewrite mnfset_sub //= !addnA. Qed. -Definition SAjoin_graph (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) := +Definition SAjoin_graph (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) := \pi_{SAset F^(m + (n + p))} (MkFormulan (nvar_join_formula f g)). -Lemma SAjoin_graphP (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) (u : 'rV[F]_m) (v : 'rV[F]_(n + p)) : +Lemma SAjoin_graphP (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) u v : (row_mx u v \in SAjoin_graph f g) = (row_mx (f u) (g u) == v). Proof. move: (lsubmx v) (rsubmx v) (hsubmxK v) => l r <- {v}. @@ -1622,7 +1643,8 @@ rewrite catA -ngraph_cat subst_env_iota_catr ?size_ngraph// -ngraph_cat. exact/holds_ngraph/inSAgraph. Qed. -Fact SAfun_SAjoin (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) : +Fact SAfun_SAjoin (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) : (SAjoin_graph f g \in SAfunc) && (SAjoin_graph f g \in SAtot). Proof. apply/andP; split. @@ -1630,10 +1652,12 @@ apply/andP; split. by apply/inSAtot => x; exists (row_mx (f x) (g x)); rewrite SAjoin_graphP. Qed. -Definition SAjoin (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) := +Definition SAjoin (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) := MkSAfun (SAfun_SAjoin f g). -Lemma SAjoinE (m n p : nat) (f : {SAfun F^m -> F^n}) (g : {SAfun F^m -> F^p}) (x : 'rV[F]_m) : +Lemma SAjoinE (m n p : nat) (f : {SAfun F^m -> F^n}) + (g : {SAfun F^m -> F^p}) x : SAjoin f g x = row_mx (f x) (g x). Proof. by apply/eqP; rewrite eq_sym -SAjoin_graphP; apply/inSAgraph. Qed. @@ -1651,7 +1675,7 @@ Qed. Definition SAadd_graph p := \pi_{SAset F^(p + p + p)} (MkFormulan (nvar_add_formula p)). -Lemma SAadd_graphP p (u : 'rV[F]_(p + p)) (v : 'rV[F]_p) : +Lemma SAadd_graphP p u v : (row_mx u v \in SAadd_graph p) = (v == lsubmx u + rsubmx u)%R. Proof. rewrite rowPE. @@ -1682,17 +1706,17 @@ Qed. Definition SAadd p := MkSAfun (SAfun_SAadd p). -Lemma SAaddE p (x y : 'rV[F]_p) : SAadd p (row_mx x y) = (x + y)%R. +Lemma SAaddE p x y : SAadd p (row_mx x y) = (x + y)%R. Proof. by apply/eqP; rewrite inSAfun SAadd_graphP row_mxKl row_mxKr. Qed. -Definition SAfun_add (n p : nat) (f g : {SAfun F^n -> F^p}) := +Definition SAfun_add n p (f g : {SAfun F^n -> F^p}) := SAcomp (SAadd p) (SAjoin f g). -Lemma SAfun_addE (n p : nat) (f g : {SAfun F^n -> F^p}) (x : 'rV[F]_n) : +Lemma SAfun_addE n p (f g : {SAfun F^n -> F^p}) x : SAfun_add f g x = (f x + g x)%R. Proof. by rewrite SAcompE/= SAjoinE SAaddE. Qed. -Definition opp_formula (p : nat) : formula F := +Definition opp_formula p : formula F := (\big[And/True]_(i : 'I_p) ('X_(p + i) == - 'X_i))%oT. Lemma nvar_opp_formula p : nvar (p + p) (opp_formula p). @@ -1705,7 +1729,7 @@ Qed. Definition SAopp_graph p := \pi_{SAset F^(p + p)} (MkFormulan (nvar_opp_formula p)). -Lemma SAopp_graphP p (u v : 'rV[F]_p) : +Lemma SAopp_graphP p u v : (row_mx u v \in SAopp_graph p) = (v == - u)%R. Proof. rewrite rowPE. @@ -1729,20 +1753,20 @@ Qed. Definition SAopp p := MkSAfun (SAfun_SAopp p). -Lemma SAoppE p (x : 'rV[F]_p) : SAopp p x = - x. +Lemma SAoppE p x : SAopp p x = - x. Proof. by apply/eqP; rewrite inSAfun SAopp_graphP. Qed. Definition SAfun_opp n p (f : {SAfun F^n -> F^p}) := SAcomp (SAopp p) f. -Lemma SAfun_oppE n p (f : {SAfun F^n -> F^p}) (x : 'rV[F]_n) : +Lemma SAfun_oppE n p (f : {SAfun F^n -> F^p}) x : SAfun_opp f x = - f x. Proof. by rewrite SAcompE/= SAoppE. Qed. -Definition SAfun_sub (n p : nat) (f g : {SAfun F^n -> F^p}) := +Definition SAfun_sub n p (f g : {SAfun F^n -> F^p}) := SAcomp (SAadd p) (SAjoin f (SAcomp (SAopp p) g)). -Lemma SAfun_subE (n p : nat) (f g : {SAfun F^n -> F^p}) (x : 'rV[F]_n) : +Lemma SAfun_subE n p (f g : {SAfun F^n -> F^p}) x : SAfun_sub f g x = f x - g x. Proof. by rewrite SAcompE/= SAjoinE SAcompE/= SAoppE SAaddE. Qed. @@ -1764,7 +1788,8 @@ move=> f; apply/eqP/SAfunE => x. by rewrite SAfun_addE SAfun_constE add0r. Qed. -Lemma SAfun_addNr n p : left_inverse (SAfun_const n (0 : 'rV[F]_p)) (@SAfun_opp n p) (@SAfun_add n p). +Lemma SAfun_addNr n p : + left_inverse (SAfun_const n (0 : 'rV[F]_p)) (@SAfun_opp n p) (@SAfun_add n p). Proof. move=> f; apply/eqP/SAfunE => x. by rewrite SAfun_addE SAfun_oppE SAfun_constE addNr. @@ -1776,8 +1801,9 @@ HB.instance Definition _ n p := GRing.isZmodule.Build {SAfun F^n -> F^p} Definition SAmul_graph p : {SAset F^(p + p + p)} := [set| \big[And/True]_(i < p) ('X_(p.*2 + i) == 'X_i * 'X_(p + i))]. -Lemma SAmul_graphP p (u : 'rV[F]_(p + p)) (v : 'rV[F]_p) : - (row_mx u v \in SAmul_graph p) = (v == \row_i (lsubmx u 0 i * rsubmx u 0 i))%R. +Lemma SAmul_graphP p u v : + (row_mx u v \in SAmul_graph p) + = (v == \row_i (lsubmx u 0 i * rsubmx u 0 i))%R. Proof. rewrite inE rcf_sat_repr_pi rcf_sat_subst -[_ (ngraph _)]cats0. rewrite subst_env_iota_catl ?size_ngraph// rcf_sat_forall rowPE. @@ -1804,20 +1830,22 @@ Qed. Definition SAmul p := MkSAfun (SAfun_SAmul p). -Lemma SAmulE p (x y : 'rV[F]_p) : SAmul p (row_mx x y) = \row_i (x 0 i * y 0 i)%R. +Lemma SAmulE p x y : SAmul p (row_mx x y) = \row_i (x 0 i * y 0 i)%R. Proof. by apply/eqP; rewrite inSAfun SAmul_graphP row_mxKl row_mxKr. Qed. -Definition SAfun_mul (n p : nat) (f g : {SAfun F^n -> F^p}) := +Definition SAfun_mul n p (f g : {SAfun F^n -> F^p}) := SAcomp (SAmul p) (SAjoin f g). -Lemma SAfun_mulE (n p : nat) (f g : {SAfun F^n -> F^p}) (x : 'rV[F]_n) : +Lemma SAfun_mulE n p (f g : {SAfun F^n -> F^p}) x : SAfun_mul f g x = \row_i (f x 0 i * g x 0 i)%R. Proof. by rewrite SAcompE/= SAjoinE SAmulE. Qed. Definition SAinv_graph n : {SAset F^(n + n)} := - [set| \big[And/True]_(i < n) ('X_i * 'X_(n + i) == Const 1%R \/ 'X_i == Const 0%R /\ 'X_(n + i) == Const 0%R)]. + [set| \big[And/True]_(i < n) + ('X_i * 'X_(n + i) == Const 1%R + \/ 'X_i == Const 0%R /\ 'X_(n + i) == Const 0%R)]. -Lemma SAinv_graphP n (x y : 'rV[F]_n) : +Lemma SAinv_graphP n x y : row_mx x y \in SAinv_graph n = (y == \row_i (x 0 i)^-1). Proof. rewrite inE rcf_sat_repr_pi rcf_sat_subst -[_ (ngraph _)]cats0. @@ -1842,30 +1870,30 @@ Qed. Definition SAinv p := MkSAfun (SAfun_SAinv p). -Lemma SAinvE p (x : 'rV[F]_p) : SAinv p x = \row_i (x 0 i)^-1. +Lemma SAinvE p x : SAinv p x = \row_i (x 0 i)^-1. Proof. by apply/eqP; rewrite inSAfun SAinv_graphP. Qed. Definition SAfun_inv n p (f : {SAfun F^n -> F^p}) := SAcomp (SAinv p) f. -Lemma SAfun_invE n p (f : {SAfun F^n -> F^p}) (x : 'rV[F]_n) : +Lemma SAfun_invE n p (f : {SAfun F^n -> F^p}) x : SAfun_inv f x = \row_i (f x 0 i)^-1. Proof. by rewrite SAcompE/= SAinvE. Qed. -Definition SAfun_div (n p : nat) (f g : {SAfun F^n -> F^p}) := +Definition SAfun_div n p (f g : {SAfun F^n -> F^p}) := SAcomp (SAmul p) (SAjoin f (SAcomp (SAinv p) g)). -Lemma SAfun_divE (n p : nat) (f g : {SAfun F^n -> F^p}) (x : 'rV[F]_n) : +Lemma SAfun_divE n p (f g : {SAfun F^n -> F^p}) x : SAfun_div f g x = \row_i (f x 0 i / g x 0 i). Proof. rewrite SAcompE/= SAjoinE SAcompE/= SAinvE SAmulE. by apply/rowP => i; rewrite !mxE. Qed. -Definition SAfun_le (n : nat) (f g : {SAfun F^n -> F^1}) := +Definition SAfun_le n (f g : {SAfun F^n -> F^1}) := SAgraph (SAfun_sub g f) :<=: (SAsetT F n) :*: (SAset_itv `[0, +oo[%R). -Lemma SAfun_leP (n : nat) (f g : {SAfun F^n -> F^1}) : +Lemma SAfun_leP n (f g : {SAfun F^n -> F^1}) : reflect (forall x, f x ord0 ord0 <= g x ord0 ord0) (SAfun_le f g). Proof. apply/(iffP (SAset_subP _ _)) => fg x. @@ -1905,10 +1933,10 @@ HB.instance Definition _ n := Order.Le_isPOrder.Build ring_display (SAfunleType n) (@SAfun_le_refl n) (@SAfun_le_anti n) (@SAfun_le_trans n). -Definition SAfun_lt (n : nat) (f g : {SAfun F^n -> F^1}) := +Definition SAfun_lt n (f g : {SAfun F^n -> F^1}) := SAgraph (SAfun_sub g f) :<=: (SAsetT F n) :*: (SAset_pos F). -Lemma SAfun_ltP (n : nat) (f g : {SAfun F^n -> F^1}) : +Lemma SAfun_ltP n (f g : {SAfun F^n -> F^1}) : reflect (forall x, f x ord0 ord0 < g x ord0 ord0) (SAfun_lt f g). Proof. apply/(iffP (SAset_subP _ _)) => fg x. @@ -1938,10 +1966,10 @@ HB.instance Definition _ n := Order.Lt_isPOrder.Build ring_display (SAfunltType n) (@SAfun_lt_irr n) (@SAfun_lt_trans n). -Definition SAmpoly_graph (n : nat) (p : {mpoly F[n]}) : {SAset F^(n + 1)} := +Definition SAmpoly_graph n (p : {mpoly F[n]}) : {SAset F^(n + 1)} := [set | 'X_n == term_mpoly p (fun i => 'X_i)]. -Lemma SAmpoly_graphP (n : nat) (p : {mpoly F[n]}) (x : 'rV[F]_n) (y : 'rV[F]_1) : +Lemma SAmpoly_graphP n (p : {mpoly F[n]}) x y : (row_mx x y \in SAmpoly_graph p) = (y ord0 ord0 == p.@[x ord0]). Proof. by apply/SAin_setP/eqP; rewrite /= eval_term_mpoly enum_ordD/= map_cat; @@ -1951,7 +1979,7 @@ by apply/SAin_setP/eqP; rewrite /= eval_term_mpoly enum_ordD/= map_cat; rewrite (nth_map i) ?size_enum_ord// nth_ord_enum/= row_mxEl. Qed. -Lemma SAfun_SAmpoly (n : nat) (p : {mpoly F[n]}) : +Lemma SAfun_SAmpoly n p : (SAmpoly_graph p \in @SAfunc _ n 1) && (SAmpoly_graph p \in @SAtot _ n 1). Proof. apply/andP; split. @@ -1961,17 +1989,17 @@ apply/andP; split. by apply/inSAtot => x; exists (\row__ p.@[x ord0]); rewrite SAmpoly_graphP mxE. Qed. -Definition SAmpoly (n : nat) (p : {mpoly F[n]}) := MkSAfun (SAfun_SAmpoly p). +Definition SAmpoly n (p : {mpoly F[n]}) := MkSAfun (SAfun_SAmpoly p). -Lemma SAmpolyE n (p : {mpoly F[n]}) (x : 'rV[F]_n) : +Lemma SAmpolyE n (p : {mpoly F[n]}) x : SAmpoly p x = \row__ p.@[x ord0]. Proof. by apply/eqP; rewrite inSAfun SAmpoly_graphP !mxE. Qed. -Definition SAselect_graph n m (x : seq nat) : {SAset F^(n + m)} := +Definition SAselect_graph n m x : {SAset F^(n + m)} := [set| \big[And/True]_(i : 'I_m) ('X_(n + i) == 'X_(if (n <= (x`_i)%R)%N then ((x`_i)%R + m)%N else x`_i))]. -Lemma SAselect_graphP n m (x : seq nat) (u : 'rV[F]_n) (v : 'rV[F]_m) : +Lemma SAselect_graphP n m x u v : (row_mx u v \in SAselect_graph n m x) = (v == \row_i (ngraph u)`_(x`_i))%R. Proof. apply/SAin_setP/eqP => /= [|->]. @@ -2002,8 +2030,9 @@ rewrite ltnNge (leq_trans ni (leq_addr _ _))/= nth_default; last first. by rewrite nth_default// size_map size_enum_ord -addnBAC// leq_addl. Qed. -Fact SAfun_SAselect n m (x : seq nat) : - (SAselect_graph n m x \in @SAfunc _ n m) && (SAselect_graph n m x \in @SAtot _ n m). +Fact SAfun_SAselect n m x : + (SAselect_graph n m x \in @SAfunc _ n m) + && (SAselect_graph n m x \in @SAtot _ n m). Proof. apply/andP; split. by apply/inSAfunc => u y1 y2; rewrite !SAselect_graphP => /eqP -> /eqP. @@ -2011,13 +2040,13 @@ apply/inSAtot => u; exists (\row_i (ngraph u)`_(x`_i))%R. by rewrite SAselect_graphP eqxx. Qed. -Definition SAselect n m (x : seq nat) := MkSAfun (SAfun_SAselect n m x). +Definition SAselect n m x := MkSAfun (SAfun_SAselect n m x). -Lemma SAselectE n m (x : seq nat) (u : 'rV[F]_n) : +Lemma SAselectE n m x u : SAselect n m x u = \row_i (ngraph u)`_(x`_i). Proof. by apply/eqP; rewrite inSAfun SAselect_graphP. Qed. -Lemma SAselect_iotal n m (x : 'rV[F]_(n + m)) : +Lemma SAselect_iotal n m x : SAselect (n + m) n (iota 0 n) x = lsubmx x. Proof. apply/rowP => i; rewrite SAselectE !mxE nth_iota//. @@ -2032,7 +2061,7 @@ apply/rowP => i; rewrite SAselectE !mxE/= nth_iota//. by rewrite -[X in _`_X]/(nat_of_ord (rshift n i)) nth_map_ord. Qed. -Lemma SAset_castE_select n m (s : {SAset F^n}) : +Lemma SAset_castE_select n m s : SAset_cast m s = SAimset (SAselect n m (iota 0 m)) s. Proof. apply/eqP/SAsetP => x. @@ -2074,8 +2103,7 @@ apply/(SAcomp (SAsum n) _). exact/(SAselect _ _ (iota 0 n)). Defined. -Lemma SAsumE n (u : 'rV[F]_n) : - SAsum n u = \row__ \sum_i (u ord0 i)%R. +Lemma SAsumE n u : SAsum n u = \row__ \sum_i (u ord0 i)%R. Proof. apply/eqP; rewrite rowPE forall_ord1 mxE; apply/eqP. elim: n u => [|n IHn] u; first by rewrite /SAsum SAfun_constE big_ord0 mxE. @@ -2098,10 +2126,11 @@ Definition SAhorner_graph n : {SAset F^(n + 1 + 1)} := \big[And/True]_(i < n) ('X_(n.+2 + i) == ('X_i * 'X_n ^+ i)))]. Lemma SAhorner_graphP n (u : 'rV[F]_(n + 1)) (v : 'rV[F]_1) : - (row_mx u v \in SAhorner_graph n) = (v == \row__ (\poly_(i < n) (ngraph u)`_i).[u ord0 (rshift n ord0)])%R. + (row_mx u v \in SAhorner_graph n) + = (v == \row__ (\poly_(i < n) (ngraph u)`_i).[u ord0 (rshift n ord0)])%R. Proof. -rewrite /SAhorner_graph. -rewrite -2![X in nquantify X]addn1 -[X in nquantify X](size_ngraph (row_mx u v)). +rewrite /SAhorner_graph -2![X in nquantify X]addn1. +rewrite -[X in nquantify X](size_ngraph (row_mx u v)). apply/SAin_setP/eqP => [/nexistsP[x]/= []|vE]. move=> /holds_subst + /holdsAnd/= xE. rewrite -cats1 subst_env_cat/= subst_env_iota_catr; first last. @@ -2130,8 +2159,9 @@ apply/SAin_setP/eqP => [/nexistsP[x]/= []|vE]. have nE: n = lshift 1 (rshift n (@ord0 0)) by rewrite /= addn0. have {2}->: i = lshift 1 (lshift 1 i) :> nat by []. by rewrite [X in _`_X ^+ _]nE !nth_map_ord !mxE !(unsplitK (inl _)) -tnth_nth. -apply/nexistsP; exists ([tuple ((ngraph u)`_i * u ord0 (rshift n ord0) ^+ i)%R | i < n]). -move=> /=; split. +apply/nexistsP. +exists ([tuple ((ngraph u)`_i * u ord0 (rshift n ord0) ^+ i)%R | i < n]) => /=. +split. apply/holds_subst. rewrite -cats1 subst_env_cat/= subst_env_iota_catr; first last. - by rewrite size_map size_enum_ord. @@ -2139,7 +2169,8 @@ move=> /=; split. rewrite nth_cat size_map size_enum_ord 2![in X in (_ < X)%N]addn1 leqnn. have nsE: n.+1 = rshift (n + 1) (@ord0 0) by rewrite /= addn0 addn1. rewrite [X in _`_X]nsE nth_map_ord mxE (unsplitK (inr _)). - suff: SAsum _ (\row_(i < n) ((ngraph u)`_i * u ord0 (rshift n ord0) ^+ i)%R) = v. + suff: SAsum _ (\row_(i < n) ((ngraph u)`_i * u ord0 (rshift n ord0) ^+ i)%R) + = v. move=> /eqP; rewrite inSAfun => /rcf_satP. rewrite ngraph_cat; congr holds; congr cat; last first. by rewrite /= enum_ordSl enum_ord0/=. @@ -2156,37 +2187,41 @@ rewrite nth_cat size_map size_enum_ord 2!{1}addn1 ltnNge (leq_addr _)/=. rewrite 2![in X in (_ - X)%N]addn1 subDnCA// subnn addn0. rewrite nth_cat size_map size_enum_ord 2![in X in (_ - X)%N]addn1. rewrite -[X in (_ < X)%N]addnA (leq_trans (ltn_ord i) (leq_addr _ _)). -rewrite nth_cat size_map size_enum_ord [X in (_ < X + 1)%N]addn1 leq_addr nth_map_ord. +rewrite nth_cat size_map size_enum_ord [X in (_ < X + 1)%N]addn1 leq_addr. +rewrite nth_map_ord. have nE: n = lshift 1 (rshift n (@ord0 0)) by rewrite /= addn0. have {1 3}->: i = lshift 1 (lshift 1 i) :> nat by []. by rewrite [X in _`_X ^+ _]nE !nth_map_ord !mxE !(unsplitK (inl _)). Qed. Fact SAfun_SAhorner n : - (SAhorner_graph n \in @SAfunc _ (n + 1) 1) && (SAhorner_graph n \in @SAtot _ (n + 1) 1). + (SAhorner_graph n \in @SAfunc _ (n + 1) 1) + && (SAhorner_graph n \in @SAtot _ (n + 1) 1). Proof. apply/andP; split. by apply/inSAfunc => u y1 y2; rewrite !SAhorner_graphP => /eqP -> /eqP. -apply/inSAtot => u; exists (\row__ (\poly_(i < n) (ngraph u)`_i).[u ord0 (rshift n ord0)])%R. +apply/inSAtot => u. +exists (\row__ (\poly_(i < n) (ngraph u)`_i).[u ord0 (rshift n ord0)])%R. by rewrite SAhorner_graphP eqxx. Qed. Definition SAhorner n := MkSAfun (SAfun_SAhorner n). -Lemma SAhornerE n (u : 'rV[F]_(n + 1)) : - SAhorner n u = (\row__ (\poly_(i < n) (ngraph u)`_i).[u ord0 (rshift n ord0)])%R. +Lemma SAhornerE n u : + SAhorner n u + = (\row__ (\poly_(i < n) (ngraph u)`_i).[u ord0 (rshift n ord0)])%R. Proof. by apply/eqP; rewrite inSAfun SAhorner_graphP. Qed. -Definition rootsR_formula n (i : nat) := - (((\big[And/True]_(j < n) ('X_j == 0)) /\ (if i == 0%N then True else False)) \/ - \big[And/True]_(j < i) subst_formula +Definition rootsR_formula n i := + (((\big[And/True]_(j < n) ('X_j == 0)) /\ (if i == 0%N then True else False)) + \/ \big[And/True]_(j < i) subst_formula (iota 0 n ++ [:: (n + j)%N; (n + i)%N]) - (SAhorner_graph n) /\ - \big[And/True]_(j < i.-1) ('X_(n + j) <% 'X_(n + j.+1))%oT /\ - ('forall 'X_(n + i), - subst_formula (iota 0 n ++ [:: (n + i)%N; (n + i).+1]) - (SAhorner_graph n) ==> - \big[Or/False]_(j < i) ('X_(n + i) == 'X_(n + j))))%oT. + (SAhorner_graph n) + /\ \big[And/True]_(j < i.-1) ('X_(n + j) <% 'X_(n + j.+1))%oT /\ + ('forall 'X_(n + i), + subst_formula (iota 0 n ++ [:: (n + i)%N; (n + i).+1]) + (SAhorner_graph n) ==> + \big[Or/False]_(j < i) ('X_(n + i) == 'X_(n + j))))%oT. Lemma rootsR_formulaE n (u : 'rV[F]_n) d (r : d.-tuple F) : rcf_sat (ngraph u ++ r) (rootsR_formula n d) @@ -2357,11 +2392,12 @@ Qed. (* Function giving the number of roots of a polynomial of degree at most n.-1 encoded in big endian in F^n *) Definition SAnbroots_graph n : {SAset F^(n + 1)} := - [set| (\big[And/True]_(i < n.+1) ('X_i == 0)) \/ \big[Or/False]_(i < n) (('X_n == Const i%:R%R) /\ - nquantify n.+1 i Exists ( - subst_formula (iota 0 n ++ iota n.+1 i) (rootsR_formula n i)))]. + [set| (\big[And/True]_(i < n.+1) ('X_i == 0)) + \/ \big[Or/False]_(i < n) (('X_n == Const i%:R%R) + /\ nquantify n.+1 i Exists ( + subst_formula (iota 0 n ++ iota n.+1 i) (rootsR_formula n i)))]. -Lemma SAnbroots_graphP n (u : 'rV[F]_n) (v : 'rV[F]_1) : +Lemma SAnbroots_graphP n u v : (row_mx u v \in SAnbroots_graph n) = (v == \row__ (size (rootsR (\poly_(i < n) (ngraph u)`_i)))%:R). Proof. rewrite inE rcf_sat_repr_pi rcf_sat_subst -[_ (ngraph _)]cats0. @@ -2372,7 +2408,8 @@ have ->: P = ((u == 0) && (v == 0)). rewrite /P; apply/forallP/andP => /= [uv0|[] /eqP -> /eqP -> i]; last first. rewrite simp_rcf_sat/=. have ilt: (val i < n + 1)%N by rewrite addn1 ltn_ord. - by rewrite (nth_map_ord _ _ (Ordinal ilt)) mxE; case: (split _) => j; rewrite mxE. + rewrite (nth_map_ord _ _ (Ordinal ilt)) mxE. + by case: (split _) => j; rewrite mxE. split. apply/eqP/rowP => i; move: (uv0 (lift ord_max i)). rewrite simp_rcf_sat/= /bump leqNgt (ltn_ord i)/=. @@ -2389,7 +2426,8 @@ under eq_existsb => /= i. rewrite rcf_sat_nexists; last first. move=> r; rewrite size_resize => ri. rewrite rcf_sat_subst subst_env_cat ngraph_cat -{1}catA. - rewrite subst_env_iota_catl ?size_ngraph// subst_env_iota_catr//; last first. + rewrite subst_env_iota_catl ?size_ngraph//. + rewrite subst_env_iota_catr//; last first. by rewrite size_cat !size_ngraph addn1. move/eqP: ri => ri. rewrite -[r]/(val (Tuple ri)) rootsR_formulaE => /= /eqP rE. @@ -2399,9 +2437,10 @@ under eq_existsb => /= i. have nE': n = rshift n (@ord0 0) by rewrite /= addn0. rewrite [X in _`_X]nE' nth_map_ord mxE (unsplitK (inr _)). rewrite rcf_sat_subst subst_env_cat ngraph_cat -{1}catA. - rewrite subst_env_iota_catl ?size_ngraph// subst_env_iota_catr//; last first. - exact/size_resize. - by rewrite -ngraph_cat. + rewrite subst_env_iota_catl ?size_ngraph//. + rewrite subst_env_iota_catr//; first last. + - exact/size_resize. + - by rewrite -ngraph_cat. move/eqP: (size_resize 0 (rootsR p) i) => ri. rewrite -[resize _ _ _]/(val (Tuple ri)) rootsR_formulaE/= resize_idE. over. @@ -2438,13 +2477,14 @@ Fact SAfun_SAnbroots n : Proof. apply/andP; split. by apply/inSAfunc => u y1 y2; rewrite !SAnbroots_graphP => /eqP -> /eqP. -apply/inSAtot => u; exists (\row__ (size (rootsR (\poly_(i < n) (ngraph u)`_i)))%:R)%R. +apply/inSAtot => u. +exists (\row__ (size (rootsR (\poly_(i < n) (ngraph u)`_i)))%:R)%R. by rewrite SAnbroots_graphP eqxx. Qed. Definition SAnbroots n := MkSAfun (SAfun_SAnbroots n). -Lemma SAnbrootsE n (u : 'rV[F]_n) : +Lemma SAnbrootsE n u : SAnbroots n u = (\row__ (size (rootsR (\poly_(i < n) (ngraph u)`_i)))%:R)%R. Proof. by apply/eqP; rewrite inSAfun SAnbroots_graphP. Qed. @@ -2458,14 +2498,16 @@ Lemma SAnthroot_graphP n m (u : 'rV[F]_n) (v : 'rV[F]_1) : == \row__ ((rootsR (\poly_(i < n) (ngraph u)`_i))`_m)). Proof. rewrite inE rcf_sat_repr_pi rcf_sat_subst -[_ (ngraph _)]cats0. -rewrite subst_env_iota_catl ?size_ngraph// rcf_sat_Or rcf_sat_forall rcf_sat_exists. +rewrite subst_env_iota_catl ?size_ngraph// rcf_sat_Or rcf_sat_forall. +rewrite rcf_sat_exists. set p := \poly_(i < n) (ngraph u)`_i. set P := [forall _, _]. have ->: P = ((u == 0) && (v == 0)). rewrite /P; apply/forallP/andP => /= [uv0|[] /eqP -> /eqP -> i]; last first. rewrite simp_rcf_sat/=. have ilt: (val i < n + 1)%N by rewrite addn1 ltn_ord. - by rewrite (nth_map_ord _ _ (Ordinal ilt)) mxE; case: (split _) => j; rewrite mxE. + rewrite (nth_map_ord _ _ (Ordinal ilt)) mxE. + by case: (split _) => j; rewrite mxE. split. apply/eqP/rowP => i; move: (uv0 (lift ord_max i)). rewrite simp_rcf_sat/= /bump leqNgt (ltn_ord i)/=. @@ -2494,9 +2536,10 @@ under eq_existsb => /= i. have nE': n = rshift n (@ord0 0) by rewrite /= addn0. rewrite [X in _`_X]nE' nth_map_ord mxE (unsplitK (inr _)). rewrite rcf_sat_subst subst_env_cat ngraph_cat -{1}catA. - rewrite subst_env_iota_catl ?size_ngraph// subst_env_iota_catr//; last first. - exact/size_resize. - by rewrite -ngraph_cat. + rewrite subst_env_iota_catl ?size_ngraph//. + rewrite subst_env_iota_catr//; first last. + - exact/size_resize. + - by rewrite -ngraph_cat. move/eqP: (size_resize 0 (rootsR p) i) => ri. rewrite -[resize _ _ _]/(val (Tuple ri)) rootsR_formulaE/= resize_idE. over. @@ -2529,7 +2572,8 @@ by rewrite coef_poly coef0 (ltn_ord i) nth_map_ord mxE. Qed. Fact SAfun_SAnthroot n i : - (SAnthroot_graph n i \in @SAfunc _ n 1) && (SAnthroot_graph n i \in @SAtot _ n 1). + (SAnthroot_graph n i \in @SAfunc _ n 1) + && (SAnthroot_graph n i \in @SAtot _ n 1). Proof. apply/andP; split. by apply/inSAfunc => u y1 y2; rewrite !SAnthroot_graphP => /eqP -> /eqP. @@ -2546,11 +2590,13 @@ Proof. by apply/eqP; rewrite inSAfun SAnthroot_graphP. Qed. (* TODO: See if rcf_sat_nexists shortens the previous proofs. *) Definition SAmulc_graph : {SAset F^((2 + 2) + 2)} := - [set| 'X_4 == ('X_0 * 'X_2 - 'X_1 * 'X_3) /\ 'X_5 == ('X_0 * 'X_3 + 'X_1 * 'X_2)]. + [set| 'X_4 == ('X_0 * 'X_2 - 'X_1 * 'X_3) + /\ 'X_5 == ('X_0 * 'X_3 + 'X_1 * 'X_2)]. Lemma SAmulc_graphP (u v w : 'rV[F]_2) : row_mx (row_mx u v) w \in SAmulc_graph = - (let x := ((u ord0 ord0 +i* u ord0 ord_max) * (v ord0 ord0 +i* v ord0 ord_max))%C in + (let x := + ((u ord0 ord0 +i* u ord0 ord_max) * (v ord0 ord0 +i* v ord0 ord_max))%C in (w == \row_i if i == 0 then complex.Re x else complex.Im x)). Proof. rewrite inE rcf_sat_repr_pi rcf_sat_subst -[_ (ngraph _)]cats0. @@ -2571,26 +2617,33 @@ by rewrite rowPE forall_ord2 !mxE/=. Qed. Fact SAfun_SAmulc : - (SAmulc_graph \in @SAfunc _ (2 + 2) 2) && (SAmulc_graph \in @SAtot _ (2 + 2) 2). + (SAmulc_graph \in @SAfunc _ (2 + 2) 2) + && (SAmulc_graph \in @SAtot _ (2 + 2) 2). Proof. apply/andP; split. - by apply/inSAfunc => u y1 y2; rewrite -[u]hsubmxK !SAmulc_graphP => /eqP -> /eqP. + apply/inSAfunc => u y1 y2; rewrite -[u]hsubmxK !SAmulc_graphP. + by move=> /eqP -> /eqP. apply/inSAtot => u. -pose x := ((lsubmx u ord0 ord0 +i* lsubmx u ord0 ord_max) * (rsubmx u ord0 ord0 +i* rsubmx u ord0 ord_max))%C. +pose x := ((lsubmx u ord0 ord0 +i* lsubmx u ord0 ord_max) + * (rsubmx u ord0 ord0 +i* rsubmx u ord0 ord_max))%C. exists (\row_i if i == 0 then complex.Re x else complex.Im x). by rewrite -[u]hsubmxK SAmulc_graphP. Qed. Definition SAmulc := MkSAfun SAfun_SAmulc. -Lemma SAmulcE (u v : 'rV[F]_2) : +Lemma SAmulcE u v : SAmulc (row_mx u v) = - (let x := ((u ord0 ord0 +i* u ord0 ord_max) * (v ord0 ord0 +i* v ord0 ord_max))%C in + (let x := ((u ord0 ord0 +i* u ord0 ord_max) + * (v ord0 ord0 +i* v ord0 ord_max))%C in \row_i if i == 0 then complex.Re x else complex.Im x). Proof. by apply/eqP; rewrite inSAfun SAmulc_graphP. Qed. Fixpoint SAexpc_subdef n : - {f : {SAfun F^2 -> F^2} | forall u : 'rV[F]_2, let x := (u ord0 ord0 +i* u ord0 ord_max)%C ^+ n in (f u = \row_(i < 2) if i == 0 then complex.Re x else complex.Im x)}. + {f : {SAfun F^2 -> F^2} | + forall u : 'rV[F]_2, + let x := (u ord0 ord0 +i* u ord0 ord_max)%C ^+ n in + (f u = \row_(i < 2) if i == 0 then complex.Re x else complex.Im x)}. Proof. case: n => [|n]. exists (SAfun_const 2 (\row_(i < 2) (i == 0)%:R)) => u/=. @@ -2606,7 +2659,7 @@ Qed. Definition SAexpc n := proj1_sig (SAexpc_subdef n). -Lemma SAexpcE n (u : 'rV[F]_2) : +Lemma SAexpcE n u : SAexpc n u = let x := (u ord0 ord0 +i* u ord0 ord_max)%C ^+ n in \row_(i < 2) if i == 0 then complex.Re x else complex.Im x. Proof. exact: (proj2_sig (SAexpc_subdef n) u). Qed. @@ -2623,7 +2676,13 @@ elim: n => [|n [f] fP]. by apply/eqP; rewrite SAfun_constE rowPE forall_ord2 !mxE/= eqxx. rewrite -[RHS](horner0 x); congr horner. by apply/eqP; rewrite -size_poly_eq0 -leqn0; apply/size_poly. -exists (SAfun_add (SAcomp f (SAselect _ _ (iota 0 n ++ iota n.+1 2))) (SAcomp SAmulc (SAjoin (SAselect _ _ [:: n; n.+3]) (SAcomp (SAexpc n) (SAselect _ _ (iota n.+1 2)))))). +exists ( + SAfun_add + (SAcomp f (SAselect _ _ (iota 0 n ++ iota n.+1 2))) + (SAcomp SAmulc + (SAjoin + (SAselect _ _ [:: n; n.+3]) + (SAcomp (SAexpc n) (SAselect _ _ (iota n.+1 2)))))). move=> u x r. rewrite SAfun_addE !SAcompE/= SAjoinE SAcompE/= 3!SAselectE SAmulcE SAexpcE. apply/rowP => i; rewrite !mxE. @@ -2679,16 +2738,29 @@ Proof. exact/(proj2_sig (SAhornerRC_subdef n)). Qed. (* Function giving the number of complex roots of a polynomial of degree at most n.-1 encoded in big endian in F^n *) Definition SAnbrootsC_graph n : {SAset F^(n + 1)} := - [set| (\big[And/True]_(i < n.+1) ('X_i == 0)) \/ \big[Or/False]_(d < n) (('X_n == Const d%:R%R) /\ - nquantify n.+1 d.*2 Exists ( - \big[And/True]_(j < d) subst_formula (iota 0 n ++ [:: n.+1 + j.*2; n.+1 + j.*2.+1; n.+1 + d.*2; n.+1 + d.*2]%N) - (SAhornerRC n) /\ - \big[And/True]_(i < d) \big[And/True]_(j < d | j != i) ('X_(n.+1 + i.*2) != 'X_(n.+1 + j.*2) \/ 'X_(n.+1 + i.*2.+1) != 'X_(n.+1 + j.*2.+1)) /\ - nquantify (n.+1 + d.*2) 2 Forall (subst_formula (iota 0 n ++ [:: n.+1 + d.*2; n.+1 + d.*2.+1; n.+1 + d.*2.+2; n.+1 + d.*2.+2]%N) - (SAhornerRC n) ==> \big[Or/False]_(j < d) ('X_(n.+1 + d.*2) == 'X_(n.+1 + j.*2) /\ 'X_(n.+1 + d.*2.+1) == 'X_(n.+1 + j.*2.+1)))))]. - -Lemma SAnbrootsC_graphP n (u : 'rV[F]_n) (v : 'rV[F]_1) : - (row_mx u v \in SAnbrootsC_graph n) = (v == \row__ (size (dec_roots (\poly_(i < n) ((ngraph u)`_i)%:C%C)))%:R). + [set| (\big[And/True]_(i < n.+1) ('X_i == 0)) + \/ \big[Or/False]_(d < n) (('X_n == Const d%:R%R) + /\ nquantify n.+1 d.*2 Exists ( + \big[And/True]_(j < d) + subst_formula + (iota 0 n ++ + [:: n.+1 + j.*2; n.+1 + j.*2.+1; n.+1 + d.*2; n.+1 + d.*2]%N) + (SAhornerRC n) + /\ \big[And/True]_(i < d) \big[And/True]_(j < d | j != i) + ('X_(n.+1 + i.*2) != 'X_(n.+1 + j.*2) + \/ 'X_(n.+1 + i.*2.+1) != 'X_(n.+1 + j.*2.+1)) + /\ nquantify (n.+1 + d.*2) 2 Forall + (subst_formula + (iota 0 n ++ + [:: n.+1 + d.*2; n.+1 + d.*2.+1; n.+1 + d.*2.+2; + n.+1 + d.*2.+2]%N) + (SAhornerRC n) ==> \big[Or/False]_(j < d) + ('X_(n.+1 + d.*2) == 'X_(n.+1 + j.*2) + /\ 'X_(n.+1 + d.*2.+1) == 'X_(n.+1 + j.*2.+1)))))]. + +Lemma SAnbrootsC_graphP n u v : + (row_mx u v \in SAnbrootsC_graph n) + = (v == \row__ (size (dec_roots (\poly_(i < n) ((ngraph u)`_i)%:C%C)))%:R). Proof. move uvE: (tval (ngraph (row_mx u v))) => uv. move: uvE; have [->|u0] := eqVneq u 0 => uvE. @@ -2696,36 +2768,39 @@ move: uvE; have [->|u0] := eqVneq u 0 => uvE. apply/polyP => i; rewrite coef_poly coef0. case: (ltnP i n) => [ilt|//]. by rewrite (nth_mktuple _ _ (Ordinal ilt)) mxE. - rewrite dec_roots0/=; apply/SAin_setP/eqP => [/= [/holdsAnd|/holdsOr-[] i]| ->]. + rewrite dec_roots0/=. + apply/SAin_setP/eqP => [/= [/holdsAnd|/holdsOr-[] i]| ->]. - move=> /(_ ord_max (mem_index_enum _) isT) /=. have nE: n = rshift n (@ord0 0) by rewrite /= addn0. rewrite [X in _`_X]nE nth_map_ord mxE (unsplitK (inr _)) => v0. by apply/eqP; rewrite rowPE forall_ord1 mxE; apply/eqP. - - move=> [_][_]/= [_]. - rewrite -[X in nquantify X]addn1 -[X in nquantify X](size_ngraph (row_mx 0 v)). + - move=> [_][_]/= [_]; rewrite -[X in nquantify X]addn1. + rewrite -[X in nquantify X](size_ngraph (row_mx 0 v)). move=> /nexistsP[r]/= [_][_]; rewrite uvE. have suvr: (n.+1 + i.*2)%N = size (uv ++ r). by rewrite -uvE size_cat size_ngraph size_tuple addn1. - rewrite suvr. - move=> /nforallP-/(_ (mktuple (fun=> 1 + \big[Order.max/0]_(x <- r) x)))%R /=. + rewrite suvr => /nforallP. + move=> /(_ (mktuple (fun=> 1 + \big[Order.max/0]_(x <- r) x)))%R /=. mp. apply/holds_subst; rewrite subst_env_cat. - rewrite -{1}uvE/= {1}enum_ordD map_cat -!catA subst_env_iota_catl; last first. - by rewrite 2!size_map size_enum_ord. + rewrite -{1}uvE/= {1}enum_ordD map_cat -!catA. + rewrite subst_env_iota_catl; last by rewrite 2!size_map size_enum_ord. rewrite catA nth_cat ltnn subnn enum_ordSl/=. rewrite nth_cat [X in (X < _)%N]addnS suvr ltnNge leqnSn/=. rewrite -suvr subnDl subSn// subnn enum_ordSl/=. rewrite nth_default; last first. by rewrite !addnS suvr size_cat/= enum_ord0/= addn2. - have /eqP: SAhornerRC n (row_mx 0 (\row__ (1 + \big[maxr/0]_(x <- r) x)%R)) = \row__ 0. - apply/eqP; rewrite SAhornerRCE rowPE forall_ord2 !mxE/= !(unsplitK (inr _)). + have: SAhornerRC n (row_mx 0 (\row__ (1 + \big[maxr/0]_(x <- r) x)%R)) + = \row__ 0. + apply/eqP; rewrite SAhornerRCE rowPE forall_ord2 !mxE/=. + rewrite !(unsplitK (inr _)). move pE : (poly _ _) => p. suff ->: p = 0 by rewrite horner0/= eqxx. apply/polyP => j; rewrite -pE coef0 coef_poly. case: (ltnP j n) => [jn|//]. rewrite ngraph_cat nth_cat size_ngraph jn. by rewrite (nth_mktuple _ _ (Ordinal jn)) mxE. - rewrite inSAfun => /rcf_satP; rewrite !ngraph_cat -catA. + move=> /eqP; rewrite inSAfun => /rcf_satP; rewrite !ngraph_cat -catA. congr (holds (_ ++ _) _); last by rewrite /= !enum_ordSl enum_ord0/= !mxE. apply/(@eq_from_nth _ 0) => [|k]; rewrite size_ngraph. by rewrite 2!size_map size_enum_ord. @@ -2744,7 +2819,8 @@ move: uvE; have [->|u0] := eqVneq u 0 => uvE. left; apply/holdsAnd; case=> i /= ilt _ _ /=. rewrite enum_ordD map_cat -2!map_comp nth_cat size_map size_enum_ord. case: (ltnP i n) => iltn. - by rewrite -/(nat_of_ord (Ordinal iltn)) nth_map_ord mxE (unsplitK (inl _)) mxE. + rewrite -/(nat_of_ord (Ordinal iltn)) nth_map_ord mxE (unsplitK (inl _)). + by rewrite mxE. have ->: i = n by apply/le_anti/andP. rewrite subnn -[X in _`_X]/(nat_of_ord (@ord0 0)) nth_map_ord mxE. by rewrite (unsplitK (inr _)) mxE. @@ -2755,19 +2831,20 @@ have pu0: \poly_(i < n) (([seq u ord0 i0 | i0 <- enum 'I_n]`_i)%:C)%C != 0. by move/complexI. have ComplexK (x : F[i]): (complex.Re x +i* complex.Im x)%C = x. by apply/eqP; rewrite eq_complex !eqxx. -rewrite inE rcf_sat_repr_pi rcf_sat_subst uvE -[uv]cats0 subst_env_iota_catl; last first. - by rewrite -uvE size_ngraph. +rewrite inE rcf_sat_repr_pi rcf_sat_subst uvE -[uv]cats0. +rewrite subst_env_iota_catl; last by rewrite -uvE size_ngraph. rewrite rcf_sat_Or rcf_sat_forall. have /negP/negPf -> /=: ~ [forall i : 'I_(n.+1), rcf_sat uv ('X_i == 0)]. move=> /forallP /= uv0. move: u0; rewrite rowPE => /forallPn/= [] i. move: (uv0 (lift ord_max i)) => /rcf_satP/=. - rewrite -uvE ngraph_cat nth_cat /bump [(n <= i)%N]leqNgt size_ngraph !(ltn_ord i)/=. - by rewrite nth_map_ord mxE => -> /eqP. + rewrite -uvE ngraph_cat nth_cat /bump [(n <= i)%N]leqNgt size_ngraph. + by rewrite !(ltn_ord i)/= nth_map_ord mxE => -> /eqP. apply/rcf_satP/eqP => [/holdsOr/=[] d [_][_]|vE]. rewrite -{1}uvE ngraph_cat nth_cat size_ngraph ltnn. rewrite subnn (nth_map_ord _ _ ord0) => -[] vE. - rewrite -[X in nquantify X]addn1 -[X in nquantify X](size_ngraph (row_mx u v)) uvE. + rewrite -[X in nquantify X]addn1. + rewrite -[X in nquantify X](size_ngraph (row_mx u v)) uvE. move=> /nexistsP[r]/= [] /holdsAnd/= rroot [] runiq rall. set r' := (mktuple (fun i : 'I_d => (r`_(val i).*2 +i* r`_(val i).*2.+1)%C)). apply/eqP; rewrite rowPE forall_ord1 vE mxE eqr_nat -(size_tuple r'). @@ -2776,7 +2853,8 @@ apply/rcf_satP/eqP => [/holdsOr/=[] d [_][_]|vE]. rewrite size_map size_enum_ord => jd. rewrite (nth_map_ord _ _ (Ordinal (ltn_trans ij jd))). rewrite (nth_map_ord _ _ (Ordinal jd)) => -[] rij rij1. - move/holdsAnd: runiq => /= /(_ (Ordinal (ltn_trans ij jd)) (mem_index_enum _) isT). + move/holdsAnd: runiq => /=. + move=> /(_ (Ordinal (ltn_trans ij jd)) (mem_index_enum _) isT). move=> /holdsAnd /= /(_ (Ordinal jd) (mem_index_enum _)). rewrite -(inj_eq val_inj)/=. mp; first by apply/eqP => ji; rewrite ji ltnn in ij. @@ -2789,7 +2867,8 @@ apply/rcf_satP/eqP => [/holdsOr/=[] d [_][_]|vE]. rewrite size_map size_enum_ord => id <-. rewrite (nth_map_ord _ _ (Ordinal id)). move: rroot => /(_ (Ordinal id) (mem_index_enum _) isT) /holds_subst. - rewrite subst_env_cat -{1}uvE ngraph_cat -catA subst_env_iota_catl ?size_ngraph//=. + rewrite subst_env_cat -{1}uvE ngraph_cat -catA. + rewrite subst_env_iota_catl ?size_ngraph//=. rewrite !nth_cat -![X in size X]uvE size_ngraph addn1. do 3 (rewrite ltnNge leq_addr/= subDnCA// subnn addn0). rewrite [r`_d.*2]nth_default ?size_tuple// => ru0. @@ -2809,15 +2888,19 @@ apply/rcf_satP/eqP => [/holdsOr/=[] d [_][_]|vE]. move: rall. have suvr: size (uv ++ r) = (n.+1 + d.*2)%N. by rewrite size_cat -uvE size_ngraph size_tuple addn1. - rewrite -suvr => /nforallP /(_ (mktuple (fun i => if i == 0 then complex.Re x else complex.Im x)))/=. + rewrite -suvr => /nforallP. + move=> /(_ (mktuple + (fun i => if i == 0 then complex.Re x else complex.Im x)))/=. mp. apply/holds_subst. - rewrite subst_env_cat -{1}uvE ngraph_cat -!catA subst_env_iota_catl ?size_ngraph//=. - rewrite catA !nth_cat ltnn subnn suvr !addnS ltnNge leqnSn/=. - rewrite ltnNge (leq_trans (leqnSn _) (leqnSn _))/=. + rewrite subst_env_cat -{1}uvE ngraph_cat -!catA. + rewrite subst_env_iota_catl ?size_ngraph//= catA !nth_cat ltnn subnn suvr. + rewrite !addnS ltnNge leqnSn/= ltnNge (leq_trans (leqnSn _) (leqnSn _))/=. rewrite subSn// subnn subSn// subSn// subnn !enum_ordSl enum_ord0/=. - suff /eqP: SAhornerRC n (row_mx u (\row_j if j == 0 then complex.Re x else complex.Im x)) = \row__ 0. - rewrite inSAfun => /rcf_satP; rewrite !ngraph_cat -catA. + suff: SAhornerRC n + (row_mx u (\row_j if j == 0 then complex.Re x else complex.Im x)) + = \row__ 0. + move=> /eqP; rewrite inSAfun => /rcf_satP; rewrite !ngraph_cat -catA. by rewrite /= !enum_ordSl enum_ord0/= !mxE/= /bump/=. rewrite SAhornerRCE/= !mxE !(unsplitK (inr _)) !mxE. apply/eqP; rewrite rowPE forall_ord2 !mxE/=. @@ -2850,21 +2933,30 @@ exists (Ordinal sn) => /=. split; first exact/mem_index_enum. split=> //. split. - by rewrite -uvE ngraph_cat nth_cat size_ngraph ltnn subnn vE/= enum_ordSl/= mxE. + rewrite -uvE ngraph_cat nth_cat size_ngraph ltnn subnn vE/= enum_ordSl/=. + by rewrite mxE. have ->: n.+1 = size uv by rewrite -uvE size_ngraph addn1. apply/nexistsP. -exists (mktuple (fun i => if odd i then complex.Im (dec_roots p)`_i./2 else complex.Re (dec_roots p)`_i./2)%N). +exists (mktuple (fun i => + if odd i + then complex.Im (dec_roots p)`_i./2 + else complex.Re (dec_roots p)`_i./2)%N). split. apply/holdsAnd => /= i _ _; apply/holds_subst. - rewrite subst_env_cat -{1}uvE ngraph_cat -!catA subst_env_iota_catl ?size_ngraph//=. + rewrite subst_env_cat -{1}uvE ngraph_cat -!catA. + rewrite subst_env_iota_catl ?size_ngraph//=. do 3 rewrite nth_cat ltnNge leq_addr/= subDnCA// subnn addn0. move: (ltn_ord i); rewrite -ltn_double => i2lt. rewrite (nth_map_ord _ _ (Ordinal i2lt))/= odd_double doubleK. move: (ltn_ord i); rewrite -ltn_Sdouble => i2slt. rewrite (nth_map_ord _ _ (Ordinal i2slt))/= odd_double/= uphalf_double. rewrite [(map _ _)`__]nth_default; last by rewrite size_map size_enum_ord. - suff /eqP: SAhornerRC n (row_mx u (\row_j if j == 0 then complex.Re (dec_roots p)`_i else complex.Im (dec_roots p)`_i)) = \row__ 0. - rewrite inSAfun => /rcf_satP; rewrite !ngraph_cat -catA. + suff: SAhornerRC n (row_mx u (\row_j + if j == 0 + then complex.Re (dec_roots p)`_i + else complex.Im (dec_roots p)`_i)) + = \row__ 0. + move=> /eqP; rewrite inSAfun => /rcf_satP; rewrite !ngraph_cat -catA. by rewrite /= !enum_ordSl enum_ord0/= !mxE/= /bump/=. rewrite SAhornerRCE/= !mxE !(unsplitK (inr _)) !mxE. apply/eqP; rewrite rowPE forall_ord2 !mxE/= ComplexK. @@ -2895,8 +2987,9 @@ split. move tE: (mktuple _) => t. rewrite -[X in (_ + X)%N](size_tuple t) -size_cat. apply/nforallP => w/= /holds_subst. -rewrite subst_env_cat -{1}uvE ngraph_cat -!catA subst_env_iota_catl ?size_ngraph//=. -rewrite !addnS -[in X in (_ + X)%N](size_tuple t) -size_cat catA. +rewrite subst_env_cat -{1}uvE ngraph_cat -!catA. +rewrite subst_env_iota_catl ?size_ngraph//= !addnS. +rewrite -[in X in (_ + X)%N](size_tuple t) -size_cat catA. rewrite nth_cat ltnNge leqnn/= subnn. rewrite nth_cat ltnNge leqnSn/= subSn// subnn. rewrite [(_ ++ _)`__]nth_default => [w0|]; last first. @@ -2931,22 +3024,28 @@ by rewrite odd_double uphalf_double iE. Qed. Fact SAfun_SAnbrootsC n : - (SAnbrootsC_graph n \in @SAfunc _ n 1) && (SAnbrootsC_graph n \in @SAtot _ n 1). + (SAnbrootsC_graph n \in @SAfunc _ n 1) + && (SAnbrootsC_graph n \in @SAtot _ n 1). Proof. apply/andP; split. by apply/inSAfunc => u y1 y2; rewrite !SAnbrootsC_graphP => /eqP -> /eqP. -apply/inSAtot => u; exists (\row__ (size (dec_roots (\poly_(i < n) ((ngraph u)`_i)%:C%C)))%:R)%R. +apply/inSAtot => u. +exists (\row__ (size (dec_roots (\poly_(i < n) ((ngraph u)`_i)%:C%C)))%:R)%R. by rewrite SAnbrootsC_graphP. Qed. Definition SAnbrootsC n := MkSAfun (SAfun_SAnbrootsC n). -Lemma SAnbrootsCE n (u : 'rV[F]_n) : - SAnbrootsC n u = (\row__ (size (dec_roots (\poly_(i < n) ((ngraph u)`_i)%:C%C)))%:R)%R. +Lemma SAnbrootsCE n u : + SAnbrootsC n u + = (\row__ (size (dec_roots (\poly_(i < n) ((ngraph u)`_i)%:C%C)))%:R)%R. Proof. by apply/eqP; rewrite inSAfun SAnbrootsC_graphP. Qed. Definition SAset_lt (s t : {SAset F^1}) := - (t != SAset0 F 1) && rcf_sat [::] ('forall 'X_0, s ==> 'forall 'X_1, subst_formula [:: 1%N] t ==> ('X_0 <% 'X_1))%oT. + (t != SAset0 F 1) + && rcf_sat [::] + ('forall 'X_0, s ==> 'forall 'X_1, subst_formula [:: 1%N] t + ==> ('X_0 <% 'X_1))%oT. Lemma SAset_ltP (s t : {SAset F^1}) : reflect (t != SAset0 F 1 /\ forall x y, x \in s -> y \in t -> x 0 0 < y 0 0) @@ -3109,7 +3208,8 @@ case: (posnP i) => i0. rewrite inSAset_bigcup; apply/hasP. exists (nth (SAset0 F 1) (partition_of_pts ((y 0 0) :: xi)) 1). by apply/mem_nth; rewrite size_map size_iota/= doubleS. - by rewrite (nth_map 0)// nth_iota//= inSAset_seq mem_seq1 rowPE forall_ord1 mxE. + rewrite (nth_map 0)// nth_iota//= inSAset_seq mem_seq1 rowPE forall_ord1. + by rewrite mxE. rewrite inSAset_bigcup; apply/hasP. exists (nth (SAset0 F 1) (partition_of_pts (x :: xi)) 2). by apply/mem_nth; rewrite size_map size_iota/= doubleS. @@ -3159,7 +3259,8 @@ rewrite inS; apply/negP => xi0. (* What??? *) move: (elimT (nthP (SAset0 F 1)) xi0) => {xi0} [] i. rewrite size_map size_iota; case: (posnP i) => [->|i0] ixi; last first. - move: xisort => /sorted_partition_of_pts /(lt_sorted_ltn_nth (SAset0 F 1 : SAsetltType)). + move: xisort => /sorted_partition_of_pts. + move=> /(lt_sorted_ltn_nth (SAset0 F 1 : SAsetltType)). move=> /(_ 0 i); rewrite !inE size_map size_iota => /(_ isT ixi). by rewrite i0 => /SAset_ltP[] /eqP + _. rewrite (nth_map 0) ?size_iota// nth_iota//= => xi0. @@ -3171,7 +3272,8 @@ rewrite leq0n => x0j; apply/(lt_le_trans _ x0j). by rewrite -subr_gt0 opprB addrCA subrr addr0. Qed. -Definition partition_of_graphs (n : nat) (xi : seq {SAfun F^n -> F^1}) : seq {SAset F^(n + 1)%N} := +Definition partition_of_graphs n + (xi : seq {SAfun F^n -> F^1}) : seq {SAset F^(n + 1)%N} := [seq if i == 0 then \big[@SAsetI F (n + 1)/SAsetT F (n + 1)]_(f <- xi) SAhypograph f @@ -3262,7 +3364,7 @@ Qed. Lemma SApreimset0 n m f : SApreimset f (SAset0 F n) = SAset0 F m. Proof. apply/eqP; rewrite -subset0; apply/SAset_subP => x. -by rewrite inSApreimset inSAset0. +by rewrite inSApreimset [_ _ \in _]inSAset0. Qed. Lemma SAset_partition_of_graphs (n : nat) (xi : seq (SAfunltType n)) : @@ -3272,7 +3374,8 @@ set S := [fset x | x in partition_of_graphs xi]. have inS x : x \in S = (x \in partition_of_graphs xi). by apply/imfsetP/idP => [[] y yS -> //|xS]; exists x. move=> /(lt_sorted_ltn_nth (0 : SAfunltType n)) xisort. -have {}xisort x : path.sorted <%O [seq (f : {SAfun F^n -> F^1}) x 0 0 | f <- xi]. +have {}xisort x : + path.sorted <%O [seq (f : {SAfun F^n -> F^1}) x 0 0 | f <- xi]. apply/path.pairwise_sorted/(pairwiseP 0) => i j. rewrite !inE size_map => ilt jlt ij. move: (xisort i j); rewrite !inE => /(_ ilt jlt); rewrite ij. @@ -3317,7 +3420,8 @@ set xi' := [seq (f : {SAfun F^n -> F^1}) (lsubmx x) 0 0 | f <- xi]. set T := [fset x | x in _] => /andP[] /andP[_] + _. have inT y : y \in T = (y \in partition_of_pts xi'). by apply/imfsetP/idP => [[] z zS -> //|yS]; exists y. -have nk k: (k < (size xi).*2.+1)%N -> nth (SAset0 F _) (partition_of_pts xi') k \in T. +have nk k: (k < (size xi).*2.+1)%N -> + nth (SAset0 F _) (partition_of_pts xi') k \in T. by rewrite inT => klt; apply/mem_nth; rewrite size_map size_iota size_map. move=> /forallP/(_ [` nk i ilt]) /forallP/(_ [` nk j jlt]) /implyP/=. rewrite nth_uniq ?size_map ?size_iota//; last first. @@ -3330,16 +3434,21 @@ rewrite (nth_map (SAset0 F _)) ?size_map ?size_iota//. by rewrite !inSAset_fiber hsubmxK xii. Qed. -Definition partition_of_graphs_above n (s : {SAset F^n}) (xi : seq {SAfun F^n -> F^1}) : {fset {SAset F^(n + 1)%N}} := +Definition partition_of_graphs_above n (s : {SAset F^n}) + (xi : seq {SAfun F^n -> F^1}) : {fset {SAset F^(n + 1)%N}} := [fset x :&: (s :*: SAsetT F 1) | x in partition_of_graphs xi]. -Lemma SAset_partition_of_graphs_above n (S : {fset {SAset F^n}}) (xi : S -> seq (SAfunltType n)) : +Lemma SAset_partition_of_graphs_above n (S : {fset {SAset F^n}}) + (xi : S -> seq (SAfunltType n)) : SAset_partition S -> (forall s, path.sorted <%O (xi s)) -> - SAset_partition (\big[fsetU/fset0]_(s : S) partition_of_graphs_above (val s) (in_tuple (xi s))). + SAset_partition + (\big[fsetU/fset0]_(s : S) + partition_of_graphs_above (val s) (in_tuple (xi s))). Proof. move=> /andP[] /andP[] S0 SI /eqP SU xisort. -have {}xisort (s : S) x : path.sorted <%O [seq (f : {SAfun F^n -> F^1}) x 0 0 | f <- xi s]. +have {}xisort (s : S) x : + path.sorted <%O [seq (f : {SAfun F^n -> F^1}) x 0 0 | f <- xi s]. apply/path.pairwise_sorted/(pairwiseP 0) => i j. rewrite !inE size_map => ilt jlt ij. move: (xisort s) => /(lt_sorted_ltn_nth (0 : SAfunltType n))/(_ i j). @@ -3357,7 +3466,9 @@ apply/andP; split; last first. move=> /imfsetP [u] /= + ->. rewrite -SAset_fiber_partition_of_graphs => /mapP[v] vxi ->. rewrite inSAset_fiber hsubmxK => xv. - have vS: v :&: \val s :*: SAsetT F 1 \in \bigcup_(s | true) partition_of_graphs_above (val s) (xi s). + have vS: + v :&: \val s :*: SAsetT F 1 + \in \bigcup_(s | true) partition_of_graphs_above (val s) (xi s). apply/bigfcupP; exists s; first by rewrite mem_index_enum. by apply/imfsetP; exists v. apply/hasP; exists [` vS ] => /=; first exact/mem_index_enum. @@ -3371,10 +3482,13 @@ apply/andP; split. have: SAset_fiber (SAset0 F (n + 1)) x = SAset0 F _. by rewrite /SAset_fiber SApreimset0. rewrite i0 /SAset_fiber SApreimsetI -/(SAset_fiber _ _). - have ->: SApreimset (SAjoin (SAfun_const 1 x) (SAid 1)) (fsval s :*: SAsetT F 1) = SAsetT F _. + have ->: + SApreimset (SAjoin (SAfun_const 1 x) (SAid 1)) (fsval s :*: SAsetT F 1) + = SAsetT F _. apply/eqP/SAsetP => y; rewrite inSApreimset SAjoinE SAfun_constE inSAsetX. by rewrite row_mxKl xs !inSAsetT. - rewrite SAsetIT -[LHS](@nth_map _ _ _ (SAset0 F _) (fun s => SAset_fiber s x)); + rewrite SAsetIT. + rewrite -[LHS](@nth_map _ _ _ (SAset0 F _) (fun s => SAset_fiber s x)); last by rewrite size_map size_iota. rewrite SAset_fiber_partition_of_graphs => {}i0. move: (SAset_partition_of_pts (xisort s x)). @@ -3389,9 +3503,11 @@ move=>/(nthP (SAset0 F _)) [j] + <- ->. rewrite size_map size_iota => jlt; apply/implyP. move: jlt; have [<- jlt ij|st _ _] := eqVneq s t; last first. rewrite /SAset_disjoint -subset0; apply/SAset_subP => x. - rewrite !inSAsetI 2!inSAsetX => /andP[] /andP[_] /andP[xs] _ /andP[_] /andP[xt] _. + rewrite !inSAsetI 2!inSAsetX. + move=> /andP[] /andP[_] /andP[xs] _ /andP[_] /andP[xt] _. move: SI => /forallP/(_ s) /forallP /(_ t) /implyP. - rewrite (inj_eq val_inj) => /(_ st); rewrite /SAset_disjoint /subset0 => /eqP st0. + rewrite (inj_eq val_inj) => /(_ st). + rewrite /SAset_disjoint /subset0 => /eqP st0. suff: lsubmx x \in SAset0 F n by rewrite inSAset0. by rewrite -st0 inSAsetI xs. case/boolP: (i == j) => [/eqP ijE|{}ij]; first by rewrite ijE eqxx in ij. @@ -3402,7 +3518,8 @@ set xi' := [seq (f : {SAfun F^n -> F^1}) (lsubmx x) 0 0 | f <- xi s]. set T := [fset x | x in _] => /andP[] /andP[_] + _. have inT y : y \in T = (y \in partition_of_pts xi'). by apply/imfsetP/idP => [[] z zS -> //|yS]; exists y. -have nk k: (k < (size (xi s)).*2.+1)%N -> nth (SAset0 F _) (partition_of_pts xi') k \in T. +have nk k: (k < (size (xi s)).*2.+1)%N -> + nth (SAset0 F _) (partition_of_pts xi') k \in T. by rewrite inT => klt; apply/mem_nth; rewrite size_map size_iota size_map. move=> /forallP/(_ [` nk i ilt]) /forallP/(_ [` nk j jlt]) /implyP/=. rewrite nth_uniq ?size_map ?size_iota//; last first. @@ -3435,8 +3552,8 @@ have [<-|[y] yi _] := set0Vmem (nth (SAset0 F _) (partition_of_pts xi') i). by apply/mem_nth; rewrite 2!size_map size_iota. exists (row_mx x y); split; last by rewrite row_mxKl. move: yi; rewrite -SAset_fiber_partition_of_graphs. -rewrite (nth_map (SAset0 F _)) ?size_map ?size_iota// inSAset_fiber inSAsetI => ->. -by rewrite inSAsetX row_mxKl row_mxKr xs inSAsetT. +rewrite (nth_map (SAset0 F _)) ?size_map ?size_iota// inSAset_fiber inSAsetI. +by move=> ->; rewrite inSAsetX row_mxKl row_mxKr xs inSAsetT. Qed. Lemma SAset_partition_cast n m (S : {fset {SAset F^n}}) : @@ -3455,7 +3572,8 @@ Lemma SAset_formula (F : rcfType) (n : nat) (s : {SAset F^n}) : Proof. exists (qf_elim s); split; first exact/rform_qf_elim. split; first exact/qf_elim_qf. -apply/eqP/SAsetP => x; apply/rcf_satP/SAin_setP => [xs|/rcf_satP/qf_elim_holdsP//]. +apply/eqP/SAsetP => x. +apply/rcf_satP/SAin_setP => [xs|/rcf_satP/qf_elim_holdsP//]. exact/rcf_satP/qf_elim_holdsP. Qed. @@ -3463,7 +3581,8 @@ Lemma SAset_nf (F : rcfType) (n : nat) (s : {SAset F^n}) : {P : seq ({mpoly F[n]} * seq {mpoly F[n]}) | s = \big[@SAsetU F n/SAset0 F n]_(p <- P) (SApreimset (SAmpoly (fst p)) (SAset_seq [:: 0]) - :&: \big[@SAsetI F n/SAsetT F n]_(q <- (snd p)) SApreimset (SAmpoly q) (SAset_pos F))}. + :&: \big[@SAsetI F n/SAsetT F n]_(q <- (snd p)) + SApreimset (SAmpoly q) (SAset_pos F))}. Proof. pose has_nf (f : {SAset F^n}) := {P : seq ({mpoly F[n]} * seq {mpoly F[n]}) @@ -3494,9 +3613,9 @@ have IHI (f g : {SAset F^n}) : forall_ord1 !mxE mevalD !mevalXn paddr_eq0 ?sqr_ge0// !sqrf_eq0. move=> /andP[]/andP[] pf10 pg10 /allP pfg20. exists (pf, pg); first by apply/allpairsP => /=; exists (pf, pg). - rewrite !inSAsetI 2!inSApreimset !SAmpolyE !inSAset1 !rowPE !forall_ord1 !mxE - pf10 pg10/= !inSAset_bigcap. - by apply/andP; split; apply/allP => p pP /=; apply/pfg20; rewrite mem_cat pP// orbT. + rewrite !inSAsetI 2!inSApreimset !SAmpolyE !inSAset1 !rowPE !forall_ord1 !mxE. + rewrite pf10 pg10/= !inSAset_bigcap; apply/andP. + by split; apply/allP => p pP /=; apply/pfg20; rewrite mem_cat pP// orbT. have IHIs (I : Type) (r : seq I) (f : I -> {SAset F^n}) : (forall i, has_nf (f i)) -> has_nf (\big[@SAsetI F n/SAsetT F n]_(i <- r) f i). @@ -3524,8 +3643,8 @@ have IHC (f : {SAset F^n}) : has_nf f -> has_nf (~: f). forall_ord1 !mxE/=. exists (0, [:: -p]). by rewrite mem_cat; apply/orP; right; apply/mapP; exists p. - by rewrite inSAsetI/= big_seq1 2!inSApreimset inSAset1 inSAset_pos !SAmpolyE - rowPE forall_ord1 !mxE meval0 eqxx mevalN oppr_gt0. + rewrite inSAsetI/= big_seq1 2!inSApreimset inSAset1 inSAset_pos !SAmpolyE. + by rewrite rowPE forall_ord1 !mxE meval0 eqxx mevalN oppr_gt0. - by left; rewrite eq_sym (lt_eqF pf10). - by left; rewrite (lt_eqF pf10). rewrite mem_cat => /= /orP [|] /mapP [q]/= qf ->. @@ -3549,12 +3668,12 @@ case: (SAset_formula s) => + [+][+] -> {s}; elim=> //=. - move=> t u /andP[] rt ru _. exists [:: (mpoly_rterm n (to_rterm (GRing.Add t (GRing.Opp u))), [::])]. rewrite big_seq1/= big_nil SAsetIT; apply/eqP/SAsetP => x. - rewrite inSApreimset SAmpolyE inSAset1 rowPE forall_ord1 !mxE mevalB subr_eq0. + rewrite inSApreimset SAmpolyE [RHS]inSAset1 rowPE forall_ord1 !mxE mevalB subr_eq0. by rewrite !meval_mpoly_rterm !evalE !eval_rterm//; apply/SAin_setP/eqP. - move=> t u /andP[] rt ru _. exists [:: (0, [:: mpoly_rterm n (to_rterm (GRing.Add u (GRing.Opp t)))])]. rewrite big_seq1/= big_seq1; apply/eqP/SAsetP => x. - rewrite inSAsetI 2!inSApreimset !SAmpolyE inSAset1 rowPE forall_ord1. + rewrite inSAsetI 2!inSApreimset !SAmpolyE [in RHS]inSAset1 rowPE forall_ord1. rewrite inSAset_pos !mxE meval0 eqxx/= mevalB subr_gt0 !meval_mpoly_rterm. by rewrite !evalE !eval_rterm//; apply/SAin_setP/idP. - move=> t u /andP[] rt ru _. @@ -3562,7 +3681,7 @@ case: (SAset_formula s) => + [+][+] -> {s}; elim=> //=. exists [:: (mpoly_rterm n (to_rterm v), [::]); (0, [:: mpoly_rterm n (to_rterm v)])]. rewrite big_cons big_nil big_seq1/= big_seq1 SAsetIT; apply/eqP/SAsetP => x. - rewrite inSAsetU inSAsetI 3!inSApreimset !SAmpolyE 2!inSAset1 !rowPE. + rewrite inSAsetU inSAsetI 3!inSApreimset !SAmpolyE 2![in RHS]inSAset1 !rowPE. rewrite !forall_ord1 inSAset_pos !mxE meval0 eqxx/= mevalB subr_gt0. rewrite subr_eq0 eq_sym -le_eqVlt !meval_mpoly_rterm !evalE !eval_rterm//. exact/SAin_setP/idP. @@ -3615,7 +3734,8 @@ have has_nfI (T : Type) (r : seq T) f : apply/eqP/SAsetP => x. by rewrite inSAsetT inSAset_itv in_itv. case: (SAset_nf s) => + -> => nf. -have mnm0 (m : 'X_{1.. 1}): [multinom [tuple m (widen_ord (leqnSn 0) i) | i < 0]] == 0%MM. +have mnm0 (m : 'X_{1.. 1}): + [multinom [tuple m (widen_ord (leqnSn 0) i) | i < 0]] == 0%MM. by apply/(@eqP (_.-tuple _))/eq_from_tnth; case. have coeffp1 (p : {mpoly F[1]}) (m : 'X_{1.. 1}) : p@_m = (map_poly (mcoeff 0) (muni p))`_(m ord0). @@ -3648,7 +3768,8 @@ apply/has_nfU => -[/=] p r; apply/has_nfI2. exists [:: `]-oo, +oo[]. apply/eqP/SAsetP => x. rewrite inSApreimset inSAset1 SAmpolyE rowPE forall_ord1 !mxE. - by rewrite meval0 big_cons big_nil SAsetUC SAset0U inSAset_itv in_itv/= eqxx. + rewrite meval0 big_cons big_nil SAsetUC SAset0U inSAset_itv in_itv/=. + by rewrite eqxx. exists [seq `[x, x] | x <- rootsR (map_poly (mcoeff 0) (muni p))]. apply/eqP/SAsetP => x. rewrite inSApreimset inSAset1 SAmpolyE rowPE forall_ord1 !mxE. @@ -3682,7 +3803,8 @@ exists ( cons `]-oo, x[ else id) ((if 0 < lead_coef q then cons `]last x r, +oo[ else id) - [seq `](x :: r)`_i, r`_i[ | i <- iota 0 (size r) & sgz q.[((x :: r)`_i + r`_i) / 2] == 1])). + [seq `](x :: r)`_i, r`_i[ | + i <- iota 0 (size r) & sgz q.[((x :: r)`_i + r`_i) / 2] == 1])). apply/eqP/SAsetP => y. rewrite inSApreimset inSAset_pos SAmpolyE mxE inSAset_bigcup. move: rE; have [->|q0 rE] := eqVneq q 0. @@ -3690,7 +3812,9 @@ move: rE; have [->|q0 rE] := eqVneq q 0. rewrite mevalp1 -/q. have /(pairwiseP 0) xr_sort: pairwise <%O (x :: r). by rewrite -lt_sorted_pairwise -rE; exact/sorted_roots. -have xr_sort': {in gtn (size (x :: r)) &, {homo nth 0 (x :: r) : i j / (i <= j)%N >-> i <= j}}. +have xr_sort': + {in gtn (size (x :: r)) &, + {homo nth 0 (x :: r) : i j / (i <= j)%N >-> i <= j}}. move=> i j ilt jlt. rewrite leq_eqVlt => /orP[/eqP ->|ij]; first exact/lexx. exact/ltW/xr_sort. @@ -3725,7 +3849,8 @@ case /boolP: (root q (y 0 0)) => qy0 /=. apply/idP/idP => [q0'|]; last first. move=> /hasP[] I. have [->|Ix] := eqVneq I `]-oo, x[. - case/boolP: ((0 < lead_coef q) (+) odd (\sum_(y0 <- rootsR q) \mu_y0 q)) => [q0'|] _; last first. + case/boolP: ((0 < lead_coef q) (+) odd (\sum_(y0 <- rootsR q) \mu_y0 q)) + => [q0'|] _; last first. case: (0 < lead_coef q); last by move=> /mapP[]. by rewrite in_cons/= => /mapP[]. rewrite inSAset_itv in_itv/= => yx. @@ -3774,12 +3899,14 @@ apply/idP/idP => [q0'|]; last first. rewrite leNgt => /negP; apply. apply/(xr_sort i i.+1) => //. by rewrite inE/= ltnS; apply/ltnW. - move: (xr_sort' j i); rewrite !inE/= [(i < _)%N]ltnS => /(_ jq (ltnW ilt) ij). + move: (xr_sort' j i). + rewrite !inE/= [(i < _)%N]ltnS => /(_ jq (ltnW ilt) ij). rewrite ji ler_pdivrMr// mulr_natr mulr2n -subr_ge0 addrKA subr_ge0. rewrite leNgt => /negP; apply. apply/(xr_sort i i.+1) => //. by rewrite inE/= ltnS; apply/ltnW. - rewrite big_mkcond big_seq -big_mkcondr [in RHS]big_mkcond [in RHS]big_seq -big_mkcondr/=. + rewrite big_mkcond big_seq -big_mkcondr [in RHS]big_mkcond [in RHS]big_seq. + rewrite -big_mkcondr/=. apply/eq_bigl => z; case/boolP: (z \in rootsR q) => // /(nthP 0)[] j. rewrite rE => jlt <- /=. rewrite ltr_pdivrMr// mulr_natr mulr2n. @@ -3787,7 +3914,8 @@ apply/idP/idP => [q0'|]; last first. move: (xr_sort' i.+1 j); rewrite !inE/= ltnS => /(_ ilt jlt ij) rij. rewrite (lt_le_trans yi rij); apply/ltr_leD => //. exact/(lt_le_trans _ rij)/(lt_trans iy). - move: (xr_sort' j i); rewrite !inE/= [(i < _)%N]ltnS => /(_ jlt (ltnW ilt) ij) rji. + move: (xr_sort' j i). + rewrite !inE/= [(i < _)%N]ltnS => /(_ jlt (ltnW ilt) ij) rji. rewrite [RHS]ltNge (ltW (le_lt_trans rji iy))/=; apply/negP/negP. rewrite -leNgt; apply/lerD => //. exact/(le_trans rji)/ltW/(lt_trans iy). @@ -3852,7 +3980,8 @@ congr (_ < _ * Posz (_ _) * _ ^+ _). rewrite leNgt => /negP; apply. apply/(xr_sort j j.+1) => //. by rewrite inE/= ltnS; apply/ltnW. -rewrite big_mkcond big_seq -big_mkcondr [in RHS]big_mkcond [in RHS]big_seq -big_mkcondr/=. +rewrite big_mkcond big_seq -big_mkcondr [in RHS]big_mkcond [in RHS]big_seq. +rewrite -big_mkcondr/=. apply/eq_bigl => z; case/boolP: (z \in rootsR q) => // /(nthP 0)[] k. rewrite rE => klt <- /=. rewrite ltr_pdivrMr// mulr_natr mulr2n. @@ -3864,7 +3993,8 @@ case: (ltnP j k) => jk. by move: (mem_nth 0 klt); rewrite -jE -rE in_rootsR (negPf qy0) andbF. rewrite yk; apply/esym/ltr_leD; first by apply/xr_sort => //; rewrite inE. by rewrite -[r`_j]/((x :: r)`_j.+1); apply/xr_sort'. -move: (xr_sort' k j); rewrite !inE/= [(j < _)%N]ltnS => /(_ klt (ltnW jr) jk) rjk. +move: (xr_sort' k j). +rewrite !inE/= [(j < _)%N]ltnS => /(_ klt (ltnW jr) jk) rjk. rewrite ltNge (ltW (le_lt_trans rjk jy))/=; apply/esym/negP/negP. rewrite -leNgt; apply/lerD => //. apply/(le_trans rjk)/ltW/(lt_trans jy). @@ -3880,10 +4010,10 @@ Variables (F : rcfType) (n : nat). Implicit Types (s : {SAset F^n}). Definition SAsetUB (s : {SAset F^1}) : {SAset F^1} := -[set | 'forall 'X_1, (subst_formula [:: 1%N] s ==> ('X_1 <=% 'X_0))%oT]. + [set | 'forall 'X_1, (subst_formula [:: 1%N] s ==> ('X_1 <=% 'X_0))%oT]. Lemma inSAsetUB (s : {SAset F^1}) (x : 'rV[F]_1) : -reflect (forall y, y \in s -> y ord0 ord0 <= x ord0 ord0) (x \in SAsetUB s). + reflect (forall y, y \in s -> y ord0 ord0 <= x ord0 ord0) (x \in SAsetUB s). Proof. apply/(iffP (SAin_setP _ _)) => /= [+ y ys|yx y]. move=> /(_ (y ord0 ord0)); rewrite holds_subst/= !nth_set_nth/= enum_ordSl/=. @@ -3895,7 +4025,7 @@ by rewrite enum_ordSl enum_ord0/= mxE; apply/rcf_satP. Qed. Lemma inSAsetUBC (s : {SAset F^1}) (x : 'rV[F]_1) : -reflect (exists y, y \in s /\ x ord0 ord0 < y ord0 ord0) (x \in ~: SAsetUB s). + reflect (exists y, y \in s /\ x ord0 ord0 < y ord0 ord0) (x \in ~: SAsetUB s). Proof. rewrite SAsetC_comprehension. apply/(iffP (SAin_setP _ _)) => [/n_forall_formula /= [y]|[y][ys] xy]. @@ -3925,7 +4055,7 @@ by rewrite !mxE -subr_ge0 opprD addrA subrr add0r leNgt oppr_lt0 ltr01. Qed. Lemma SAsetUBU (s t : {SAset F^1}) : -SAsetUB (s :|: t) = SAsetUB s :&: SAsetUB t. + SAsetUB (s :|: t) = SAsetUB s :&: SAsetUB t. Proof. apply/eqP/SAsetP => x; rewrite inSAsetI. apply/inSAsetUB/andP => [xst|[] /inSAsetUB xs/inSAsetUB xt y]; last first. @@ -3934,8 +4064,8 @@ by split; apply/inSAsetUB => y yst; apply/xst; rewrite inSAsetU yst// orbT. Qed. Lemma SAsetUBbigcup (I : Type) (r : seq I) (P : pred I) (f : I -> {SAset F^1}) : -SAsetUB (\big[@SAsetU F 1/SAset0 F 1]_(i <- r | P i) f i) - = \big[@SAsetI F 1/SAsetT F 1]_(i <- r | P i) (SAsetUB (f i)). + SAsetUB (\big[@SAsetU F 1/SAset0 F 1]_(i <- r | P i) f i) + = \big[@SAsetI F 1/SAsetT F 1]_(i <- r | P i) (SAsetUB (f i)). Proof. elim: r => [|i r IHr]; first by rewrite !big_nil SAsetUB0. by rewrite !big_cons; case: (P i) => //; rewrite SAsetUBU IHr. diff --git a/subresultant.v b/subresultant.v index d8b1c6f..be6c291 100644 --- a/subresultant.v +++ b/subresultant.v @@ -1,11 +1,7 @@ -From mathcomp -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq order div fintype tuple. -From mathcomp -Require Import finfun bigop fingroup perm ssralg zmodp matrix mxalgebra interval. -From mathcomp -Require Import binomial poly polydiv mxpoly ssrnum. -From mathcomp -Require Import ssrint. +From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq order. +From mathcomp Require Import div fintype tuple finfun bigop fingroup perm. +From mathcomp Require Import ssralg zmodp matrix mxalgebra interval binomial. +From mathcomp Require Import ssrint poly polydiv mxpoly ssrnum. From mathcomp Require Import polyorder polyrcf qe_rcf_th. Require Import extra_ssr auxresults. @@ -393,8 +389,9 @@ Qed. Fact gt_size_gcd (R : idomainType) (p q u v : {poly R}) j : p != 0 -> q != 0 -> u != 0 -> - (j < size (gcdp p q))%N -> (j <= (size q).-1)%N -> (size u <= (size q).-1 - j)%N -> - (size (u * p + v * q)%R <= j)%N -> (j < (size (gcdp p q)).-1)%N. + (j < size (gcdp p q))%N -> (j <= (size q).-1)%N -> + (size u <= (size q).-1 - j)%N -> (size (u * p + v * q)%R <= j)%N -> + (j < (size (gcdp p q)).-1)%N. Proof. move=> p0 q0 u0 gt_sg_j ge_sq_j ge_sqmj_u. set l := _ * _ + _ * _ => sl; have /eqP : l = 0. @@ -565,7 +562,8 @@ Qed. Lemma subresultant_mod (R : idomainType) j (p q : {poly R}) (c := (-1) ^+ 'C((size p).-1 - (size q).-1, 2) * lead_coef q ^+ ((size p).-1 - (size (p %% q)).-1)) : - p != 0 -> q != 0 -> ((size q).-1 <= (size p).-1)%N -> (j <= (size (p %% q)%R).-1)%N -> + p != 0 -> q != 0 -> ((size q).-1 <= (size p).-1)%N -> + (j <= (size (p %% q)%R).-1)%N -> subresultant j (lead_coef q ^+ (scalp p q) *: p) q = c * subresultant j q (- (p %% q)). Proof. @@ -615,7 +613,8 @@ by rewrite addbC odd_bin2B ?leq_sub // addKb addnC. Qed. Lemma subresultant_gt_mod (R : idomainType) j (p q : {poly R}) : - p != 0 -> q != 0 -> ((size q).-1 <= (size p).-1)%N -> (size (p %% q)%R <= j < (size q)%R.-1)%N -> + p != 0 -> q != 0 -> ((size q).-1 <= (size p).-1)%N -> + (size (p %% q)%R <= j < (size q)%R.-1)%N -> subresultant j p q = 0. Proof. move=> p_neq0 q_neq0 le_pq /andP[le_rj] le_jq. @@ -632,8 +631,8 @@ Qed. Lemma subresultantp0 (R : idomainType) j (p : {poly R}) : (j < (size p).-1)%N -> subresultant j p 0 = 0. Proof. -rewrite ltnNge -subn_eq0 => /negPf sp; apply/eqP; rewrite subresultant_eq0 size_poly0/= sub0n. -rewrite /SylvesterHabicht_mx band0. +rewrite ltnNge -subn_eq0 => /negPf sp; apply/eqP. +rewrite subresultant_eq0 size_poly0/= sub0n /SylvesterHabicht_mx band0. case: ((size p).-1 - j)%N sp => //= n _. apply/det0P; exists (row_mx 0 (\row_i (i == ord_max)%:R)). apply/eqP => /rowP /(_ (unsplit (inr ord_max))). @@ -649,7 +648,8 @@ move=> jq; apply/eqP; rewrite subresultantC mulf_eq0; apply/orP; right. exact/eqP/subresultantp0. Qed. -Lemma subresultant_map_poly (A B : ringType) i (p q : {poly A}) (f : {rmorphism A -> B}) : +Lemma subresultant_map_poly (A B : ringType) i + (p q : {poly A}) (f : {rmorphism A -> B}) : f (lead_coef p) != 0 -> f (lead_coef q) != 0 -> subresultant i (map_poly f p) (map_poly f q) = f (subresultant i p q). Proof. @@ -700,9 +700,10 @@ have trig: is_trig_mx rewrite /subresultant subnn det_trig//. under eq_bigr => i _. suff ->: rsubmx - (block_mx (perm_mx extra_ssr.perm_rev) 0 0 1%:M *m - SylvesterHabicht_mx 0 ((size q).-1 - (size p).-1) - ((size p).-1 + (0 + ((size q).-1 - (size p).-1))) q p) i i = lead_coef p. + (block_mx (perm_mx extra_ssr.perm_rev) 0 0 1%:M *m + SylvesterHabicht_mx 0 ((size q).-1 - (size p).-1) + ((size p).-1 + (0 + ((size q).-1 - (size p).-1))) q p) i i + = lead_coef p. over. rewrite !mxE; under eq_bigr => /= k _. suff ->: block_mx (perm_mx extra_ssr.perm_rev) 0 0 1%:M i k * @@ -732,10 +733,16 @@ Import Num.Theory Order.POrderTheory Pdiv.Field. Theorem pmv_subresultant (R : rcfType) (p q : {poly R}) : (size q < size p)%N -> - pmv ([seq (lead_coef p) *+ (i.+1 == size p) | i <- rev (iota (size q) (size p - size q))] ++ [seq subresultant i p q | i <- rev (iota 0 (size q))]) = cindexR q p. + pmv ( + [seq (lead_coef p) *+ (i.+1 == size p) | + i <- rev (iota (size q) (size p - size q))] ++ + [seq subresultant i p q | i <- rev (iota 0 (size q))]) + = cindexR q p. Proof. move sq: (size q) => n; move: p q sq. -apply/(@Wf_nat.lt_wf_ind n (fun n => forall p q : {poly R}, size q = n -> (n < size p)%N -> _)) => {}n IHn p q sq sp. +apply/(@Wf_nat.lt_wf_ind n + (fun n => forall p q : {poly R}, size q = n -> (n < size p)%N -> _)). +move=> {}n IHn p q sq sp. case/boolP: (q == 0) sq sp => [/eqP ->|q0 sq sp]. rewrite size_poly0 => <- {n IHn} sp /=; rewrite cindexR0p cats0 subn0. case: (size p) => // n. @@ -745,17 +752,21 @@ case/boolP: (q == 0) sq sp => [/eqP ->|q0 sq sp]. have p0: p != 0. by apply/eqP => p0; move: sp; rewrite p0 size_poly0. have sr : (size (p %% q)%R < n)%N by rewrite -sq; apply/ltn_modpN0. -have subr: pmv [seq subresultant i p q | i <- rev (iota 0 (size (p %% q)))] = pmv [seq subresultant i q (- (p %% q)) | i <- rev (iota 0 (size (p %% q)))]. +have subr: pmv [seq subresultant i p q | i <- rev (iota 0 (size (p %% q)))] + = pmv [seq subresultant i q (- (p %% q)) | + i <- rev (iota 0 (size (p %% q)))]. rewrite -[RHS](pmvZ (a:=(-1) ^+ 'C((size p).-1 - (size q).-1, 2) * lead_coef q ^+ ((size p).-1 - (size (p %% q)).-1))); last first. rewrite mulf_neq0//; first by apply/expf_neq0; rewrite oppr_eq0 oner_neq0. by apply/expf_neq0; rewrite lead_coef_eq0. - rewrite -map_comp/= [in RHS](eq_map_seq (g:=fun i => subresultant i ((lead_coef q ^+ scalp p q) *: p) q)); last first. - move=> i /=; rewrite mem_rev mem_iota/= add0n => ilt. - apply/esym/subresultant_mod => //. - - by rewrite -!subn1 sq; apply/leq_sub2r/ltnW. - by rewrite -subn1 leq_leq_subRL ?add1n//; apply/(leq_trans _ ilt). - by rewrite scalpE expr0 scale1r. + rewrite -map_comp/=. + rewrite [in RHS](eq_map_seq + (g:=fun i => subresultant i ((lead_coef q ^+ scalp p q) *: p) q)). + by rewrite scalpE expr0 scale1r. + move=> i /=; rewrite mem_rev mem_iota/= add0n => ilt. + apply/esym/subresultant_mod => //. + by rewrite -!subn1 sq; apply/leq_sub2r/ltnW. + by rewrite -subn1 leq_leq_subRL ?add1n//; apply/(leq_trans _ ilt). have subq: pmv [seq subresultant i p q | i <- rev (iota 0 n)] = pmv @@ -770,7 +781,8 @@ have subq: move=> x /mapP [i] + ->; rewrite mem_rev mem_iota => /andP[pqi ipq]. apply/eqP/subresultant_gt_mod => //; rewrite sq. by rewrite -!subn1; apply/leq_sub2r/ltnW. - by rewrite pqi/=; move: ipq; rewrite -ltnS -addnS -mE subnKC ?ltn_predRL// ltnW. + rewrite pqi/=; move: ipq; rewrite -ltnS -addnS -mE subnKC ?ltn_predRL//. + by rewrite ltnW. rewrite -addnS -mE subnKC; last exact/ltnW. have /ltn_eqF nSn: (n < n.+1)%N by []. rewrite eqxx mulr1n. @@ -780,15 +792,20 @@ have subq: rewrite -[RHS]/(pmv (_ :: _)) pmv_cat00// => x /mapP [i]. rewrite mem_rev mem_iota/= add0n mE eqSS => /ltn_eqF -> ->. by rewrite mulr0n. - move sr: (size (p %% q)) sr subr mE pq0 => k; case: k sr => [|k] sr kn subr mE pq0. + move sr: (size (p %% q)) sr subr mE pq0 => k. + case: k sr => [|k] sr kn subr mE pq0. by move/eqP: sr; rewrite size_poly_eq0 (negPf r0). move: subr; rewrite iotanS add0n rev_rcons/= => subr. - have kqr : subresultant k q (- (p %% q)) = (-1) ^+ 'C(m.+1, 2) * lead_coef (- (p %% q)) ^+ m.+1. + have kqr : subresultant k q (- (p %% q)) + = (-1) ^+ 'C(m.+1, 2) * lead_coef (- (p %% q)) ^+ m.+1. rewrite subresultantE size_opp sr -pred_Sn subnn sq. rewrite /SylvesterHabicht_mx col_flat_mx det_rsub_band; last first. by rewrite size_opp sr -pred_Sn. by rewrite add0n -predn_sub -subnS mE bin0n/= addn0. - have kpq: subresultant k p q = (-1) ^+ 'C((size p).-1 - n.-1, 2) * lead_coef q ^+ ((size p).-1 - k) * subresultant k q (- (p %% q)). + have kpq: subresultant k p q + = (-1) ^+ 'C((size p).-1 - n.-1, 2) * + lead_coef q ^+ ((size p).-1 - k) * + subresultant k q (- (p %% q)). have {1}->: p = lead_coef q ^+ scalp p q *: p. by rewrite scalpE expr0 scale1r. rewrite subresultant_mod//. @@ -796,13 +813,15 @@ have subq: - by rewrite sq -!subn1; apply/leq_sub2r/ltnW. - by rewrite sr -pred_Sn. rewrite -[LHS]/(pmv (_ :: _)) pmv_cat0s//; last first. - by rewrite kpq kqr !mulf_eq0 !signr_eq0 !expf_eq0/= lead_coefN oppr_eq0 !lead_coef_eq0 (negPf q0) andbF. + rewrite kpq kqr !mulf_eq0 !signr_eq0 !expf_eq0/= lead_coefN oppr_eq0. + by rewrite !lead_coef_eq0 (negPf q0) andbF. rewrite size_map size_rev size_iota/= {}subr. rewrite -[in RHS]/(pmv (_ :: _)) pmv_cat0s//; first last. - move=> x /mapP[i]; rewrite mem_rev mem_iota => /andP[] ki + ->. rewrite -ltnS -addnS -mE subnKC// ?(ltnW kn)// => ilt. by rewrite (ltn_eqF ilt) !mulr0n. - - by rewrite kqr mulf_eq0 signr_eq0 expf_eq0 lead_coefN oppr_eq0 lead_coef_eq0. + - rewrite kqr mulf_eq0 signr_eq0 expf_eq0 lead_coefN oppr_eq0. + by rewrite lead_coef_eq0. rewrite size_map size_rev size_iota/=; congr (_ + _). rewrite -!mulrnAr; congr (_ * _). rewrite [(_ + _)%N]pred_Sn -addnS -mE subnKC ?(ltnW kn)//.