From e0d682ec25282a348d35c5b169abafec48555690 Mon Sep 17 00:00:00 2001 From: Stephane Glondu Date: Mon, 20 Aug 2012 18:27:01 +0200 Subject: Imported Upstream version 8.4dfsg --- plugins/setoid_ring/Ring_polynom.v | 1310 ++++++++++++++---------------------- 1 file changed, 507 insertions(+), 803 deletions(-) (limited to 'plugins/setoid_ring/Ring_polynom.v') diff --git a/plugins/setoid_ring/Ring_polynom.v b/plugins/setoid_ring/Ring_polynom.v index b722a31b..b23ba352 100644 --- a/plugins/setoid_ring/Ring_polynom.v +++ b/plugins/setoid_ring/Ring_polynom.v @@ -1,20 +1,16 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* R -> Prop. (* Ring properties *) - Variable Rsth : Setoid_Theory R req. + Variable Rsth : Equivalence req. Variable Reqe : ring_eq_ext radd rmul ropp req. Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req. @@ -37,7 +33,7 @@ Section MakeRingPol. Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi. - (* Power coefficients *) + (* Power coefficients *) Variable Cpow : Type. Variable Cp_phi : N -> Cpow. Variable rpow : R -> Cpow -> R. @@ -50,26 +46,47 @@ Section MakeRingPol. (* R notations *) Notation "0" := rO. Notation "1" := rI. - Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y). - Notation "x - y " := (rsub x y). Notation "- x" := (ropp x). - Notation "x == y" := (req x y). + Infix "+" := radd. Infix "*" := rmul. + Infix "-" := rsub. Notation "- x" := (ropp x). + Infix "==" := req. + Infix "^" := (pow_pos rmul). (* C notations *) - Notation "x +! y" := (cadd x y). Notation "x *! y " := (cmul x y). - Notation "x -! y " := (csub x y). Notation "-! x" := (copp x). - Notation " x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x). + Infix "+!" := cadd. Infix "*!" := cmul. + Infix "-! " := csub. Notation "-! x" := (copp x). + Infix "?=!" := ceqb. Notation "[ x ]" := (phi x). (* Useful tactics *) - Add Setoid R req Rsth as R_set1. - Ltac rrefl := gen_reflexivity Rsth. - Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed. - Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed. - Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed. - Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed. + Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed. + Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed. + Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed. + Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed. Ltac rsimpl := gen_srewrite Rsth Reqe ARth. + Ltac add_push := gen_add_push radd Rsth Reqe ARth. Ltac mul_push := gen_mul_push rmul Rsth Reqe ARth. + Ltac add_permut_rec t := + match t with + | ?x + ?y => add_permut_rec y || add_permut_rec x + | _ => add_push t; apply (Radd_ext Reqe); [|reflexivity] + end. + + Ltac add_permut := + repeat (reflexivity || + match goal with |- ?t == _ => add_permut_rec t end). + + Ltac mul_permut_rec t := + match t with + | ?x * ?y => mul_permut_rec y || mul_permut_rec x + | _ => mul_push t; apply (Rmul_ext Reqe); [|reflexivity] + end. + + Ltac mul_permut := + repeat (reflexivity || + match goal with |- ?t == _ => mul_permut_rec t end). + + (* Definition of multivariable polynomials with coefficients in C : Type [Pol] represents [X1 ... Xn]. The representation is Horner's where a [n] variable polynomial @@ -116,19 +133,19 @@ Section MakeRingPol. | _, _ => false end. - Notation " P ?== P' " := (Peq P P'). + Infix "?==" := Peq. Definition mkPinj j P := match P with | Pc _ => P - | Pinj j' Q => Pinj ((j + j'):positive) Q + | Pinj j' Q => Pinj (j + j') Q | _ => Pinj j P end. Definition mkPinj_pred j P:= match j with | xH => P - | xO j => Pinj (Pdouble_minus_one j) P + | xO j => Pinj (Pos.pred_double j) P | xI j => Pinj (xO j) P end. @@ -156,14 +173,14 @@ Section MakeRingPol. (** Addition et subtraction *) - Fixpoint PaddC (P:Pol) (c:C) {struct P} : Pol := + Fixpoint PaddC (P:Pol) (c:C) : Pol := match P with | Pc c1 => Pc (c1 +! c) | Pinj j Q => Pinj j (PaddC Q c) | PX P i Q => PX P i (PaddC Q c) end. - Fixpoint PsubC (P:Pol) (c:C) {struct P} : Pol := + Fixpoint PsubC (P:Pol) (c:C) : Pol := match P with | Pc c1 => Pc (c1 -! c) | Pinj j Q => Pinj j (PsubC Q c) @@ -175,11 +192,11 @@ Section MakeRingPol. Variable Pop : Pol -> Pol -> Pol. Variable Q : Pol. - Fixpoint PaddI (j:positive) (P:Pol){struct P} : Pol := + Fixpoint PaddI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PaddC Q c) | Pinj j' Q' => - match ZPminus j' j with + match Z.pos_sub j' j with | Zpos k => mkPinj j (Pop (Pinj k Q') Q) | Z0 => mkPinj j (Pop Q' Q) | Zneg k => mkPinj j' (PaddI k Q') @@ -187,16 +204,16 @@ Section MakeRingPol. | PX P i Q' => match j with | xH => PX P i (Pop Q' Q) - | xO j => PX P i (PaddI (Pdouble_minus_one j) Q') + | xO j => PX P i (PaddI (Pos.pred_double j) Q') | xI j => PX P i (PaddI (xO j) Q') end end. - Fixpoint PsubI (j:positive) (P:Pol){struct P} : Pol := + Fixpoint PsubI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PaddC (--Q) c) | Pinj j' Q' => - match ZPminus j' j with + match Z.pos_sub j' j with | Zpos k => mkPinj j (Pop (Pinj k Q') Q) | Z0 => mkPinj j (Pop Q' Q) | Zneg k => mkPinj j' (PsubI k Q') @@ -204,41 +221,41 @@ Section MakeRingPol. | PX P i Q' => match j with | xH => PX P i (Pop Q' Q) - | xO j => PX P i (PsubI (Pdouble_minus_one j) Q') + | xO j => PX P i (PsubI (Pos.pred_double j) Q') | xI j => PX P i (PsubI (xO j) Q') end end. Variable P' : Pol. - Fixpoint PaddX (i':positive) (P:Pol) {struct P} : Pol := + Fixpoint PaddX (i':positive) (P:Pol) : Pol := match P with | Pc c => PX P' i' P | Pinj j Q' => match j with | xH => PX P' i' Q' - | xO j => PX P' i' (Pinj (Pdouble_minus_one j) Q') + | xO j => PX P' i' (Pinj (Pos.pred_double j) Q') | xI j => PX P' i' (Pinj (xO j) Q') end | PX P i Q' => - match ZPminus i i' with + match Z.pos_sub i i' with | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' | Z0 => mkPX (Pop P P') i Q' | Zneg k => mkPX (PaddX k P) i Q' end end. - Fixpoint PsubX (i':positive) (P:Pol) {struct P} : Pol := + Fixpoint PsubX (i':positive) (P:Pol) : Pol := match P with | Pc c => PX (--P') i' P | Pinj j Q' => match j with | xH => PX (--P') i' Q' - | xO j => PX (--P') i' (Pinj (Pdouble_minus_one j) Q') + | xO j => PX (--P') i' (Pinj (Pos.pred_double j) Q') | xI j => PX (--P') i' (Pinj (xO j) Q') end | PX P i Q' => - match ZPminus i i' with + match Z.pos_sub i i' with | Zpos k => mkPX (Pop (PX P k P0) P') i' Q' | Z0 => mkPX (Pop P P') i Q' | Zneg k => mkPX (PsubX k P) i Q' @@ -258,18 +275,18 @@ Section MakeRingPol. | Pinj j Q => match j with | xH => PX P' i' (Padd Q Q') - | xO j => PX P' i' (Padd (Pinj (Pdouble_minus_one j) Q) Q') + | xO j => PX P' i' (Padd (Pinj (Pos.pred_double j) Q) Q') | xI j => PX P' i' (Padd (Pinj (xO j) Q) Q') end | PX P i Q => - match ZPminus i i' with + match Z.pos_sub i i' with | Zpos k => mkPX (Padd (PX P k P0) P') i' (Padd Q Q') | Z0 => mkPX (Padd P P') i (Padd Q Q') | Zneg k => mkPX (PaddX Padd P' k P) i (Padd Q Q') end end end. - Notation "P ++ P'" := (Padd P P'). + Infix "++" := Padd. Fixpoint Psub (P P': Pol) {struct P'} : Pol := match P' with @@ -281,22 +298,22 @@ Section MakeRingPol. | Pinj j Q => match j with | xH => PX (--P') i' (Psub Q Q') - | xO j => PX (--P') i' (Psub (Pinj (Pdouble_minus_one j) Q) Q') + | xO j => PX (--P') i' (Psub (Pinj (Pos.pred_double j) Q) Q') | xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q') end | PX P i Q => - match ZPminus i i' with + match Z.pos_sub i i' with | Zpos k => mkPX (Psub (PX P k P0) P') i' (Psub Q Q') | Z0 => mkPX (Psub P P') i (Psub Q Q') | Zneg k => mkPX (PsubX Psub P' k P) i (Psub Q Q') end end end. - Notation "P -- P'" := (Psub P P'). + Infix "--" := Psub. (** Multiplication *) - Fixpoint PmulC_aux (P:Pol) (c:C) {struct P} : Pol := + Fixpoint PmulC_aux (P:Pol) (c:C) : Pol := match P with | Pc c' => Pc (c' *! c) | Pinj j Q => mkPinj j (PmulC_aux Q c) @@ -310,11 +327,11 @@ Section MakeRingPol. Section PmulI. Variable Pmul : Pol -> Pol -> Pol. Variable Q : Pol. - Fixpoint PmulI (j:positive) (P:Pol) {struct P} : Pol := + Fixpoint PmulI (j:positive) (P:Pol) : Pol := match P with | Pc c => mkPinj j (PmulC Q c) | Pinj j' Q' => - match ZPminus j' j with + match Z.pos_sub j' j with | Zpos k => mkPinj j (Pmul (Pinj k Q') Q) | Z0 => mkPinj j (Pmul Q' Q) | Zneg k => mkPinj j' (PmulI k Q') @@ -322,13 +339,12 @@ Section MakeRingPol. | PX P' i' Q' => match j with | xH => mkPX (PmulI xH P') i' (Pmul Q' Q) - | xO j' => mkPX (PmulI j P') i' (PmulI (Pdouble_minus_one j') Q') + | xO j' => mkPX (PmulI j P') i' (PmulI (Pos.pred_double j') Q') | xI j' => mkPX (PmulI j P') i' (PmulI (xO j') Q') end end. End PmulI. -(* A symmetric version of the multiplication *) Fixpoint Pmul (P P'' : Pol) {struct P''} : Pol := match P'' with @@ -341,7 +357,7 @@ Section MakeRingPol. let QQ' := match j with | xH => Pmul Q Q' - | xO j => Pmul (Pinj (Pdouble_minus_one j) Q) Q' + | xO j => Pmul (Pinj (Pos.pred_double j) Q) Q' | xI j => Pmul (Pinj (xO j) Q) Q' end in mkPX (Pmul P P') i' QQ' @@ -354,25 +370,7 @@ Section MakeRingPol. end end. -(* Non symmetric *) -(* - Fixpoint Pmul_aux (P P' : Pol) {struct P'} : Pol := - match P' with - | Pc c' => PmulC P c' - | Pinj j' Q' => PmulI Pmul_aux Q' j' P - | PX P' i' Q' => - (mkPX (Pmul_aux P P') i' P0) ++ (PmulI Pmul_aux Q' xH P) - end. - - Definition Pmul P P' := - match P with - | Pc c => PmulC P' c - | Pinj j Q => PmulI Pmul_aux Q j P' - | PX P i Q => - (mkPX (Pmul_aux P P') i P0) ++ (PmulI Pmul_aux Q xH P') - end. -*) - Notation "P ** P'" := (Pmul P P'). + Infix "**" := Pmul. Fixpoint Psquare (P:Pol) : Pol := match P with @@ -387,26 +385,26 @@ Section MakeRingPol. (** Monomial **) + (** A monomial is X1^k1...Xi^ki. Its representation + is a simplified version of the polynomial representation: + + - [mon0] correspond to the polynom [P1]. + - [(zmon j M)] corresponds to [(Pinj j ...)], + i.e. skip j variable indices. + - [(vmon i M)] is X^i*M with X the current variable, + its corresponds to (PX P1 i ...)] + *) + Inductive Mon: Set := - mon0: Mon + | mon0: Mon | zmon: positive -> Mon -> Mon | vmon: positive -> Mon -> Mon. - Fixpoint Mphi(l:list R) (M: Mon) {struct M} : R := - match M with - mon0 => rI - | zmon j M1 => Mphi (jump j l) M1 - | vmon i M1 => - let x := hd 0 l in - let xi := pow_pos rmul x i in - (Mphi (tail l) M1) * xi - end. - Definition mkZmon j M := match M with mon0 => mon0 | _ => zmon j M end. Definition zmon_pred j M := - match j with xH => M | _ => mkZmon (Ppred j) M end. + match j with xH => M | _ => mkZmon (Pos.pred j) M end. Definition mkVmon i M := match M with @@ -421,7 +419,7 @@ Section MakeRingPol. | Pinj j1 P1 => let (R,S) := CFactor P1 c in (mkPinj j1 R, mkPinj j1 S) - | PX P1 i Q1 => + | PX P1 i Q1 => let (R1, S1) := CFactor P1 c in let (R2, S2) := CFactor Q1 c in (mkPX R1 i R2, mkPX S1 i S2) @@ -429,10 +427,7 @@ Section MakeRingPol. Fixpoint MFactor (P: Pol) (c: C) (M: Mon) {struct P}: Pol * Pol := match P, M with - _, mon0 => - if (ceqb c cI) then (Pc cO, P) else -(* if (ceqb c (copp cI)) then (Pc cO, Popp P) else Not in almost ring *) - CFactor P c + _, mon0 => if (ceqb c cI) then (Pc cO, P) else CFactor P c | Pc _, _ => (P, Pc cO) | Pinj j1 P1, zmon j2 M1 => match j1 ?= j2 with @@ -468,7 +463,7 @@ Section MakeRingPol. | _ => Some (Padd Q1 (Pmul P2 R1)) end. - Fixpoint PNSubst1 (P1: Pol) (cM1: C * Mon) (P2: Pol) (n: nat) {struct n}: Pol := + Fixpoint PNSubst1 (P1: Pol) (cM1: C * Mon) (P2: Pol) (n: nat) : Pol := match POneSubst P1 cM1 P2 with Some P3 => match n with S n1 => PNSubst1 P3 cM1 P2 n1 | _ => P3 end | _ => P1 @@ -480,14 +475,13 @@ Section MakeRingPol. | _ => None end. - Fixpoint PSubstL1 (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) {struct LM1}: - Pol := + Fixpoint PSubstL1 (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) : Pol := match LM1 with cons (M1,P2) LM2 => PSubstL1 (PNSubst1 P1 M1 P2 n) LM2 n | _ => P1 end. - Fixpoint PSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) {struct LM1}: option Pol := + Fixpoint PSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) : option Pol := match LM1 with cons (M1,P2) LM2 => match PNSubst P1 M1 P2 n with @@ -497,7 +491,7 @@ Section MakeRingPol. | _ => None end. - Fixpoint PNSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (m n: nat) {struct m}: Pol := + Fixpoint PNSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (m n: nat) : Pol := match PSubstL P1 LM1 n with Some P3 => match m with S m1 => PNSubstL P3 LM1 m1 n | _ => P3 end | _ => P1 @@ -505,658 +499,409 @@ Section MakeRingPol. (** Evaluation of a polynomial towards R *) - Fixpoint Pphi(l:list R) (P:Pol) {struct P} : R := + Local Notation hd := (List.hd 0). + + Fixpoint Pphi(l:list R) (P:Pol) : R := match P with | Pc c => [c] | Pinj j Q => Pphi (jump j l) Q - | PX P i Q => - let x := hd 0 l in - let xi := pow_pos rmul x i in - (Pphi l P) * xi + (Pphi (tail l) Q) + | PX P i Q => Pphi l P * (hd l) ^ i + Pphi (tail l) Q end. Reserved Notation "P @ l " (at level 10, no associativity). Notation "P @ l " := (Pphi l P). + + (** Evaluation of a monomial towards R *) + + Fixpoint Mphi(l:list R) (M: Mon) : R := + match M with + | mon0 => rI + | zmon j M1 => Mphi (jump j l) M1 + | vmon i M1 => Mphi (tail l) M1 * (hd l) ^ i + end. + + Notation "M @@ l" := (Mphi l M) (at level 10, no associativity). + (** Proofs *) - Lemma ZPminus_spec : forall x y, - match ZPminus x y with - | Z0 => x = y - | Zpos k => x = (y + k)%positive - | Zneg k => y = (x + k)%positive + + Ltac destr_pos_sub := + match goal with |- context [Z.pos_sub ?x ?y] => + generalize (Z.pos_sub_discr x y); destruct (Z.pos_sub x y) end. + + Lemma jump_add' i j (l:list R) : jump (i + j) l = jump j (jump i l). + Proof. rewrite Pos.add_comm. apply jump_add. Qed. + + Lemma Peq_ok P P' : (P ?== P') = true -> forall l, P@l == P'@ l. Proof. - induction x;destruct y. - replace (ZPminus (xI x) (xI y)) with (Zdouble (ZPminus x y));trivial. - assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial. - replace (ZPminus (xI x) (xO y)) with (Zdouble_plus_one (ZPminus x y));trivial. - assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_plus_one;rewrite H;trivial. - apply Pplus_xI_double_minus_one. - simpl;trivial. - replace (ZPminus (xO x) (xI y)) with (Zdouble_minus_one (ZPminus x y));trivial. - assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_minus_one;rewrite H;trivial. - apply Pplus_xI_double_minus_one. - replace (ZPminus (xO x) (xO y)) with (Zdouble (ZPminus x y));trivial. - assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial. - replace (ZPminus (xO x) xH) with (Zpos (Pdouble_minus_one x));trivial. - rewrite <- Pplus_one_succ_l. - rewrite Psucc_o_double_minus_one_eq_xO;trivial. - replace (ZPminus xH (xI y)) with (Zneg (xO y));trivial. - replace (ZPminus xH (xO y)) with (Zneg (Pdouble_minus_one y));trivial. - rewrite <- Pplus_one_succ_l. - rewrite Psucc_o_double_minus_one_eq_xO;trivial. - simpl;trivial. + revert P';induction P;destruct P';simpl; intros H l; try easy. + - now apply (morph_eq CRmorph). + - destruct (Pos.compare_spec p p0); [ subst | easy | easy ]. + now rewrite IHP. + - specialize (IHP1 P'1); specialize (IHP2 P'2). + destruct (Pos.compare_spec p p0); [ subst | easy | easy ]. + destruct (P2 ?== P'1); [|easy]. + rewrite H in *. + now rewrite IHP1, IHP2. Qed. - Lemma Peq_ok : forall P P', - (P ?== P') = true -> forall l, P@l == P'@ l. + Lemma Peq_spec P P' : + BoolSpec (forall l, P@l == P'@l) True (P ?== P'). Proof. - induction P;destruct P';simpl;intros;try discriminate;trivial. - apply (morph_eq CRmorph);trivial. - assert (H1 := Pos.compare_eq p p0); destruct (p ?= p0); - try discriminate H. - rewrite (IHP P' H); rewrite H1;trivial;rrefl. - assert (H1 := Pos.compare_eq p p0); destruct (p ?= p0); - try discriminate H. - rewrite H1;trivial. clear H1. - assert (H1 := IHP1 P'1);assert (H2 := IHP2 P'2); - destruct (P2 ?== P'1);[destruct (P3 ?== P'2); [idtac|discriminate H] - |discriminate H]. - rewrite (H1 H);rewrite (H2 H);rrefl. + generalize (Peq_ok P P'). destruct (P ?== P'); auto. Qed. - Lemma Pphi0 : forall l, P0@l == 0. + Lemma Pphi0 l : P0@l == 0. Proof. - intros;simpl;apply (morph0 CRmorph). + simpl;apply (morph0 CRmorph). Qed. - Lemma Pphi1 : forall l, P1@l == 1. + Lemma Pphi1 l : P1@l == 1. Proof. - intros;simpl;apply (morph1 CRmorph). + simpl;apply (morph1 CRmorph). Qed. - Lemma mkPinj_ok : forall j l P, (mkPinj j P)@l == P@(jump j l). + Lemma mkPinj_ok j l P : (mkPinj j P)@l == P@(jump j l). Proof. - intros j l p;destruct p;simpl;rsimpl. - rewrite <-jump_Pplus;rewrite Pplus_comm;rrefl. + destruct P;simpl;rsimpl. + now rewrite jump_add'. Qed. - Let pow_pos_Pplus := - pow_pos_Pplus rmul Rsth Reqe.(Rmul_ext) ARth.(ARmul_comm) ARth.(ARmul_assoc). + Lemma pow_pos_add x i j : x^(j + i) == x^i * x^j. + Proof. + rewrite Pos.add_comm. + apply (pow_pos_add Rsth Reqe.(Rmul_ext) ARth.(ARmul_assoc)). + Qed. - Lemma mkPX_ok : forall l P i Q, - (mkPX P i Q)@l == P@l*(pow_pos rmul (hd 0 l) i) + Q@(tail l). + Lemma ceqb_spec c c' : BoolSpec ([c] == [c']) True (c ?=! c'). Proof. - intros l P i Q;unfold mkPX. - destruct P;try (simpl;rrefl). - assert (H := morph_eq CRmorph c cO);destruct (c ?=! cO);simpl;try rrefl. - rewrite (H (refl_equal true));rewrite (morph0 CRmorph). - rewrite mkPinj_ok;rsimpl;simpl;rrefl. - assert (H := @Peq_ok P3 P0);destruct (P3 ?== P0);simpl;try rrefl. - rewrite (H (refl_equal true));trivial. - rewrite Pphi0. rewrite pow_pos_Pplus;rsimpl. + generalize (morph_eq CRmorph c c'). + destruct (c ?=! c'); auto. Qed. - Ltac Esimpl := - repeat (progress ( - match goal with - | |- context [?P@?l] => - match P with - | P0 => rewrite (Pphi0 l) - | P1 => rewrite (Pphi1 l) - | (mkPinj ?j ?P) => rewrite (mkPinj_ok j l P) - | (mkPX ?P ?i ?Q) => rewrite (mkPX_ok l P i Q) - end - | |- context [[?c]] => - match c with - | cO => rewrite (morph0 CRmorph) - | cI => rewrite (morph1 CRmorph) - | ?x +! ?y => rewrite ((morph_add CRmorph) x y) - | ?x *! ?y => rewrite ((morph_mul CRmorph) x y) - | ?x -! ?y => rewrite ((morph_sub CRmorph) x y) - | -! ?x => rewrite ((morph_opp CRmorph) x) - end - end)); - rsimpl; simpl. - - Lemma PaddC_ok : forall c P l, (PaddC P c)@l == P@l + [c]. + Lemma mkPX_ok l P i Q : + (mkPX P i Q)@l == P@l * (hd l)^i + Q@(tail l). Proof. - induction P;simpl;intros;Esimpl;trivial. - rewrite IHP2;rsimpl. + unfold mkPX. destruct P. + - case ceqb_spec; intros H; simpl; try reflexivity. + rewrite H, (morph0 CRmorph), mkPinj_ok; rsimpl. + - reflexivity. + - case Peq_spec; intros H; simpl; try reflexivity. + rewrite H, Pphi0, Pos.add_comm, pow_pos_add; rsimpl. Qed. - Lemma PsubC_ok : forall c P l, (PsubC P c)@l == P@l - [c]. + Hint Rewrite + Pphi0 + Pphi1 + mkPinj_ok + mkPX_ok + (morph0 CRmorph) + (morph1 CRmorph) + (morph0 CRmorph) + (morph_add CRmorph) + (morph_mul CRmorph) + (morph_sub CRmorph) + (morph_opp CRmorph) + : Esimpl. + + (* Quicker than autorewrite with Esimpl :-) *) + Ltac Esimpl := try rewrite_db Esimpl; rsimpl; simpl. + + Lemma PaddC_ok c P l : (PaddC P c)@l == P@l + [c]. Proof. - induction P;simpl;intros. - Esimpl. - rewrite IHP;rsimpl. + revert l;induction P;simpl;intros;Esimpl;trivial. rewrite IHP2;rsimpl. Qed. - Lemma PmulC_aux_ok : forall c P l, (PmulC_aux P c)@l == P@l * [c]. + Lemma PsubC_ok c P l : (PsubC P c)@l == P@l - [c]. Proof. - induction P;simpl;intros;Esimpl;trivial. - rewrite IHP1;rewrite IHP2;rsimpl. - mul_push ([c]);rrefl. + revert l;induction P;simpl;intros. + - Esimpl. + - rewrite IHP;rsimpl. + - rewrite IHP2;rsimpl. Qed. - Lemma PmulC_ok : forall c P l, (PmulC P c)@l == P@l * [c]. + Lemma PmulC_aux_ok c P l : (PmulC_aux P c)@l == P@l * [c]. Proof. - intros c P l; unfold PmulC. - assert (H:= morph_eq CRmorph c cO);destruct (c ?=! cO). - rewrite (H (refl_equal true));Esimpl. - assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI). - rewrite (H1 (refl_equal true));Esimpl. - apply PmulC_aux_ok. + revert l;induction P;simpl;intros;Esimpl;trivial. + rewrite IHP1, IHP2;rsimpl. add_permut. mul_permut. Qed. - Lemma Popp_ok : forall P l, (--P)@l == - P@l. + Lemma PmulC_ok c P l : (PmulC P c)@l == P@l * [c]. Proof. - induction P;simpl;intros. - Esimpl. - apply IHP. - rewrite IHP1;rewrite IHP2;rsimpl. + unfold PmulC. + case ceqb_spec; intros H. + - rewrite H; Esimpl. + - case ceqb_spec; intros H'. + + rewrite H'; Esimpl. + + apply PmulC_aux_ok. Qed. - Ltac Esimpl2 := - Esimpl; - repeat (progress ( - match goal with - | |- context [(PaddC ?P ?c)@?l] => rewrite (PaddC_ok c P l) - | |- context [(PsubC ?P ?c)@?l] => rewrite (PsubC_ok c P l) - | |- context [(PmulC ?P ?c)@?l] => rewrite (PmulC_ok c P l) - | |- context [(--?P)@?l] => rewrite (Popp_ok P l) - end)); Esimpl. - - Lemma Padd_ok : forall P' P l, (P ++ P')@l == P@l + P'@l. + Lemma Popp_ok P l : (--P)@l == - P@l. Proof. - induction P';simpl;intros;Esimpl2. - generalize P p l;clear P p l. - induction P;simpl;intros. - Esimpl2;apply (ARadd_comm ARth). - assert (H := ZPminus_spec p p0);destruct (ZPminus p p0). - rewrite H;Esimpl. rewrite IHP';rrefl. - rewrite H;Esimpl. rewrite IHP';Esimpl. - rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl. - rewrite H;Esimpl. rewrite IHP. - rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl. - destruct p0;simpl. - rewrite IHP2;simpl;rsimpl. - rewrite IHP2;simpl. - rewrite jump_Pdouble_minus_one;rsimpl. - rewrite IHP';rsimpl. - destruct P;simpl. - Esimpl2;add_push [c];rrefl. - destruct p0;simpl;Esimpl2. - rewrite IHP'2;simpl. - rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl. - rewrite IHP'2;simpl. - rewrite jump_Pdouble_minus_one;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl. - rewrite IHP'2;rsimpl. add_push (P @ (tail l));rrefl. - assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2. - rewrite IHP'1;rewrite IHP'2;rsimpl. - add_push (P3 @ (tail l));rewrite H;rrefl. - rewrite IHP'1;rewrite IHP'2;simpl;Esimpl. - rewrite H;rewrite Pplus_comm. - rewrite pow_pos_Pplus;rsimpl. - add_push (P3 @ (tail l));rrefl. - assert (forall P k l, - (PaddX Padd P'1 k P) @ l == P@l + P'1@l * pow_pos rmul (hd 0 l) k). - induction P;simpl;intros;try apply (ARadd_comm ARth). - destruct p2;simpl;try apply (ARadd_comm ARth). - rewrite jump_Pdouble_minus_one;apply (ARadd_comm ARth). - assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2. - rewrite IHP'1;rsimpl; rewrite H1;add_push (P5 @ (tail l0));rrefl. - rewrite IHP'1;simpl;Esimpl. - rewrite H1;rewrite Pplus_comm. - rewrite pow_pos_Pplus;simpl;Esimpl. - add_push (P5 @ (tail l0));rrefl. - rewrite IHP1;rewrite H1;rewrite Pplus_comm. - rewrite pow_pos_Pplus;simpl;rsimpl. - add_push (P5 @ (tail l0));rrefl. - rewrite H0;rsimpl. - add_push (P3 @ (tail l)). - rewrite H;rewrite Pplus_comm. - rewrite IHP'2;rewrite pow_pos_Pplus;rsimpl. - add_push (P3 @ (tail l));rrefl. + revert l;induction P;simpl;intros. + - Esimpl. + - apply IHP. + - rewrite IHP1, IHP2;rsimpl. Qed. - Lemma Psub_ok : forall P' P l, (P -- P')@l == P@l - P'@l. + Hint Rewrite PaddC_ok PsubC_ok PmulC_ok Popp_ok : Esimpl. + + Lemma PaddX_ok P' P k l : + (forall P l, (P++P')@l == P@l + P'@l) -> + (PaddX Padd P' k P) @ l == P@l + P'@l * (hd l)^k. Proof. - induction P';simpl;intros;Esimpl2;trivial. - generalize P p l;clear P p l. - induction P;simpl;intros. - Esimpl2;apply (ARadd_comm ARth). - assert (H := ZPminus_spec p p0);destruct (ZPminus p p0). - rewrite H;Esimpl. rewrite IHP';rsimpl. - rewrite H;Esimpl. rewrite IHP';Esimpl. - rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl. - rewrite H;Esimpl. rewrite IHP. - rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl. - destruct p0;simpl. - rewrite IHP2;simpl;rsimpl. - rewrite IHP2;simpl. - rewrite jump_Pdouble_minus_one;rsimpl. - rewrite IHP';rsimpl. - destruct P;simpl. - repeat rewrite Popp_ok;Esimpl2;rsimpl;add_push [c];try rrefl. - destruct p0;simpl;Esimpl2. - rewrite IHP'2;simpl;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));trivial. - add_push (P @ (jump p0 (jump p0 (tail l))));rrefl. - rewrite IHP'2;simpl;rewrite jump_Pdouble_minus_one;rsimpl. - add_push (- (P'1 @ l * pow_pos rmul (hd 0 l) p));rrefl. - rewrite IHP'2;rsimpl;add_push (P @ (tail l));rrefl. - assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2. - rewrite IHP'1; rewrite IHP'2;rsimpl. - add_push (P3 @ (tail l));rewrite H;rrefl. - rewrite IHP'1; rewrite IHP'2;rsimpl;simpl;Esimpl. - rewrite H;rewrite Pplus_comm. - rewrite pow_pos_Pplus;rsimpl. - add_push (P3 @ (tail l));rrefl. - assert (forall P k l, - (PsubX Psub P'1 k P) @ l == P@l + - P'1@l * pow_pos rmul (hd 0 l) k). - induction P;simpl;intros. - rewrite Popp_ok;rsimpl;apply (ARadd_comm ARth);trivial. - destruct p2;simpl;rewrite Popp_ok;rsimpl. - apply (ARadd_comm ARth);trivial. - rewrite jump_Pdouble_minus_one;apply (ARadd_comm ARth);trivial. - apply (ARadd_comm ARth);trivial. - assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2;rsimpl. - rewrite IHP'1;rsimpl;add_push (P5 @ (tail l0));rewrite H1;rrefl. - rewrite IHP'1;rewrite H1;rewrite Pplus_comm. - rewrite pow_pos_Pplus;simpl;Esimpl. - add_push (P5 @ (tail l0));rrefl. - rewrite IHP1;rewrite H1;rewrite Pplus_comm. - rewrite pow_pos_Pplus;simpl;rsimpl. - add_push (P5 @ (tail l0));rrefl. - rewrite H0;rsimpl. - rewrite IHP'2;rsimpl;add_push (P3 @ (tail l)). - rewrite H;rewrite Pplus_comm. - rewrite pow_pos_Pplus;rsimpl. + intros IHP'. + revert k l. induction P;simpl;intros. + - add_permut. + - destruct p; simpl; + rewrite ?jump_pred_double; add_permut. + - destr_pos_sub; intros ->;Esimpl. + + rewrite IHP';rsimpl. add_permut. + + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut. + + rewrite IHP1, pow_pos_add;rsimpl. add_permut. Qed. -(* Proof for the symmetriv version *) - Lemma PmulI_ok : - forall P', - (forall (P : Pol) (l : list R), (Pmul P P') @ l == P @ l * P' @ l) -> - forall (P : Pol) (p : positive) (l : list R), - (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l). + Lemma Padd_ok P' P l : (P ++ P')@l == P@l + P'@l. Proof. - induction P;simpl;intros. - Esimpl2;apply (ARmul_comm ARth). - assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2. - rewrite H1; rewrite H;rrefl. - rewrite H1; rewrite H. - rewrite Pplus_comm. - rewrite jump_Pplus;simpl;rrefl. - rewrite H1;rewrite Pplus_comm. - rewrite jump_Pplus;rewrite IHP;rrefl. - destruct p0;Esimpl2. - rewrite IHP1;rewrite IHP2;simpl;rsimpl. - mul_push (pow_pos rmul (hd 0 l) p);rrefl. - rewrite IHP1;rewrite IHP2;simpl;rsimpl. - mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_Pdouble_minus_one;rrefl. - rewrite IHP1;simpl;rsimpl. - mul_push (pow_pos rmul (hd 0 l) p). - rewrite H;rrefl. + revert P l; induction P';simpl;intros;Esimpl. + - revert p l; induction P;simpl;intros. + + Esimpl; add_permut. + + destr_pos_sub; intros ->;Esimpl. + * now rewrite IHP'. + * rewrite IHP';Esimpl. now rewrite jump_add'. + * rewrite IHP. now rewrite jump_add'. + + destruct p0;simpl. + * rewrite IHP2;simpl. rsimpl. + * rewrite IHP2;simpl. rewrite jump_pred_double. rsimpl. + * rewrite IHP'. rsimpl. + - destruct P;simpl. + + Esimpl. add_permut. + + destruct p0;simpl;Esimpl; rewrite IHP'2; simpl. + * rsimpl. add_permut. + * rewrite jump_pred_double. rsimpl. add_permut. + * rsimpl. add_permut. + + destr_pos_sub; intros ->; Esimpl. + * rewrite IHP'1, IHP'2;rsimpl. add_permut. + * rewrite IHP'1, IHP'2;simpl;Esimpl. + rewrite pow_pos_add;rsimpl. add_permut. + * rewrite PaddX_ok by trivial; rsimpl. + rewrite IHP'2, pow_pos_add; rsimpl. add_permut. Qed. -(* - Lemma PmulI_ok : - forall P', - (forall (P : Pol) (l : list R), (Pmul_aux P P') @ l == P @ l * P' @ l) -> - forall (P : Pol) (p : positive) (l : list R), - (PmulI Pmul_aux P' p P) @ l == P @ l * P' @ (jump p l). + Lemma PsubX_ok P' P k l : + (forall P l, (P--P')@l == P@l - P'@l) -> + (PsubX Psub P' k P) @ l == P@l - P'@l * (hd l)^k. Proof. - induction P;simpl;intros. - Esimpl2;apply (ARmul_comm ARth). - assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2. - rewrite H1; rewrite H;rrefl. - rewrite H1; rewrite H. - rewrite Pplus_comm. - rewrite jump_Pplus;simpl;rrefl. - rewrite H1;rewrite Pplus_comm. - rewrite jump_Pplus;rewrite IHP;rrefl. - destruct p0;Esimpl2. - rewrite IHP1;rewrite IHP2;simpl;rsimpl. - mul_push (pow_pos rmul (hd 0 l) p);rrefl. - rewrite IHP1;rewrite IHP2;simpl;rsimpl. - mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_Pdouble_minus_one;rrefl. - rewrite IHP1;simpl;rsimpl. - mul_push (pow_pos rmul (hd 0 l) p). - rewrite H;rrefl. + intros IHP'. + revert k l. induction P;simpl;intros. + - rewrite Popp_ok;rsimpl; add_permut. + - destruct p; simpl; + rewrite Popp_ok;rsimpl; + rewrite ?jump_pred_double; add_permut. + - destr_pos_sub; intros ->; Esimpl. + + rewrite IHP';rsimpl. add_permut. + + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut. + + rewrite IHP1, pow_pos_add;rsimpl. add_permut. Qed. - Lemma Pmul_aux_ok : forall P' P l,(Pmul_aux P P')@l == P@l * P'@l. + Lemma Psub_ok P' P l : (P -- P')@l == P@l - P'@l. Proof. - induction P';simpl;intros. - Esimpl2;trivial. - apply PmulI_ok;trivial. - rewrite Padd_ok;Esimpl2. - rewrite (PmulI_ok P'2 IHP'2). rewrite IHP'1. rrefl. + revert P l; induction P';simpl;intros;Esimpl. + - revert p l; induction P;simpl;intros. + + Esimpl; add_permut. + + destr_pos_sub; intros ->;Esimpl. + * rewrite IHP';rsimpl. + * rewrite IHP';Esimpl. now rewrite jump_add'. + * rewrite IHP. now rewrite jump_add'. + + destruct p0;simpl. + * rewrite IHP2;simpl. rsimpl. + * rewrite IHP2;simpl. rewrite jump_pred_double. rsimpl. + * rewrite IHP'. rsimpl. + - destruct P;simpl. + + Esimpl; add_permut. + + destruct p0;simpl;Esimpl; rewrite IHP'2; simpl. + * rsimpl. add_permut. + * rewrite jump_pred_double. rsimpl. add_permut. + * rsimpl. add_permut. + + destr_pos_sub; intros ->; Esimpl. + * rewrite IHP'1, IHP'2;rsimpl. add_permut. + * rewrite IHP'1, IHP'2;simpl;Esimpl. + rewrite pow_pos_add;rsimpl. add_permut. + * rewrite PsubX_ok by trivial;rsimpl. + rewrite IHP'2, pow_pos_add;rsimpl. add_permut. Qed. -*) -(* Proof for the symmetric version *) - Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l. + Lemma PmulI_ok P' : + (forall P l, (Pmul P P') @ l == P @ l * P' @ l) -> + forall P p l, (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l). Proof. - intros P P';generalize P;clear P;induction P';simpl;intros. - apply PmulC_ok. apply PmulI_ok;trivial. - destruct P. - rewrite (ARmul_comm ARth);Esimpl2;Esimpl2. - Esimpl2. rewrite IHP'1;Esimpl2. - assert (match p0 with - | xI j => Pinj (xO j) P ** P'2 - | xO j => Pinj (Pdouble_minus_one j) P ** P'2 - | 1 => P ** P'2 - end @ (tail l) == P @ (jump p0 l) * P'2 @ (tail l)). - destruct p0;simpl;rewrite IHP'2;Esimpl. - rewrite jump_Pdouble_minus_one;Esimpl. - rewrite H;Esimpl. - rewrite Padd_ok; Esimpl2. rewrite Padd_ok; Esimpl2. - repeat (rewrite IHP'1 || rewrite IHP'2);simpl. - rewrite PmulI_ok;trivial. - mul_push (P'1@l). simpl. mul_push (P'2 @ (tail l)). Esimpl. + intros IHP'. + induction P;simpl;intros. + - Esimpl; mul_permut. + - destr_pos_sub; intros ->;Esimpl. + + now rewrite IHP'. + + now rewrite IHP', jump_add'. + + now rewrite IHP, jump_add'. + - destruct p0;Esimpl; rewrite ?IHP1, ?IHP2; rsimpl. + + f_equiv. mul_permut. + + rewrite jump_pred_double. f_equiv. mul_permut. + + rewrite IHP'. f_equiv. mul_permut. Qed. -(* -Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l. + Lemma Pmul_ok P P' l : (P**P')@l == P@l * P'@l. Proof. - destruct P;simpl;intros. - Esimpl2;apply (ARmul_comm ARth). - rewrite (PmulI_ok P (Pmul_aux_ok P)). - apply (ARmul_comm ARth). - rewrite Padd_ok; Esimpl2. - rewrite (PmulI_ok P3 (Pmul_aux_ok P3));trivial. - rewrite Pmul_aux_ok;mul_push (P' @ l). - rewrite (ARmul_comm ARth (P' @ l));rrefl. + revert P l;induction P';simpl;intros. + - apply PmulC_ok. + - apply PmulI_ok;trivial. + - destruct P. + + rewrite (ARmul_comm ARth). Esimpl. + + Esimpl. f_equiv. rewrite IHP'1; Esimpl. + destruct p0;rewrite IHP'2;Esimpl. + rewrite jump_pred_double; Esimpl. + + rewrite Padd_ok, !mkPX_ok, Padd_ok, !mkPX_ok, + !IHP'1, !IHP'2, PmulI_ok; trivial. simpl. Esimpl. + add_permut; f_equiv; mul_permut. Qed. -*) - Lemma Psquare_ok : forall P l, (Psquare P)@l == P@l * P@l. + Lemma Psquare_ok P l : (Psquare P)@l == P@l * P@l. Proof. - induction P;simpl;intros;Esimpl2. - apply IHP. rewrite Padd_ok. rewrite Pmul_ok;Esimpl2. - rewrite IHP1;rewrite IHP2. - mul_push (pow_pos rmul (hd 0 l) p). mul_push (P2@l). - rrefl. + revert l;induction P;simpl;intros;Esimpl. + - apply IHP. + - rewrite Padd_ok, Pmul_ok;Esimpl. + rewrite IHP1, IHP2. + mul_push ((hd l)^p). now mul_push (P2@l). Qed. - - Lemma mkZmon_ok: forall M j l, - Mphi l (mkZmon j M) == Mphi l (zmon j M). - intros M j l; case M; simpl; intros; rsimpl. + Lemma mkZmon_ok M j l : + (mkZmon j M) @@ l == (zmon j M) @@ l. + Proof. + destruct M; simpl; rsimpl. Qed. - Lemma zmon_pred_ok : forall M j l, - Mphi (tail l) (zmon_pred j M) == Mphi l (zmon j M). + Lemma zmon_pred_ok M j l : + (zmon_pred j M) @@ (tail l) == (zmon j M) @@ l. Proof. - destruct j; simpl;intros auto; rsimpl. - rewrite mkZmon_ok;rsimpl. - rewrite mkZmon_ok;simpl. rewrite jump_Pdouble_minus_one; rsimpl. + destruct j; simpl; rewrite ?mkZmon_ok; simpl; rsimpl. + rewrite jump_pred_double; rsimpl. Qed. - Lemma mkVmon_ok : forall M i l, Mphi l (mkVmon i M) == Mphi l M*pow_pos rmul (hd 0 l) i. + Lemma mkVmon_ok M i l : + (mkVmon i M)@@l == M@@l * (hd l)^i. Proof. destruct M;simpl;intros;rsimpl. - rewrite zmon_pred_ok;simpl;rsimpl. - rewrite Pplus_comm;rewrite pow_pos_Pplus;rsimpl. + - rewrite zmon_pred_ok;simpl;rsimpl. + - rewrite pow_pos_add;rsimpl. Qed. - Lemma Mcphi_ok: forall P c l, - let (Q,R) := CFactor P c in - P@l == Q@l + (phi c) * (R@l). + Ltac destr_factor := match goal with + | H : context [CFactor ?P _] |- context [CFactor ?P ?c] => + destruct (CFactor P c); destr_factor; rewrite H; clear H + | H : context [MFactor ?P _ _] |- context [MFactor ?P ?c ?M] => + specialize (H M); destruct (MFactor P c M); destr_factor; rewrite H; clear H + | _ => idtac + end. + + Lemma Mcphi_ok P c l : + let (Q,R) := CFactor P c in + P@l == Q@l + [c] * R@l. Proof. - intros P; elim P; simpl; auto; clear P. - intros c c1 l; generalize (div_th.(div_eucl_th) c c1); case cdiv. - intros q r H; rewrite H. - Esimpl. - rewrite (ARadd_comm ARth); rsimpl. - intros i P Hrec c l. - generalize (Hrec c (jump i l)); case CFactor. - intros R1 S1; Esimpl; auto. - intros Q1 Qrec i R1 Rrec c l. - generalize (Qrec c l); case CFactor; intros S1 S2 HS. - generalize (Rrec c (tail l)); case CFactor; intros S3 S4 HS1. - rewrite HS; rewrite HS1; Esimpl. - apply (Radd_ext Reqe); rsimpl. - repeat rewrite <- (ARadd_assoc ARth). - apply (Radd_ext Reqe); rsimpl. - rewrite (ARadd_comm ARth); rsimpl. + revert l. + induction P as [c0 | j P IH | P1 IH1 i P2 IH2]; intros l; Esimpl. + - assert (H := div_th.(div_eucl_th) c0 c). + destruct cdiv as (q,r). rewrite H; Esimpl. add_permut. + - destr_factor. Esimpl. + - destr_factor. Esimpl. add_permut. Qed. - Lemma Mphi_ok: forall P (cM: C * Mon) l, - let (c,M) := cM in - let (Q,R) := MFactor P c M in - P@l == Q@l + (phi c) * (Mphi l M) * (R@l). + Lemma Mphi_ok P (cM: C * Mon) l : + let (c,M) := cM in + let (Q,R) := MFactor P c M in + P@l == Q@l + [c] * M@@l * R@l. Proof. - intros P; elim P; simpl; auto; clear P. - intros c (c1, M) l; case M; simpl; auto. - assert (H1:= morph_eq CRmorph c1 cI);destruct (c1 ?=! cI). - rewrite (H1 (refl_equal true));Esimpl. - try rewrite (morph0 CRmorph); rsimpl. - generalize (div_th.(div_eucl_th) c c1); case (cdiv c c1). - intros q r H; rewrite H; clear H H1. - Esimpl. - rewrite (ARadd_comm ARth); rsimpl. - intros p m; Esimpl. - intros p m; Esimpl. - intros i P Hrec (c,M) l; case M; simpl; clear M. - assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI). - rewrite (H1 (refl_equal true));Esimpl. - Esimpl. - generalize (Mcphi_ok P c (jump i l)); case CFactor. - intros R1 Q1 HH; rewrite HH; Esimpl. - intros j M. - case_eq (i ?= j); intros He; simpl. - rewrite (Pos.compare_eq _ _ He). - generalize (Hrec (c, M) (jump j l)); case (MFactor P c M); - simpl; intros P2 Q2 H; repeat rewrite mkPinj_ok; auto. - generalize (Hrec (c, (zmon (j -i) M)) (jump i l)); - case (MFactor P c (zmon (j -i) M)); simpl. - intros P2 Q2 H; repeat rewrite mkPinj_ok; auto. - rewrite <- (Pplus_minus _ _ (ZC2 _ _ He)). - rewrite Pplus_comm; rewrite jump_Pplus; auto. - rewrite (morph0 CRmorph); rsimpl. - intros P2 m; rewrite (morph0 CRmorph); rsimpl. - - intros P2 Hrec1 i Q2 Hrec2 (c, M) l; case M; simpl; auto. - assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI). - rewrite (H1 (refl_equal true));Esimpl. - Esimpl. - generalize (Mcphi_ok P2 c l); case CFactor. - intros S1 S2 HS. - generalize (Mcphi_ok Q2 c (tail l)); case CFactor. - intros S3 S4 HS1; Esimpl; rewrite HS; rewrite HS1. - rsimpl. - apply (Radd_ext Reqe); rsimpl. - repeat rewrite <- (ARadd_assoc ARth). - apply (Radd_ext Reqe); rsimpl. - rewrite (ARadd_comm ARth); rsimpl. - intros j M1. - generalize (Hrec1 (c,zmon j M1) l); - case (MFactor P2 c (zmon j M1)). - intros R1 S1 H1. - generalize (Hrec2 (c, zmon_pred j M1) (List.tail l)); - case (MFactor Q2 c (zmon_pred j M1)); simpl. - intros R2 S2 H2; rewrite H1; rewrite H2. - repeat rewrite mkPX_ok; simpl. - rsimpl. - apply radd_ext; rsimpl. - rewrite (ARadd_comm ARth); rsimpl. - apply radd_ext; rsimpl. - rewrite (ARadd_comm ARth); rsimpl. - rewrite zmon_pred_ok;rsimpl. - intros j M1. - case_eq (i ?= j); intros He; simpl. - rewrite (Pos.compare_eq _ _ He). - generalize (Hrec1 (c, mkZmon xH M1) l); case (MFactor P2 c (mkZmon xH M1)); - simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto. - rewrite H; rewrite mkPX_ok; rsimpl. - repeat (rewrite <-(ARadd_assoc ARth)). - apply radd_ext; rsimpl. - rewrite (ARadd_comm ARth); rsimpl. - apply radd_ext; rsimpl. - repeat (rewrite <-(ARmul_assoc ARth)). - rewrite mkZmon_ok. - apply rmul_ext; rsimpl. - repeat (rewrite <-(ARmul_assoc ARth)). - apply rmul_ext; rsimpl. - rewrite (ARmul_comm ARth); rsimpl. - generalize (Hrec1 (c, vmon (j - i) M1) l); - case (MFactor P2 c (vmon (j - i) M1)); - simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto. - rewrite H; rsimpl; repeat rewrite mkPinj_ok; auto. - rewrite mkPX_ok; rsimpl. - repeat (rewrite <-(ARadd_assoc ARth)). - apply radd_ext; rsimpl. - rewrite (ARadd_comm ARth); rsimpl. - apply radd_ext; rsimpl. - repeat (rewrite <-(ARmul_assoc ARth)). - apply rmul_ext; rsimpl. - rewrite (ARmul_comm ARth); rsimpl. - apply rmul_ext; rsimpl. - rewrite <- (ARmul_comm ARth (Mphi (tail l) M1)); rsimpl. - repeat (rewrite <-(ARmul_assoc ARth)). - apply rmul_ext; rsimpl. - rewrite <- pow_pos_Pplus. - rewrite (Pplus_minus _ _ (ZC2 _ _ He)); rsimpl. - generalize (Hrec1 (c, mkZmon 1 M1) l); - case (MFactor P2 c (mkZmon 1 M1)); - simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto. - rewrite H; rsimpl. - rewrite mkPX_ok; rsimpl. - repeat (rewrite <-(ARadd_assoc ARth)). - apply radd_ext; rsimpl. - rewrite (ARadd_comm ARth); rsimpl. - apply radd_ext; rsimpl. - rewrite mkZmon_ok. - repeat (rewrite <-(ARmul_assoc ARth)). - apply rmul_ext; rsimpl. - rewrite (ARmul_comm ARth); rsimpl. - rewrite mkPX_ok; simpl; rsimpl. - rewrite (morph0 CRmorph); rsimpl. - repeat (rewrite <-(ARmul_assoc ARth)). - rewrite (ARmul_comm ARth (Q3@l)); rsimpl. - apply rmul_ext; rsimpl. - rewrite (ARmul_comm ARth); rsimpl. - repeat (rewrite <- (ARmul_assoc ARth)). - apply rmul_ext; rsimpl. - rewrite <- pow_pos_Pplus. - rewrite (Pplus_minus _ _ He); rsimpl. + destruct cM as (c,M). revert M l. + induction P; destruct M; intros l; simpl; auto; + try (case ceqb_spec; intro He); + try (case Pos.compare_spec; intros He); rewrite ?He; + destr_factor; simpl; Esimpl. + - assert (H := div_th.(div_eucl_th) c0 c). + destruct cdiv as (q,r). rewrite H; Esimpl. add_permut. + - assert (H := Mcphi_ok P c). destr_factor. Esimpl. + - now rewrite <- jump_add, Pos.sub_add. + - assert (H2 := Mcphi_ok P2 c). assert (H3 := Mcphi_ok P3 c). + destr_factor. Esimpl. add_permut. + - rewrite zmon_pred_ok. simpl. add_permut. + - rewrite mkZmon_ok. simpl. add_permut. mul_permut. + - add_permut. mul_permut. + rewrite <- pow_pos_add, Pos.add_comm, Pos.sub_add by trivial; rsimpl. + - rewrite mkZmon_ok. simpl. Esimpl. add_permut. mul_permut. + rewrite <- pow_pos_add, Pos.sub_add by trivial; rsimpl. Qed. -(* Proof for the symmetric version *) - - Lemma POneSubst_ok: forall P1 M1 P2 P3 l, - POneSubst P1 M1 P2 = Some P3 -> phi (fst M1) * Mphi l (snd M1) == P2@l -> P1@l == P3@l. + Lemma POneSubst_ok P1 cM1 P2 P3 l : + POneSubst P1 cM1 P2 = Some P3 -> + [fst cM1] * (snd cM1)@@l == P2@l -> P1@l == P3@l. Proof. - intros P2 (cc,M1) P3 P4 l; unfold POneSubst. - generalize (Mphi_ok P2 (cc, M1) l); case (MFactor P2 cc M1); simpl; auto. - intros Q1 R1; case R1. - intros c H; rewrite H. - generalize (morph_eq CRmorph c cO); - case (c ?=! cO); simpl; auto. - intros H1 H2; rewrite H1; auto; rsimpl. - discriminate. - intros _ H1 H2; injection H1; intros; subst. - rewrite H2; rsimpl. - (* new version *) - rewrite Padd_ok; rewrite PmulC_ok; rsimpl. - intros i P5 H; rewrite H. - intros HH H1; injection HH; intros; subst; rsimpl. - rewrite Padd_ok; rewrite PmulI_ok by (intros;apply Pmul_ok). rewrite H1; rsimpl. - intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3. - assert (P4 = Q1 ++ P3 ** PX i P5 P6). - injection H2; intros; subst;trivial. - rewrite H;rewrite Padd_ok;rewrite Pmul_ok;rsimpl. - Qed. -(* - Lemma POneSubst_ok: forall P1 M1 P2 P3 l, - POneSubst P1 M1 P2 = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l. -Proof. - intros P2 M1 P3 P4 l; unfold POneSubst. - generalize (Mphi_ok P2 M1 l); case (MFactor P2 M1); simpl; auto. - intros Q1 R1; case R1. - intros c H; rewrite H. - generalize (morph_eq CRmorph c cO); - case (c ?=! cO); simpl; auto. - intros H1 H2; rewrite H1; auto; rsimpl. - discriminate. - intros _ H1 H2; injection H1; intros; subst. - rewrite H2; rsimpl. - rewrite Padd_ok; rewrite Pmul_ok; rsimpl. - intros i P5 H; rewrite H. - intros HH H1; injection HH; intros; subst; rsimpl. - rewrite Padd_ok; rewrite Pmul_ok. rewrite H1; rsimpl. - intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3. - injection H2; intros; subst; rsimpl. - rewrite Padd_ok. - rewrite Pmul_ok; rsimpl. + destruct cM1 as (cc,M1). + unfold POneSubst. + assert (H := Mphi_ok P1 (cc, M1) l). simpl in H. + destruct MFactor as (R1,S1); simpl. rewrite H. clear H. + intros EQ EQ'. replace P3 with (R1 ++ P2 ** S1). + - rewrite EQ', Padd_ok, Pmul_ok; rsimpl. + - revert EQ. destruct S1; try now injection 1. + case ceqb_spec; now inversion 2. Qed. -*) - Lemma PNSubst1_ok: forall n P1 M1 P2 l, - [fst M1] * Mphi l (snd M1) == P2@l -> P1@l == (PNSubst1 P1 M1 P2 n)@l. + + Lemma PNSubst1_ok n P1 cM1 P2 l : + [fst cM1] * (snd cM1)@@l == P2@l -> + P1@l == (PNSubst1 P1 cM1 P2 n)@l. Proof. - intros n; elim n; simpl; auto. - intros P2 M1 P3 l H. - generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l); - case (POneSubst P2 M1 P3); [idtac | intros; rsimpl]. - intros P4 Hrec; rewrite (Hrec P4); auto; rsimpl. - intros n1 Hrec P2 M1 P3 l H. - generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l); - case (POneSubst P2 M1 P3); [idtac | intros; rsimpl]. - intros P4 Hrec1; rewrite (Hrec1 P4); auto; rsimpl. + revert P1. induction n; simpl; intros P1; + generalize (POneSubst_ok P1 cM1 P2); destruct POneSubst; + intros; rewrite <- ?IHn; auto; reflexivity. Qed. - Lemma PNSubst_ok: forall n P1 M1 P2 l P3, - PNSubst P1 M1 P2 n = Some P3 -> [fst M1] * Mphi l (snd M1) == P2@l -> P1@l == P3@l. + Lemma PNSubst_ok n P1 cM1 P2 l P3 : + PNSubst P1 cM1 P2 n = Some P3 -> + [fst cM1] * (snd cM1)@@l == P2@l -> P1@l == P3@l. Proof. - intros n P2 (cc, M1) P3 l P4; unfold PNSubst. - generalize (fun P4 => @POneSubst_ok P2 (cc,M1) P3 P4 l); - case (POneSubst P2 (cc,M1) P3); [idtac | intros; discriminate]. - intros P5 H1; case n; try (intros; discriminate). - intros n1 H2; injection H2; intros; subst. - rewrite <- PNSubst1_ok; auto. + unfold PNSubst. + assert (H := POneSubst_ok P1 cM1 P2); destruct POneSubst; try discriminate. + destruct n; inversion_clear 1. + intros. rewrite <- PNSubst1_ok; auto. Qed. - Fixpoint MPcond (LM1: list (C * Mon * Pol)) (l: list R) {struct LM1} : Prop := - match LM1 with - cons (M1,P2) LM2 => ([fst M1] * Mphi l (snd M1) == P2@l) /\ (MPcond LM2 l) - | _ => True - end. + Fixpoint MPcond (LM1: list (C * Mon * Pol)) (l: list R) : Prop := + match LM1 with + | (M1,P2) :: LM2 => ([fst M1] * (snd M1)@@l == P2@l) /\ MPcond LM2 l + | _ => True + end. - Lemma PSubstL1_ok: forall n LM1 P1 l, - MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l. + Lemma PSubstL1_ok n LM1 P1 l : + MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l. Proof. - intros n LM1; elim LM1; simpl; auto. - intros; rsimpl. - intros (M2,P2) LM2 Hrec P3 l [H H1]. - rewrite <- Hrec; auto. - apply PNSubst1_ok; auto. + revert P1; induction LM1 as [|(M2,P2) LM2 IH]; simpl; intros. + - reflexivity. + - rewrite <- IH by intuition. now apply PNSubst1_ok. Qed. - Lemma PSubstL_ok: forall n LM1 P1 P2 l, - PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l. + Lemma PSubstL_ok n LM1 P1 P2 l : + PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l. Proof. - intros n LM1; elim LM1; simpl; auto. - intros; discriminate. - intros (M2,P2) LM2 Hrec P3 P4 l. - generalize (PNSubst_ok n P3 M2 P2); case (PNSubst P3 M2 P2 n). - intros P5 H0 H1 [H2 H3]; injection H1; intros; subst. - rewrite <- PSubstL1_ok; auto. - intros l1 H [H1 H2]; auto. + revert P1. induction LM1 as [|(M2,P2') LM2 IH]; simpl; intros. + - discriminate. + - assert (H':=PNSubst_ok n P3 M2 P2'). destruct PNSubst. + * injection H; intros <-. rewrite <- PSubstL1_ok; intuition. + * now apply IH. Qed. - Lemma PNSubstL_ok: forall m n LM1 P1 l, - MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l. + Lemma PNSubstL_ok m n LM1 P1 l : + MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l. Proof. - intros m; elim m; simpl; auto. - intros n LM1 P2 l H; generalize (fun P3 => @PSubstL_ok n LM1 P2 P3 l); - case (PSubstL P2 LM1 n); intros; rsimpl; auto. - intros m1 Hrec n LM1 P2 l H. - generalize (fun P3 => @PSubstL_ok n LM1 P2 P3 l); - case (PSubstL P2 LM1 n); intros; rsimpl; auto. - rewrite <- Hrec; auto. + revert LM1 P1. induction m; simpl; intros; + assert (H' := PSubstL_ok n LM1 P2); destruct PSubstL; + auto; try reflexivity. + rewrite <- IHm; auto. Qed. (** Definition of polynomial expressions *) @@ -1190,58 +935,22 @@ Strategy expand [PEeval]. (** Correctness proofs *) - Lemma mkX_ok : forall p l, nth 0 p l == (mk_X p) @ l. + Lemma mkX_ok p l : nth 0 p l == (mk_X p) @ l. Proof. destruct p;simpl;intros;Esimpl;trivial. - rewrite <-jump_tl;rewrite nth_jump;rrefl. - rewrite <- nth_jump. - rewrite nth_Pdouble_minus_one;rrefl. + - now rewrite <-jump_tl, nth_jump. + - now rewrite <- nth_jump, nth_pred_double. Qed. - Ltac Esimpl3 := - repeat match goal with - | |- context [(?P1 ++ ?P2)@?l] => rewrite (Padd_ok P2 P1 l) - | |- context [(?P1 -- ?P2)@?l] => rewrite (Psub_ok P2 P1 l) - end;Esimpl2;try rrefl;try apply (ARadd_comm ARth). - -(* Power using the chinise algorithm *) -(*Section POWER. - Variable subst_l : Pol -> Pol. - Fixpoint Ppow_pos (P:Pol) (p:positive){struct p} : Pol := - match p with - | xH => P - | xO p => subst_l (Psquare (Ppow_pos P p)) - | xI p => subst_l (Pmul P (Psquare (Ppow_pos P p))) - end. - - Definition Ppow_N P n := - match n with - | N0 => P1 - | Npos p => Ppow_pos P p - end. - - Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) -> - forall P p, (Ppow_pos P p)@l == (pow_pos Pmul P p)@l. - Proof. - intros l subst_l_ok P. - induction p;simpl;intros;try rrefl;try rewrite subst_l_ok. - repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl. - repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl. - Qed. - - Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) -> - forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. - Proof. destruct n;simpl. rrefl. apply Ppow_pos_ok. trivial. Qed. - - End POWER. *) + Hint Rewrite Padd_ok Psub_ok : Esimpl. Section POWER. Variable subst_l : Pol -> Pol. - Fixpoint Ppow_pos (res P:Pol) (p:positive){struct p} : Pol := + Fixpoint Ppow_pos (res P:Pol) (p:positive) : Pol := match p with - | xH => subst_l (Pmul res P) + | xH => subst_l (res ** P) | xO p => Ppow_pos (Ppow_pos res P p) P p - | xI p => subst_l (Pmul (Ppow_pos (Ppow_pos res P p) P p) P) + | xI p => subst_l ((Ppow_pos (Ppow_pos res P p) P p) ** P) end. Definition Ppow_N P n := @@ -1250,17 +959,23 @@ Section POWER. | Npos p => Ppow_pos P1 P p end. - Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) -> - forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l. + Lemma Ppow_pos_ok l : + (forall P, subst_l P@l == P@l) -> + forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l. Proof. - intros l subst_l_ok res P p. generalize res;clear res. - induction p;simpl;intros;try rewrite subst_l_ok; repeat rewrite Pmul_ok;repeat rewrite IHp. - rsimpl. mul_push (P@l);rsimpl. rsimpl. rrefl. + intros subst_l_ok res P p. revert res. + induction p;simpl;intros; rewrite ?subst_l_ok, ?Pmul_ok, ?IHp; + mul_permut. Qed. - Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) -> - forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. - Proof. destruct n;simpl. rrefl. rewrite Ppow_pos_ok by trivial. Esimpl. Qed. + Lemma Ppow_N_ok l : + (forall P, subst_l P@l == P@l) -> + forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l. + Proof. + destruct n;simpl. + - reflexivity. + - rewrite Ppow_pos_ok by trivial. Esimpl. + Qed. End POWER. @@ -1277,69 +992,66 @@ Section POWER. match pe with | PEc c => Pc c | PEX j => mk_X j - | PEadd (PEopp pe1) pe2 => Psub (norm_aux pe2) (norm_aux pe1) - | PEadd pe1 (PEopp pe2) => - Psub (norm_aux pe1) (norm_aux pe2) - | PEadd pe1 pe2 => Padd (norm_aux pe1) (norm_aux pe2) - | PEsub pe1 pe2 => Psub (norm_aux pe1) (norm_aux pe2) - | PEmul pe1 pe2 => Pmul (norm_aux pe1) (norm_aux pe2) - | PEopp pe1 => Popp (norm_aux pe1) + | PEadd (PEopp pe1) pe2 => (norm_aux pe2) -- (norm_aux pe1) + | PEadd pe1 (PEopp pe2) => (norm_aux pe1) -- (norm_aux pe2) + | PEadd pe1 pe2 => (norm_aux pe1) ++ (norm_aux pe2) + | PEsub pe1 pe2 => (norm_aux pe1) -- (norm_aux pe2) + | PEmul pe1 pe2 => (norm_aux pe1) ** (norm_aux pe2) + | PEopp pe1 => -- (norm_aux pe1) | PEpow pe1 n => Ppow_N (fun p => p) (norm_aux pe1) n end. Definition norm_subst pe := subst_l (norm_aux pe). - (* - Fixpoint norm_subst (pe:PExpr) : Pol := + (** Internally, [norm_aux] is expanded in a large number of cases. + To speed-up proofs, we use an alternative definition. *) + + Definition get_PEopp pe := match pe with - | PEc c => Pc c - | PEX j => subst_l (mk_X j) - | PEadd (PEopp pe1) pe2 => Psub (norm_subst pe2) (norm_subst pe1) - | PEadd pe1 (PEopp pe2) => - Psub (norm_subst pe1) (norm_subst pe2) - | PEadd pe1 pe2 => Padd (norm_subst pe1) (norm_subst pe2) - | PEsub pe1 pe2 => Psub (norm_subst pe1) (norm_subst pe2) - | PEmul pe1 pe2 => Pmul_subst (norm_subst pe1) (norm_subst pe2) - | PEopp pe1 => Popp (norm_subst pe1) - | PEpow pe1 n => Ppow_subst (norm_subst pe1) n + | PEopp pe' => Some pe' + | _ => None end. - Lemma norm_subst_spec : - forall l pe, MPcond lmp l -> - PEeval l pe == (norm_subst pe)@l. + Lemma norm_aux_PEadd pe1 pe2 : + norm_aux (PEadd pe1 pe2) = + match get_PEopp pe1, get_PEopp pe2 with + | Some pe1', _ => (norm_aux pe2) -- (norm_aux pe1') + | None, Some pe2' => (norm_aux pe1) -- (norm_aux pe2') + | None, None => (norm_aux pe1) ++ (norm_aux pe2) + end. Proof. - intros;assert (subst_l_ok:forall P, (subst_l P)@l == P@l). - unfold subst_l;intros. - rewrite <- PNSubstL_ok;trivial. rrefl. - assert (Pms_ok:forall P1 P2, (Pmul_subst P1 P2)@l == P1@l*P2@l). - intros;unfold Pmul_subst;rewrite subst_l_ok;rewrite Pmul_ok;rrefl. - induction pe;simpl;Esimpl3. - rewrite subst_l_ok;apply mkX_ok. - rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3. - rewrite IHpe1;rewrite IHpe2;rrefl. - rewrite Pms_ok;rewrite IHpe1;rewrite IHpe2;rrefl. - rewrite IHpe;rrefl. - unfold Ppow_subst. rewrite Ppow_N_ok. trivial. - rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3. - induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok; - repeat rewrite Pmul_ok;rrefl. + simpl (norm_aux (PEadd _ _)). + destruct pe1; [ | | | | | reflexivity | ]; + destruct pe2; simpl get_PEopp; reflexivity. Qed. -*) - Lemma norm_aux_spec : - forall l pe, MPcond lmp l -> - PEeval l pe == (norm_aux pe)@l. + + Lemma norm_aux_PEopp pe : + match get_PEopp pe with + | Some pe' => norm_aux pe = -- (norm_aux pe') + | None => True + end. + Proof. + now destruct pe. + Qed. + + Lemma norm_aux_spec l pe : + PEeval l pe == (norm_aux pe)@l. Proof. intros. - induction pe;simpl;Esimpl3. - apply mkX_ok. - rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3. - rewrite IHpe1;rewrite IHpe2;rrefl. - rewrite IHpe1;rewrite IHpe2. rewrite Pmul_ok. rrefl. - rewrite IHpe;rrefl. - rewrite Ppow_N_ok by (intros;rrefl). - rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3. - induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok; - repeat rewrite Pmul_ok;rrefl. + induction pe. + - reflexivity. + - apply mkX_ok. + - simpl PEeval. rewrite IHpe1, IHpe2. + assert (H1 := norm_aux_PEopp pe1). + assert (H2 := norm_aux_PEopp pe2). + rewrite norm_aux_PEadd. + do 2 destruct get_PEopp; rewrite ?H1, ?H2; Esimpl; add_permut. + - simpl. rewrite IHpe1, IHpe2. Esimpl. + - simpl. rewrite IHpe1, IHpe2. now rewrite Pmul_ok. + - simpl. rewrite IHpe. Esimpl. + - simpl. rewrite Ppow_N_ok by reflexivity. + rewrite pow_th.(rpow_pow_N). destruct n0; simpl; Esimpl. + induction p;simpl; now rewrite ?IHp, ?IHpe, ?Pms_ok, ?Pmul_ok. Qed. Lemma norm_subst_spec : @@ -1347,7 +1059,7 @@ Section POWER. PEeval l pe == (norm_subst pe)@l. Proof. intros;unfold norm_subst. - unfold subst_l;rewrite <- PNSubstL_ok;trivial. apply norm_aux_spec. trivial. + unfold subst_l;rewrite <- PNSubstL_ok;trivial. apply norm_aux_spec. Qed. End NORM_SUBST_REC. @@ -1514,27 +1226,27 @@ Section POWER. (rP:R) (P:Pol) (fv:list R) (n:N) (lm:list (R*positive)) {struct P} : R := match P with | Pc c => - let lm := add_pow_list (hd 0 fv) n lm in + let lm := add_pow_list (hd fv) n lm in mkadd_mult rP c lm | Pinj j Q => - add_mult_dev rP Q (jump j fv) N0 (add_pow_list (hd 0 fv) n lm) + add_mult_dev rP Q (jump j fv) N0 (add_pow_list (hd fv) n lm) | PX P i Q => - let rP := add_mult_dev rP P fv (Nplus (Npos i) n) lm in + let rP := add_mult_dev rP P fv (N.add (Npos i) n) lm in if Q ?== P0 then rP - else add_mult_dev rP Q (tail fv) N0 (add_pow_list (hd 0 fv) n lm) + else add_mult_dev rP Q (tail fv) N0 (add_pow_list (hd fv) n lm) end. Fixpoint mult_dev (P:Pol) (fv : list R) (n:N) (lm:list (R*positive)) {struct P} : R := (* P@l * (hd 0 l)^n * lm *) match P with - | Pc c => mkmult_c c (add_pow_list (hd 0 fv) n lm) - | Pinj j Q => mult_dev Q (jump j fv) N0 (add_pow_list (hd 0 fv) n lm) + | Pc c => mkmult_c c (add_pow_list (hd fv) n lm) + | Pinj j Q => mult_dev Q (jump j fv) N0 (add_pow_list (hd fv) n lm) | PX P i Q => - let rP := mult_dev P fv (Nplus (Npos i) n) lm in + let rP := mult_dev P fv (N.add (Npos i) n) lm in if Q ?== P0 then rP else - let lmq := add_pow_list (hd 0 fv) n lm in + let lmq := add_pow_list (hd fv) n lm in add_mult_dev rP Q (tail fv) N0 lmq end. @@ -1575,7 +1287,7 @@ Section POWER. (forall l lr : list (R * positive), r_list_pow (rev_append l lr) == r_list_pow lr * r_list_pow l). induction l;intros;simpl;Esimpl. destruct a;rewrite IHl;Esimpl. - rewrite (ARmul_comm ARth (pow_pos rmul r p)). rrefl. + rewrite (ARmul_comm ARth (pow_pos rmul r p)). reflexivity. intros;unfold rev'. rewrite H;simpl;Esimpl. Qed. @@ -1617,11 +1329,11 @@ Qed. Qed. Lemma add_mult_dev_ok : forall P rP fv n lm, - add_mult_dev rP P fv n lm == rP + P@fv*pow_N rI rmul (hd 0 fv) n * r_list_pow lm. + add_mult_dev rP P fv n lm == rP + P@fv*pow_N rI rmul (hd fv) n * r_list_pow lm. Proof. induction P;simpl;intros. - rewrite mkadd_mult_ok. rewrite add_pow_list_ok; Esimpl. - rewrite IHP. simpl. rewrite add_pow_list_ok; Esimpl. + rewrite mkadd_mult_ok. rewrite add_pow_list_ok; Esimpl. + rewrite IHP. simpl. rewrite add_pow_list_ok; Esimpl. change (match P3 with | Pc c => c ?=! cO | Pinj _ _ => false @@ -1630,17 +1342,19 @@ Qed. change match n with | N0 => Npos p | Npos q => Npos (p + q) - end with (Nplus (Npos p) n);trivial. + end with (N.add (Npos p) n);trivial. assert (H := Peq_ok P3 P0). destruct (P3 ?== P0). - rewrite (H (refl_equal true)). - rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl. - rewrite IHP2. - rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl. + rewrite (H eq_refl). + rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. + add_permut. mul_permut. + rewrite IHP2. + rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. + add_permut. mul_permut. Qed. Lemma mult_dev_ok : forall P fv n lm, - mult_dev P fv n lm == P@fv * pow_N rI rmul (hd 0 fv) n * r_list_pow lm. + mult_dev P fv n lm == P@fv * pow_N rI rmul (hd fv) n * r_list_pow lm. Proof. induction P;simpl;intros;Esimpl. rewrite mkmult_c_ok;rewrite add_pow_list_ok;Esimpl. @@ -1653,13 +1367,15 @@ Qed. change match n with | N0 => Npos p | Npos q => Npos (p + q) - end with (Nplus (Npos p) n);trivial. + end with (N.add (Npos p) n);trivial. assert (H := Peq_ok P3 P0). destruct (P3 ?== P0). - rewrite (H (refl_equal true)). - rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl. + rewrite (H eq_refl). + rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. + mul_permut. rewrite add_mult_dev_ok. rewrite IHP1; rewrite add_pow_list_ok. - destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl. + destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl. + add_permut; mul_permut. Qed. Lemma Pphi_avoid_ok : forall P fv, Pphi_avoid fv P == P@fv. @@ -1676,18 +1392,18 @@ Qed. let mkmult_pow r x p := rmul r (mkpow x p) in Pphi_avoid mkpow mkopp_pow mkmult_pow. - Lemma local_mkpow_ok : - forall (r : R) (p : positive), + Lemma local_mkpow_ok r p : match p with | xI _ => rpow r (Cp_phi (Npos p)) | xO _ => rpow r (Cp_phi (Npos p)) | 1 => r end == pow_pos rmul r p. - Proof. intros r p;destruct p;try rewrite pow_th.(rpow_pow_N);reflexivity. Qed. + Proof. destruct p; now rewrite ?pow_th.(rpow_pow_N). Qed. Lemma Pphi_pow_ok : forall P fv, Pphi_pow fv P == P@fv. Proof. - unfold Pphi_pow;intros;apply Pphi_avoid_ok;intros;try rewrite local_mkpow_ok;rrefl. + unfold Pphi_pow;intros;apply Pphi_avoid_ok;intros; + now rewrite ?local_mkpow_ok. Qed. Lemma ring_rw_pow_correct : forall n lH l, @@ -1697,7 +1413,7 @@ Qed. PEeval l pe == Pphi_pow l npe. Proof. intros n lH l H1 lmp Heq1 pe npe Heq2. - rewrite Pphi_pow_ok. rewrite <- Heq2;rewrite <- Heq1. + rewrite Pphi_pow_ok, <- Heq2, <- Heq1. apply norm_subst_ok. trivial. Qed. @@ -1711,58 +1427,48 @@ Qed. Definition mkpow x p := match p with | xH => x - | xO p => mkmult_pow x x (Pdouble_minus_one p) + | xO p => mkmult_pow x x (Pos.pred_double p) | xI p => mkmult_pow x x (xO p) end. Definition mkopp_pow x p := match p with | xH => -x - | xO p => mkmult_pow (-x) x (Pdouble_minus_one p) + | xO p => mkmult_pow (-x) x (Pos.pred_double p) | xI p => mkmult_pow (-x) x (xO p) end. Definition Pphi_dev := Pphi_avoid mkpow mkopp_pow mkmult_pow. - Lemma mkmult_pow_ok : forall p r x, mkmult_pow r x p == r*pow_pos rmul x p. + Lemma mkmult_pow_ok p r x : mkmult_pow r x p == r * x^p. Proof. - induction p;intros;simpl;Esimpl. - repeat rewrite IHp;Esimpl. - repeat rewrite IHp;Esimpl. + revert r; induction p;intros;simpl;Esimpl;rewrite !IHp;Esimpl. Qed. - Lemma mkpow_ok : forall p x, mkpow x p == pow_pos rmul x p. + Lemma mkpow_ok p x : mkpow x p == x^p. Proof. destruct p;simpl;intros;Esimpl. - repeat rewrite mkmult_pow_ok;Esimpl. - rewrite mkmult_pow_ok;Esimpl. - pattern x at 1;replace x with (pow_pos rmul x 1). - rewrite <- pow_pos_Pplus. - rewrite <- Pplus_one_succ_l. - rewrite Psucc_o_double_minus_one_eq_xO. - simpl;Esimpl. - trivial. + - rewrite !mkmult_pow_ok;Esimpl. + - rewrite mkmult_pow_ok;Esimpl. + change x with (x^1) at 1. + now rewrite <- pow_pos_add, Pos.add_1_r, Pos.succ_pred_double. Qed. - Lemma mkopp_pow_ok : forall p x, mkopp_pow x p == - pow_pos rmul x p. + Lemma mkopp_pow_ok p x : mkopp_pow x p == - x^p. Proof. destruct p;simpl;intros;Esimpl. - repeat rewrite mkmult_pow_ok;Esimpl. - rewrite mkmult_pow_ok;Esimpl. - pattern x at 1;replace x with (pow_pos rmul x 1). - rewrite <- pow_pos_Pplus. - rewrite <- Pplus_one_succ_l. - rewrite Psucc_o_double_minus_one_eq_xO. - simpl;Esimpl. - trivial. + - rewrite !mkmult_pow_ok;Esimpl. + - rewrite mkmult_pow_ok;Esimpl. + change x with (x^1) at 1. + now rewrite <- pow_pos_add, Pos.add_1_r, Pos.succ_pred_double. Qed. Lemma Pphi_dev_ok : forall P fv, Pphi_dev fv P == P@fv. Proof. unfold Pphi_dev;intros;apply Pphi_avoid_ok. - intros;apply mkpow_ok. - intros;apply mkopp_pow_ok. - intros;apply mkmult_pow_ok. + - intros;apply mkpow_ok. + - intros;apply mkopp_pow_ok. + - intros;apply mkmult_pow_ok. Qed. Lemma ring_rw_correct : forall n lH l, @@ -1776,6 +1482,4 @@ Qed. apply norm_subst_ok. trivial. Qed. - End MakeRingPol. - -- cgit v1.2.3