summaryrefslogtreecommitdiff
path: root/plugins/setoid_ring
diff options
context:
space:
mode:
authorGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
committerGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
commit5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 (patch)
tree631ad791a7685edafeb1fb2e8faeedc8379318ae /plugins/setoid_ring
parentda178a880e3ace820b41d38b191d3785b82991f5 (diff)
Imported Upstream snapshot 8.3~beta0+13298
Diffstat (limited to 'plugins/setoid_ring')
-rw-r--r--plugins/setoid_ring/ArithRing.v60
-rw-r--r--plugins/setoid_ring/BinList.v93
-rw-r--r--plugins/setoid_ring/Field.v10
-rw-r--r--plugins/setoid_ring/Field_tac.v571
-rw-r--r--plugins/setoid_ring/Field_theory.v1946
-rw-r--r--plugins/setoid_ring/InitialRing.v908
-rw-r--r--plugins/setoid_ring/NArithRing.v21
-rw-r--r--plugins/setoid_ring/RealField.v134
-rw-r--r--plugins/setoid_ring/Ring.v44
-rw-r--r--plugins/setoid_ring/Ring_base.v17
-rw-r--r--plugins/setoid_ring/Ring_equiv.v74
-rw-r--r--plugins/setoid_ring/Ring_polynom.v1781
-rw-r--r--plugins/setoid_ring/Ring_tac.v434
-rw-r--r--plugins/setoid_ring/Ring_theory.v608
-rw-r--r--plugins/setoid_ring/ZArithRing.v60
-rw-r--r--plugins/setoid_ring/newring.ml41164
-rw-r--r--plugins/setoid_ring/newring_plugin.mllib2
-rw-r--r--plugins/setoid_ring/vo.itarget15
18 files changed, 7942 insertions, 0 deletions
diff --git a/plugins/setoid_ring/ArithRing.v b/plugins/setoid_ring/ArithRing.v
new file mode 100644
index 00000000..e5a4c8d1
--- /dev/null
+++ b/plugins/setoid_ring/ArithRing.v
@@ -0,0 +1,60 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Mult.
+Require Import BinNat.
+Require Import Nnat.
+Require Export Ring.
+Set Implicit Arguments.
+
+Lemma natSRth : semi_ring_theory O (S O) plus mult (@eq nat).
+ Proof.
+ constructor. exact plus_0_l. exact plus_comm. exact plus_assoc.
+ exact mult_1_l. exact mult_0_l. exact mult_comm. exact mult_assoc.
+ exact mult_plus_distr_r.
+ Qed.
+
+Lemma nat_morph_N :
+ semi_morph 0 1 plus mult (eq (A:=nat))
+ 0%N 1%N Nplus Nmult Neq_bool nat_of_N.
+Proof.
+ constructor;trivial.
+ exact nat_of_Nplus.
+ exact nat_of_Nmult.
+ intros x y H;rewrite (Neq_bool_ok _ _ H);trivial.
+Qed.
+
+Ltac natcst t :=
+ match isnatcst t with
+ true => constr:(N_of_nat t)
+ | _ => constr:InitialRing.NotConstant
+ end.
+
+Ltac Ss_to_add f acc :=
+ match f with
+ | S ?f1 => Ss_to_add f1 (S acc)
+ | _ => constr:(acc + f)%nat
+ end.
+
+Ltac natprering :=
+ match goal with
+ |- context C [S ?p] =>
+ match p with
+ O => fail 1 (* avoid replacing 1 with 1+0 ! *)
+ | p => match isnatcst p with
+ | true => fail 1
+ | false => let v := Ss_to_add p (S 0) in
+ fold v; natprering
+ end
+ end
+ | _ => idtac
+ end.
+
+Add Ring natr : natSRth
+ (morphism nat_morph_N, constants [natcst], preprocess [natprering]).
+
diff --git a/plugins/setoid_ring/BinList.v b/plugins/setoid_ring/BinList.v
new file mode 100644
index 00000000..d403c9ef
--- /dev/null
+++ b/plugins/setoid_ring/BinList.v
@@ -0,0 +1,93 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Set Implicit Arguments.
+Require Import BinPos.
+Require Export List.
+Require Export ListTactics.
+Open Local Scope positive_scope.
+
+Section MakeBinList.
+ Variable A : Type.
+ Variable default : A.
+
+ Fixpoint jump (p:positive) (l:list A) {struct p} : list A :=
+ match p with
+ | xH => tail l
+ | xO p => jump p (jump p l)
+ | xI p => jump p (jump p (tail l))
+ end.
+
+ Fixpoint nth (p:positive) (l:list A) {struct p} : A:=
+ match p with
+ | xH => hd default l
+ | xO p => nth p (jump p l)
+ | xI p => nth p (jump p (tail l))
+ end.
+
+ Lemma jump_tl : forall j l, tail (jump j l) = jump j (tail l).
+ Proof.
+ induction j;simpl;intros.
+ repeat rewrite IHj;trivial.
+ repeat rewrite IHj;trivial.
+ trivial.
+ Qed.
+
+ Lemma jump_Psucc : forall j l,
+ (jump (Psucc j) l) = (jump 1 (jump j l)).
+ Proof.
+ induction j;simpl;intros.
+ repeat rewrite IHj;simpl;repeat rewrite jump_tl;trivial.
+ repeat rewrite jump_tl;trivial.
+ trivial.
+ Qed.
+
+ Lemma jump_Pplus : forall i j l,
+ (jump (i + j) l) = (jump i (jump j l)).
+ Proof.
+ induction i;intros.
+ rewrite xI_succ_xO;rewrite Pplus_one_succ_r.
+ rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
+ repeat rewrite IHi.
+ rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite jump_Psucc;trivial.
+ rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
+ repeat rewrite IHi;trivial.
+ rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite jump_Psucc;trivial.
+ Qed.
+
+ Lemma jump_Pdouble_minus_one : forall i l,
+ (jump (Pdouble_minus_one i) (tail l)) = (jump i (jump i l)).
+ Proof.
+ induction i;intros;simpl.
+ repeat rewrite jump_tl;trivial.
+ rewrite IHi. do 2 rewrite <- jump_tl;rewrite IHi;trivial.
+ trivial.
+ Qed.
+
+
+ Lemma nth_jump : forall p l, nth p (tail l) = hd default (jump p l).
+ Proof.
+ induction p;simpl;intros.
+ rewrite <-jump_tl;rewrite IHp;trivial.
+ rewrite <-jump_tl;rewrite IHp;trivial.
+ trivial.
+ Qed.
+
+ Lemma nth_Pdouble_minus_one :
+ forall p l, nth (Pdouble_minus_one p) (tail l) = nth p (jump p l).
+ Proof.
+ induction p;simpl;intros.
+ repeat rewrite jump_tl;trivial.
+ rewrite jump_Pdouble_minus_one.
+ repeat rewrite <- jump_tl;rewrite IHp;trivial.
+ trivial.
+ Qed.
+
+End MakeBinList.
+
+
diff --git a/plugins/setoid_ring/Field.v b/plugins/setoid_ring/Field.v
new file mode 100644
index 00000000..a944ba5f
--- /dev/null
+++ b/plugins/setoid_ring/Field.v
@@ -0,0 +1,10 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Export Field_theory.
+Require Export Field_tac.
diff --git a/plugins/setoid_ring/Field_tac.v b/plugins/setoid_ring/Field_tac.v
new file mode 100644
index 00000000..9d82d1fd
--- /dev/null
+++ b/plugins/setoid_ring/Field_tac.v
@@ -0,0 +1,571 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Ring_tac BinList Ring_polynom InitialRing.
+Require Export Field_theory.
+
+ (* syntaxification *)
+ Ltac mkFieldexpr C Cst CstPow radd rmul rsub ropp rdiv rinv rpow t fv :=
+ let rec mkP t :=
+ let f :=
+ match Cst t with
+ | InitialRing.NotConstant =>
+ match t with
+ | (radd ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(FEadd e1 e2)
+ | (rmul ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(FEmul e1 e2)
+ | (rsub ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(FEsub e1 e2)
+ | (ropp ?t1) =>
+ fun _ => let e1 := mkP t1 in constr:(FEopp e1)
+ | (rdiv ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(FEdiv e1 e2)
+ | (rinv ?t1) =>
+ fun _ => let e1 := mkP t1 in constr:(FEinv e1)
+ | (rpow ?t1 ?n) =>
+ match CstPow n with
+ | InitialRing.NotConstant =>
+ fun _ =>
+ let p := Find_at t fv in
+ constr:(@FEX C p)
+ | ?c => fun _ => let e1 := mkP t1 in constr:(FEpow e1 c)
+ end
+ | _ =>
+ fun _ =>
+ let p := Find_at t fv in
+ constr:(@FEX C p)
+ end
+ | ?c => fun _ => constr:(FEc c)
+ end in
+ f ()
+ in mkP t.
+
+Ltac FFV Cst CstPow add mul sub opp div inv pow t fv :=
+ let rec TFV t fv :=
+ match Cst t with
+ | InitialRing.NotConstant =>
+ match t with
+ | (add ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv)
+ | (mul ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv)
+ | (sub ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv)
+ | (opp ?t1) => TFV t1 fv
+ | (div ?t1 ?t2) => TFV t2 ltac:(TFV t1 fv)
+ | (inv ?t1) => TFV t1 fv
+ | (pow ?t1 ?n) =>
+ match CstPow n with
+ | InitialRing.NotConstant =>
+ AddFvTail t fv
+ | _ => TFV t1 fv
+ end
+ | _ => AddFvTail t fv
+ end
+ | _ => fv
+ end
+ in TFV t fv.
+
+(* packaging the field structure *)
+
+(* TODO: inline PackField into field_lookup *)
+Ltac PackField F req Cst_tac Pow_tac L1 L2 L3 L4 cond_ok pre post :=
+ let FLD :=
+ match type of L1 with
+ | context [req (@FEeval ?R ?rO ?radd ?rmul ?rsub ?ropp ?rdiv ?rinv
+ ?C ?phi ?Cpow ?Cp_phi ?rpow _ _) _ ] =>
+ (fun proj =>
+ proj Cst_tac Pow_tac pre post
+ req radd rmul rsub ropp rdiv rinv rpow C L1 L2 L3 L4 cond_ok)
+ | _ => fail 1 "field anomaly: bad correctness lemma (parse)"
+ end in
+ F FLD.
+
+Ltac get_FldPre FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ pre).
+
+Ltac get_FldPost FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ post).
+
+Ltac get_L1 FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ L1).
+
+Ltac get_SimplifyEqLemma FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ L2).
+
+Ltac get_SimplifyLemma FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ L3).
+
+Ltac get_L4 FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ L4).
+
+Ltac get_CondLemma FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ cond_ok).
+
+Ltac get_FldEq FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ req).
+
+Ltac get_FldCarrier FLD :=
+ let req := get_FldEq FLD in
+ relation_carrier req.
+
+Ltac get_RingFV FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ FV Cst_tac Pow_tac radd rmul rsub ropp rpow).
+
+Ltac get_FFV FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ FFV Cst_tac Pow_tac radd rmul rsub ropp rdiv rinv rpow).
+
+Ltac get_RingMeta FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ mkPolexpr C Cst_tac Pow_tac radd rmul rsub ropp rpow).
+
+Ltac get_Meta FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ mkFieldexpr C Cst_tac Pow_tac radd rmul rsub ropp rdiv rinv rpow).
+
+Ltac get_Hyp_tac FLD :=
+ FLD ltac:
+ (fun Cst_tac Pow_tac pre post req radd rmul rsub ropp rdiv rinv rpow C
+ L1 L2 L3 L4 cond_ok =>
+ let mkPol := mkPolexpr C Cst_tac Pow_tac radd rmul rsub ropp rpow in
+ fun fv lH => mkHyp_tac C req ltac:(fun t => mkPol t fv) lH).
+
+Ltac get_FEeval FLD :=
+ let L1 := get_L1 FLD in
+ match type of L1 with
+ | context
+ [(@FEeval
+ ?R ?r0 ?add ?mul ?sub ?opp ?div ?inv ?C ?phi ?Cpow ?powphi ?pow _ _)] =>
+ constr:(@FEeval R r0 add mul sub opp div inv C phi Cpow powphi pow)
+ | _ => fail 1 "field anomaly: bad correctness lemma (get_FEeval)"
+ end.
+
+(* simplifying the non-zero condition... *)
+
+Ltac fold_field_cond req :=
+ let rec fold_concl t :=
+ match t with
+ ?x /\ ?y =>
+ let fx := fold_concl x in let fy := fold_concl y in constr:(fx/\fy)
+ | req ?x ?y -> False => constr:(~ req x y)
+ | _ => t
+ end in
+ let ft := fold_concl Get_goal in
+ change ft.
+
+Ltac simpl_PCond FLD :=
+ let req := get_FldEq FLD in
+ let lemma := get_CondLemma FLD in
+ try apply lemma;
+ protect_fv "field_cond";
+ fold_field_cond req;
+ try exact I.
+
+Ltac simpl_PCond_BEURK FLD :=
+ let req := get_FldEq FLD in
+ let lemma := get_CondLemma FLD in
+ apply lemma;
+ protect_fv "field_cond";
+ fold_field_cond req.
+
+(* Rewriting (field_simplify) *)
+Ltac Field_norm_gen f n FLD lH rl :=
+ let mkFV := get_RingFV FLD in
+ let mkFFV := get_FFV FLD in
+ let mkFE := get_Meta FLD in
+ let fv0 := FV_hypo_tac mkFV ltac:(get_FldEq FLD) lH in
+ let lemma_tac fv kont :=
+ let lemma := get_SimplifyLemma FLD in
+ (* reify equations of the context *)
+ let lpe := get_Hyp_tac FLD fv lH in
+ let vlpe := fresh "hyps" in
+ pose (vlpe := lpe);
+ let prh := proofHyp_tac lH in
+ (* compute the normal form of the reified hyps *)
+ let vlmp := fresh "hyps'" in
+ let vlmp_eq := fresh "hyps_eq" in
+ let mk_monpol := get_MonPol lemma in
+ compute_assertion vlmp_eq vlmp (mk_monpol vlpe);
+ (* partially instantiate the lemma *)
+ let lem := fresh "f_rw_lemma" in
+ (assert (lem := lemma n vlpe fv prh vlmp vlmp_eq)
+ || fail "type error when building the rewriting lemma");
+ (* continuation will call main_tac for all reified terms *)
+ kont lem;
+ (* at the end, cleanup *)
+ (clear lem vlmp_eq vlmp vlpe||idtac"Field_norm_gen:cleanup failed") in
+ (* each instance of the lemma is simplified then passed to f *)
+ let main_tac H := protect_fv "field" in H; f H in
+ (* generate and use equations for each expression *)
+ ReflexiveRewriteTactic mkFFV mkFE lemma_tac main_tac fv0 rl;
+ try simpl_PCond FLD.
+
+Ltac Field_simplify_gen f FLD lH rl :=
+ get_FldPre FLD ();
+ Field_norm_gen f ring_subst_niter FLD lH rl;
+ get_FldPost FLD ().
+
+Ltac Field_simplify :=
+ Field_simplify_gen ltac:(fun H => rewrite H).
+
+Tactic Notation (at level 0) "field_simplify" constr_list(rl) :=
+ let G := Get_goal in
+ field_lookup (PackField Field_simplify) [] rl G.
+
+Tactic Notation (at level 0)
+ "field_simplify" "[" constr_list(lH) "]" constr_list(rl) :=
+ let G := Get_goal in
+ field_lookup (PackField Field_simplify) [lH] rl G.
+
+Tactic Notation "field_simplify" constr_list(rl) "in" hyp(H):=
+ let G := Get_goal in
+ let t := type of H in
+ let g := fresh "goal" in
+ set (g:= G);
+ revert H;
+ field_lookup (PackField Field_simplify) [] rl t;
+ intro H;
+ unfold g;clear g.
+
+Tactic Notation "field_simplify"
+ "["constr_list(lH) "]" constr_list(rl) "in" hyp(H):=
+ let G := Get_goal in
+ let t := type of H in
+ let g := fresh "goal" in
+ set (g:= G);
+ revert H;
+ field_lookup (PackField Field_simplify) [lH] rl t;
+ intro H;
+ unfold g;clear g.
+
+(*
+Ltac Field_simplify_in hyp:=
+ Field_simplify_gen ltac:(fun H => rewrite H in hyp).
+
+Tactic Notation (at level 0)
+ "field_simplify" constr_list(rl) "in" hyp(h) :=
+ let t := type of h in
+ field_lookup (Field_simplify_in h) [] rl t.
+
+Tactic Notation (at level 0)
+ "field_simplify" "[" constr_list(lH) "]" constr_list(rl) "in" hyp(h) :=
+ let t := type of h in
+ field_lookup (Field_simplify_in h) [lH] rl t.
+*)
+
+(** Generic tactic for solving equations *)
+
+Ltac Field_Scheme Simpl_tac n lemma FLD lH :=
+ let req := get_FldEq FLD in
+ let mkFV := get_RingFV FLD in
+ let mkFFV := get_FFV FLD in
+ let mkFE := get_Meta FLD in
+ let Main_eq t1 t2 :=
+ let fv := FV_hypo_tac mkFV req lH in
+ let fv := mkFFV t1 fv in
+ let fv := mkFFV t2 fv in
+ let lpe := get_Hyp_tac FLD fv lH in
+ let prh := proofHyp_tac lH in
+ let vlpe := fresh "list_hyp" in
+ let fe1 := mkFE t1 fv in
+ let fe2 := mkFE t2 fv in
+ pose (vlpe := lpe);
+ let nlemma := fresh "field_lemma" in
+ (assert (nlemma := lemma n fv vlpe fe1 fe2 prh)
+ || fail "field anomaly:failed to build lemma");
+ ProveLemmaHyps nlemma
+ ltac:(fun ilemma =>
+ apply ilemma
+ || fail "field anomaly: failed in applying lemma";
+ [ Simpl_tac | simpl_PCond FLD]);
+ clear nlemma;
+ subst vlpe in
+ OnEquation req Main_eq.
+
+(* solve completely a field equation, leaving non-zero conditions to be
+ proved (field) *)
+
+Ltac FIELD FLD lH rl :=
+ let Simpl := vm_compute; reflexivity || fail "not a valid field equation" in
+ let lemma := get_L1 FLD in
+ get_FldPre FLD ();
+ Field_Scheme Simpl Ring_tac.ring_subst_niter lemma FLD lH;
+ try exact I;
+ get_FldPost FLD().
+
+Tactic Notation (at level 0) "field" :=
+ let G := Get_goal in
+ field_lookup (PackField FIELD) [] G.
+
+Tactic Notation (at level 0) "field" "[" constr_list(lH) "]" :=
+ let G := Get_goal in
+ field_lookup (PackField FIELD) [lH] G.
+
+(* transforms a field equation to an equivalent (simplified) ring equation,
+ and leaves non-zero conditions to be proved (field_simplify_eq) *)
+Ltac FIELD_SIMPL FLD lH rl :=
+ let Simpl := (protect_fv "field") in
+ let lemma := get_SimplifyEqLemma FLD in
+ get_FldPre FLD ();
+ Field_Scheme Simpl Ring_tac.ring_subst_niter lemma FLD lH;
+ get_FldPost FLD ().
+
+Tactic Notation (at level 0) "field_simplify_eq" :=
+ let G := Get_goal in
+ field_lookup (PackField FIELD_SIMPL) [] G.
+
+Tactic Notation (at level 0) "field_simplify_eq" "[" constr_list(lH) "]" :=
+ let G := Get_goal in
+ field_lookup (PackField FIELD_SIMPL) [lH] G.
+
+(* Same as FIELD_SIMPL but in hypothesis *)
+
+Ltac Field_simplify_eq n FLD lH :=
+ let req := get_FldEq FLD in
+ let mkFV := get_RingFV FLD in
+ let mkFFV := get_FFV FLD in
+ let mkFE := get_Meta FLD in
+ let lemma := get_L4 FLD in
+ let hyp := fresh "hyp" in
+ intro hyp;
+ OnEquationHyp req hyp ltac:(fun t1 t2 =>
+ let fv := FV_hypo_tac mkFV req lH in
+ let fv := mkFFV t1 fv in
+ let fv := mkFFV t2 fv in
+ let lpe := get_Hyp_tac FLD fv lH in
+ let prh := proofHyp_tac lH in
+ let fe1 := mkFE t1 fv in
+ let fe2 := mkFE t2 fv in
+ let vlpe := fresh "vlpe" in
+ ProveLemmaHyps (lemma n fv lpe fe1 fe2 prh)
+ ltac:(fun ilemma =>
+ match type of ilemma with
+ | req _ _ -> _ -> ?EQ =>
+ let tmp := fresh "tmp" in
+ assert (tmp : EQ);
+ [ apply ilemma; [ exact hyp | simpl_PCond_BEURK FLD]
+ | protect_fv "field" in tmp; revert tmp ];
+ clear hyp
+ end)).
+
+Ltac FIELD_SIMPL_EQ FLD lH rl :=
+ get_FldPre FLD ();
+ Field_simplify_eq Ring_tac.ring_subst_niter FLD lH;
+ get_FldPost FLD ().
+
+Tactic Notation (at level 0) "field_simplify_eq" "in" hyp(H) :=
+ let t := type of H in
+ generalize H;
+ field_lookup (PackField FIELD_SIMPL_EQ) [] t;
+ [ try exact I
+ | clear H;intro H].
+
+
+Tactic Notation (at level 0)
+ "field_simplify_eq" "[" constr_list(lH) "]" "in" hyp(H) :=
+ let t := type of H in
+ generalize H;
+ field_lookup (PackField FIELD_SIMPL_EQ) [lH] t;
+ [ try exact I
+ |clear H;intro H].
+
+(* More generic tactics to build variants of field *)
+
+(* This tactic reifies c and pass to F:
+ - the FLD structure gathering all info in the field DB
+ - the atom list
+ - the expression (FExpr)
+ *)
+Ltac gen_with_field F c :=
+ let MetaExpr FLD _ rl :=
+ let R := get_FldCarrier FLD in
+ let mkFFV := get_FFV FLD in
+ let mkFE := get_Meta FLD in
+ let csr :=
+ match rl with
+ | List.cons ?r _ => r
+ | _ => fail 1 "anomaly: ill-formed list"
+ end in
+ let fv := mkFFV csr (@List.nil R) in
+ let expr := mkFE csr fv in
+ F FLD fv expr in
+ field_lookup (PackField MetaExpr) [] (c=c).
+
+
+(* pushes the equation expr = ope(expr) in the goal, and
+ discharge it with field *)
+Ltac prove_field_eqn ope FLD fv expr :=
+ let res := ope expr in
+ let expr' := fresh "input_expr" in
+ pose (expr' := expr);
+ let res' := fresh "result" in
+ pose (res' := res);
+ let lemma := get_L1 FLD in
+ let lemma :=
+ constr:(lemma O fv List.nil expr' res' I List.nil (refl_equal _)) in
+ let ty := type of lemma in
+ let lhs := match ty with
+ forall _, ?lhs=_ -> _ => lhs
+ end in
+ let rhs := match ty with
+ forall _, _=_ -> forall _, ?rhs=_ -> _ => rhs
+ end in
+ let lhs' := fresh "lhs" in let lhs_eq := fresh "lhs_eq" in
+ let rhs' := fresh "rhs" in let rhs_eq := fresh "rhs_eq" in
+ compute_assertion lhs_eq lhs' lhs;
+ compute_assertion rhs_eq rhs' rhs;
+ let H := fresh "fld_eqn" in
+ refine (_ (lemma lhs' lhs_eq rhs' rhs_eq _ _));
+ (* main goal *)
+ [intro H;protect_fv "field" in H; revert H
+ (* ring-nf(lhs') = ring-nf(rhs') *)
+ | vm_compute; reflexivity || fail "field cannot prove this equality"
+ (* denominator condition *)
+ | simpl_PCond FLD];
+ clear lhs_eq rhs_eq; subst lhs' rhs'.
+
+Ltac prove_with_field ope c :=
+ gen_with_field ltac:(prove_field_eqn ope) c.
+
+(* Prove an equation x=ope(x) and rewrite with it *)
+Ltac prove_rw ope x :=
+ prove_with_field ope x;
+ [ let H := fresh "Heq_maple" in
+ intro H; rewrite H; clear H
+ |..].
+
+(* Apply ope (FExpr->FExpr) on an expression *)
+Ltac reduce_field_expr ope kont FLD fv expr :=
+ let evfun := get_FEeval FLD in
+ let res := ope expr in
+ let c := (eval simpl_field_expr in (evfun fv res)) in
+ kont c.
+
+(* Hack to let a Ltac return a term in the context of a primitive tactic *)
+Ltac return_term x := generalize (refl_equal x).
+Ltac get_term :=
+ match goal with
+ | |- ?x = _ -> _ => x
+ end.
+
+(* Turn an operation on field expressions (FExpr) into a reduction
+ on terms (in the field carrier). Because of field_lookup,
+ the tactic cannot return a term directly, so it is returned
+ via the conclusion of the goal (return_term). *)
+Ltac reduce_field_ope ope c :=
+ gen_with_field ltac:(reduce_field_expr ope return_term) c.
+
+
+(* Adding a new field *)
+
+Ltac ring_of_field f :=
+ match type of f with
+ | almost_field_theory _ _ _ _ _ _ _ _ _ => constr:(AF_AR f)
+ | field_theory _ _ _ _ _ _ _ _ _ => constr:(F_R f)
+ | semi_field_theory _ _ _ _ _ _ _ => constr:(SF_SR f)
+ end.
+
+Ltac coerce_to_almost_field set ext f :=
+ match type of f with
+ | almost_field_theory _ _ _ _ _ _ _ _ _ => f
+ | field_theory _ _ _ _ _ _ _ _ _ => constr:(F2AF set ext f)
+ | semi_field_theory _ _ _ _ _ _ _ => constr:(SF2AF set f)
+ end.
+
+Ltac field_elements set ext fspec pspec sspec dspec rk :=
+ let afth := coerce_to_almost_field set ext fspec in
+ let rspec := ring_of_field fspec in
+ ring_elements set ext rspec pspec sspec dspec rk
+ ltac:(fun arth ext_r morph p_spec s_spec d_spec f => f afth ext_r morph p_spec s_spec d_spec).
+
+Ltac field_lemmas set ext inv_m fspec pspec sspec dspec rk :=
+ let get_lemma :=
+ match pspec with None => fun x y => x | _ => fun x y => y end in
+ let simpl_eq_lemma := get_lemma
+ Field_simplify_eq_correct Field_simplify_eq_pow_correct in
+ let simpl_eq_in_lemma := get_lemma
+ Field_simplify_eq_in_correct Field_simplify_eq_pow_in_correct in
+ let rw_lemma := get_lemma
+ Field_rw_correct Field_rw_pow_correct in
+ field_elements set ext fspec pspec sspec dspec rk
+ ltac:(fun afth ext_r morph p_spec s_spec d_spec =>
+ match morph with
+ | _ =>
+ let field_ok1 := constr:(Field_correct set ext_r inv_m afth morph) in
+ match p_spec with
+ | mkhypo ?pp_spec =>
+ let field_ok2 := constr:(field_ok1 _ _ _ pp_spec) in
+ match s_spec with
+ | mkhypo ?ss_spec =>
+ let field_ok3 := constr:(field_ok2 _ ss_spec) in
+ match d_spec with
+ | mkhypo ?dd_spec =>
+ let field_ok := constr:(field_ok3 _ dd_spec) in
+ let mk_lemma lemma :=
+ constr:(lemma _ _ _ _ _ _ _ _ _ _
+ set ext_r inv_m afth
+ _ _ _ _ _ _ _ _ _ morph
+ _ _ _ pp_spec _ ss_spec _ dd_spec) in
+ let field_simpl_eq_ok := mk_lemma simpl_eq_lemma in
+ let field_simpl_ok := mk_lemma rw_lemma in
+ let field_simpl_eq_in := mk_lemma simpl_eq_in_lemma in
+ let cond1_ok :=
+ constr:(Pcond_simpl_gen set ext_r afth morph pp_spec dd_spec) in
+ let cond2_ok :=
+ constr:(Pcond_simpl_complete set ext_r afth morph pp_spec dd_spec) in
+ (fun f =>
+ f afth ext_r morph field_ok field_simpl_ok field_simpl_eq_ok field_simpl_eq_in
+ cond1_ok cond2_ok)
+ | _ => fail 4 "field: bad coefficiant division specification"
+ end
+ | _ => fail 3 "field: bad sign specification"
+ end
+ | _ => fail 2 "field: bad power specification"
+ end
+ | _ => fail 1 "field internal error : field_lemmas, please report"
+ end).
diff --git a/plugins/setoid_ring/Field_theory.v b/plugins/setoid_ring/Field_theory.v
new file mode 100644
index 00000000..9617d409
--- /dev/null
+++ b/plugins/setoid_ring/Field_theory.v
@@ -0,0 +1,1946 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Ring.
+Import Ring_polynom Ring_tac Ring_theory InitialRing Setoid List.
+Require Import ZArith_base.
+(*Require Import Omega.*)
+Set Implicit Arguments.
+
+Section MakeFieldPol.
+
+(* Field elements *)
+ Variable R:Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R->R).
+ Variable (rdiv : R -> R -> R) (rinv : R -> R).
+ Variable req : R -> R -> Prop.
+
+ 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 / y" := (rdiv x y).
+ Notation "- x" := (ropp x). Notation "/ x" := (rinv x).
+ Notation "x == y" := (req x y) (at level 70, no associativity).
+
+ (* Equality properties *)
+ Variable Rsth : Setoid_Theory R req.
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Variable SRinv_ext : forall p q, p == q -> / p == / q.
+
+ (* Field properties *)
+ Record almost_field_theory : Prop := mk_afield {
+ AF_AR : almost_ring_theory rO rI radd rmul rsub ropp req;
+ AF_1_neq_0 : ~ 1 == 0;
+ AFdiv_def : forall p q, p / q == p * / q;
+ AFinv_l : forall p, ~ p == 0 -> / p * p == 1
+ }.
+
+Section AlmostField.
+
+ Variable AFth : almost_field_theory.
+ Let ARth := AFth.(AF_AR).
+ Let rI_neq_rO := AFth.(AF_1_neq_0).
+ Let rdiv_def := AFth.(AFdiv_def).
+ Let rinv_l := AFth.(AFinv_l).
+
+ (* Coefficients *)
+ Variable C: Type.
+ Variable (cO cI: C) (cadd cmul csub : C->C->C) (copp : C->C).
+ Variable ceqb : C->C->bool.
+ Variable phi : C -> R.
+
+ Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req
+ cO cI cadd cmul csub copp ceqb phi.
+
+Lemma ceqb_rect : forall c1 c2 (A:Type) (x y:A) (P:A->Type),
+ (phi c1 == phi c2 -> P x) -> P y -> P (if ceqb c1 c2 then x else y).
+Proof.
+intros.
+generalize (fun h => X (morph_eq CRmorph c1 c2 h)).
+case (ceqb c1 c2); auto.
+Qed.
+
+
+ (* C notations *)
+ Notation "x +! y" := (cadd x y) (at level 50).
+ Notation "x *! y " := (cmul x y) (at level 40).
+ Notation "x -! y " := (csub x y) (at level 50).
+ Notation "-! x" := (copp x) (at level 35).
+ Notation " x ?=! y" := (ceqb x y) (at level 70, no associativity).
+ Notation "[ x ]" := (phi x) (at level 0).
+
+
+ (* Useful tactics *)
+ Add Setoid R req Rsth as R_set1.
+ 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 rinv : rinv_ext. exact SRinv_ext. Qed.
+
+Let eq_trans := Setoid.Seq_trans _ _ Rsth.
+Let eq_sym := Setoid.Seq_sym _ _ Rsth.
+Let eq_refl := Setoid.Seq_refl _ _ Rsth.
+
+Hint Resolve eq_refl rdiv_def rinv_l rI_neq_rO CRmorph.(morph1) .
+Hint Resolve (Rmul_ext Reqe) (Rmul_ext Reqe) (Radd_ext Reqe)
+ (ARsub_ext Rsth Reqe ARth) (Ropp_ext Reqe) SRinv_ext.
+Hint Resolve (ARadd_0_l ARth) (ARadd_comm ARth) (ARadd_assoc ARth)
+ (ARmul_1_l ARth) (ARmul_0_l ARth)
+ (ARmul_comm ARth) (ARmul_assoc ARth) (ARdistr_l ARth)
+ (ARopp_mul_l ARth) (ARopp_add ARth)
+ (ARsub_def ARth) .
+
+ (* Power coefficients *)
+ Variable Cpow : Set.
+ Variable Cp_phi : N -> Cpow.
+ Variable rpow : R -> Cpow -> R.
+ Variable pow_th : power_theory rI rmul req Cp_phi rpow.
+ (* sign function *)
+ Variable get_sign : C -> option C.
+ Variable get_sign_spec : sign_theory copp ceqb get_sign.
+
+ Variable cdiv:C -> C -> C*C.
+ Variable cdiv_th : div_theory req cadd cmul phi cdiv.
+
+Notation NPEeval := (PEeval rO radd rmul rsub ropp phi Cp_phi rpow).
+Notation Nnorm:= (norm_subst cO cI cadd cmul csub copp ceqb cdiv).
+
+Notation NPphi_dev := (Pphi_dev rO rI radd rmul rsub ropp cO cI ceqb phi get_sign).
+Notation NPphi_pow := (Pphi_pow rO rI radd rmul rsub ropp cO cI ceqb phi Cp_phi rpow get_sign).
+
+(* add abstract semi-ring to help with some proofs *)
+Add Ring Rring : (ARth_SRth ARth).
+
+
+(* additional ring properties *)
+
+Lemma rsub_0_l : forall r, 0 - r == - r.
+intros; rewrite (ARsub_def ARth) in |- *;ring.
+Qed.
+
+Lemma rsub_0_r : forall r, r - 0 == r.
+intros; rewrite (ARsub_def ARth) in |- *.
+rewrite (ARopp_zero Rsth Reqe ARth) in |- *; ring.
+Qed.
+
+(***************************************************************************
+
+ Properties of division
+
+ ***************************************************************************)
+
+Theorem rdiv_simpl: forall p q, ~ q == 0 -> q * (p / q) == p.
+intros p q H.
+rewrite rdiv_def in |- *.
+transitivity (/ q * q * p); [ ring | idtac ].
+rewrite rinv_l in |- *; auto.
+Qed.
+Hint Resolve rdiv_simpl .
+
+Theorem SRdiv_ext:
+ forall p1 p2, p1 == p2 -> forall q1 q2, q1 == q2 -> p1 / q1 == p2 / q2.
+intros p1 p2 H q1 q2 H0.
+transitivity (p1 * / q1); auto.
+transitivity (p2 * / q2); auto.
+Qed.
+Hint Resolve SRdiv_ext .
+
+ Add Morphism rdiv : rdiv_ext. exact SRdiv_ext. Qed.
+
+Lemma rmul_reg_l : forall p q1 q2,
+ ~ p == 0 -> p * q1 == p * q2 -> q1 == q2.
+intros.
+rewrite <- (@rdiv_simpl q1 p) in |- *; trivial.
+rewrite <- (@rdiv_simpl q2 p) in |- *; trivial.
+repeat rewrite rdiv_def in |- *.
+repeat rewrite (ARmul_assoc ARth) in |- *.
+auto.
+Qed.
+
+Theorem field_is_integral_domain : forall r1 r2,
+ ~ r1 == 0 -> ~ r2 == 0 -> ~ r1 * r2 == 0.
+Proof.
+red in |- *; intros.
+apply H0.
+transitivity (1 * r2); auto.
+transitivity (/ r1 * r1 * r2); auto.
+rewrite <- (ARmul_assoc ARth) in |- *.
+rewrite H1 in |- *.
+apply ARmul_0_r with (1 := Rsth) (2 := ARth).
+Qed.
+
+Theorem ropp_neq_0 : forall r,
+ ~ -(1) == 0 -> ~ r == 0 -> ~ -r == 0.
+intros.
+setoid_replace (- r) with (- (1) * r).
+ apply field_is_integral_domain; trivial.
+ rewrite <- (ARopp_mul_l ARth) in |- *.
+ rewrite (ARmul_1_l ARth) in |- *.
+ reflexivity.
+Qed.
+
+Theorem rdiv_r_r : forall r, ~ r == 0 -> r / r == 1.
+intros.
+rewrite (AFdiv_def AFth) in |- *.
+rewrite (ARmul_comm ARth) in |- *.
+apply (AFinv_l AFth).
+trivial.
+Qed.
+
+Theorem rdiv1: forall r, r == r / 1.
+intros r; transitivity (1 * (r / 1)); auto.
+Qed.
+
+Theorem rdiv2:
+ forall r1 r2 r3 r4,
+ ~ r2 == 0 ->
+ ~ r4 == 0 ->
+ r1 / r2 + r3 / r4 == (r1 * r4 + r3 * r2) / (r2 * r4).
+Proof.
+intros r1 r2 r3 r4 H H0.
+assert (~ r2 * r4 == 0) by complete (apply field_is_integral_domain; trivial).
+apply rmul_reg_l with (r2 * r4); trivial.
+rewrite rdiv_simpl in |- *; trivial.
+rewrite (ARdistr_r Rsth Reqe ARth) in |- *.
+apply (Radd_ext Reqe).
+ transitivity (r2 * (r1 / r2) * r4); [ ring | auto ].
+ transitivity (r2 * (r4 * (r3 / r4))); auto.
+ transitivity (r2 * r3); auto.
+Qed.
+
+
+Theorem rdiv2b:
+ forall r1 r2 r3 r4 r5,
+ ~ (r2*r5) == 0 ->
+ ~ (r4*r5) == 0 ->
+ r1 / (r2*r5) + r3 / (r4*r5) == (r1 * r4 + r3 * r2) / (r2 * (r4 * r5)).
+Proof.
+intros r1 r2 r3 r4 r5 H H0.
+assert (HH1: ~ r2 == 0) by (intros HH; case H; rewrite HH; ring).
+assert (HH2: ~ r5 == 0) by (intros HH; case H; rewrite HH; ring).
+assert (HH3: ~ r4 == 0) by (intros HH; case H0; rewrite HH; ring).
+assert (HH4: ~ r2 * (r4 * r5) == 0)
+ by complete (repeat apply field_is_integral_domain; trivial).
+apply rmul_reg_l with (r2 * (r4 * r5)); trivial.
+rewrite rdiv_simpl in |- *; trivial.
+rewrite (ARdistr_r Rsth Reqe ARth) in |- *.
+apply (Radd_ext Reqe).
+ transitivity ((r2 * r5) * (r1 / (r2 * r5)) * r4); [ ring | auto ].
+ transitivity ((r4 * r5) * (r3 / (r4 * r5)) * r2); [ ring | auto ].
+Qed.
+
+Theorem rdiv5: forall r1 r2, - (r1 / r2) == - r1 / r2.
+intros r1 r2.
+transitivity (- (r1 * / r2)); auto.
+transitivity (- r1 * / r2); auto.
+Qed.
+Hint Resolve rdiv5 .
+
+Theorem rdiv3:
+ forall r1 r2 r3 r4,
+ ~ r2 == 0 ->
+ ~ r4 == 0 ->
+ r1 / r2 - r3 / r4 == (r1 * r4 - r3 * r2) / (r2 * r4).
+intros r1 r2 r3 r4 H H0.
+assert (~ r2 * r4 == 0) by (apply field_is_integral_domain; trivial).
+transitivity (r1 / r2 + - (r3 / r4)); auto.
+transitivity (r1 / r2 + - r3 / r4); auto.
+transitivity ((r1 * r4 + - r3 * r2) / (r2 * r4)); auto.
+apply rdiv2; auto.
+apply SRdiv_ext; auto.
+transitivity (r1 * r4 + - (r3 * r2)); symmetry; auto.
+Qed.
+
+
+Theorem rdiv3b:
+ forall r1 r2 r3 r4 r5,
+ ~ (r2 * r5) == 0 ->
+ ~ (r4 * r5) == 0 ->
+ r1 / (r2*r5) - r3 / (r4*r5) == (r1 * r4 - r3 * r2) / (r2 * (r4 * r5)).
+Proof.
+intros r1 r2 r3 r4 r5 H H0.
+transitivity (r1 / (r2 * r5) + - (r3 / (r4 * r5))); auto.
+transitivity (r1 / (r2 * r5) + - r3 / (r4 * r5)); auto.
+transitivity ((r1 * r4 + - r3 * r2) / (r2 * (r4 * r5))).
+apply rdiv2b; auto; try ring.
+apply (SRdiv_ext); auto.
+transitivity (r1 * r4 + - (r3 * r2)); symmetry; auto.
+Qed.
+
+Theorem rdiv6:
+ forall r1 r2,
+ ~ r1 == 0 -> ~ r2 == 0 -> / (r1 / r2) == r2 / r1.
+intros r1 r2 H H0.
+assert (~ r1 / r2 == 0) as Hk.
+ intros H1; case H.
+ transitivity (r2 * (r1 / r2)); auto.
+ rewrite H1 in |- *; ring.
+ apply rmul_reg_l with (r1 / r2); auto.
+ transitivity (/ (r1 / r2) * (r1 / r2)); auto.
+ transitivity 1; auto.
+ repeat rewrite rdiv_def in |- *.
+ transitivity (/ r1 * r1 * (/ r2 * r2)); [ idtac | ring ].
+ repeat rewrite rinv_l in |- *; auto.
+Qed.
+Hint Resolve rdiv6 .
+
+ Theorem rdiv4:
+ forall r1 r2 r3 r4,
+ ~ r2 == 0 ->
+ ~ r4 == 0 ->
+ (r1 / r2) * (r3 / r4) == (r1 * r3) / (r2 * r4).
+Proof.
+intros r1 r2 r3 r4 H H0.
+assert (~ r2 * r4 == 0) by complete (apply field_is_integral_domain; trivial).
+apply rmul_reg_l with (r2 * r4); trivial.
+rewrite rdiv_simpl in |- *; trivial.
+transitivity (r2 * (r1 / r2) * (r4 * (r3 / r4))); [ ring | idtac ].
+repeat rewrite rdiv_simpl in |- *; trivial.
+Qed.
+
+ Theorem rdiv4b:
+ forall r1 r2 r3 r4 r5 r6,
+ ~ r2 * r5 == 0 ->
+ ~ r4 * r6 == 0 ->
+ ((r1 * r6) / (r2 * r5)) * ((r3 * r5) / (r4 * r6)) == (r1 * r3) / (r2 * r4).
+Proof.
+intros r1 r2 r3 r4 r5 r6 H H0.
+rewrite rdiv4; auto.
+transitivity ((r5 * r6) * (r1 * r3) / ((r5 * r6) * (r2 * r4))).
+apply SRdiv_ext; ring.
+assert (HH: ~ r5*r6 == 0).
+ apply field_is_integral_domain.
+ intros H1; case H; rewrite H1; ring.
+ intros H1; case H0; rewrite H1; ring.
+rewrite <- rdiv4 ; auto.
+ rewrite rdiv_r_r; auto.
+
+ apply field_is_integral_domain.
+ intros H1; case H; rewrite H1; ring.
+ intros H1; case H0; rewrite H1; ring.
+Qed.
+
+
+Theorem rdiv7:
+ forall r1 r2 r3 r4,
+ ~ r2 == 0 ->
+ ~ r3 == 0 ->
+ ~ r4 == 0 ->
+ (r1 / r2) / (r3 / r4) == (r1 * r4) / (r2 * r3).
+Proof.
+intros.
+rewrite (rdiv_def (r1 / r2)) in |- *.
+rewrite rdiv6 in |- *; trivial.
+apply rdiv4; trivial.
+Qed.
+
+Theorem rdiv7b:
+ forall r1 r2 r3 r4 r5 r6,
+ ~ r2 * r6 == 0 ->
+ ~ r3 * r5 == 0 ->
+ ~ r4 * r6 == 0 ->
+ ((r1 * r5) / (r2 * r6)) / ((r3 * r5) / (r4 * r6)) == (r1 * r4) / (r2 * r3).
+Proof.
+intros.
+rewrite rdiv7; auto.
+transitivity ((r5 * r6) * (r1 * r4) / ((r5 * r6) * (r2 * r3))).
+apply SRdiv_ext; ring.
+assert (HH: ~ r5*r6 == 0).
+ apply field_is_integral_domain.
+ intros H2; case H0; rewrite H2; ring.
+ intros H2; case H1; rewrite H2; ring.
+rewrite <- rdiv4 ; auto.
+rewrite rdiv_r_r; auto.
+ apply field_is_integral_domain.
+ intros H2; case H; rewrite H2; ring.
+ intros H2; case H0; rewrite H2; ring.
+Qed.
+
+
+Theorem rdiv8: forall r1 r2, ~ r2 == 0 -> r1 == 0 -> r1 / r2 == 0.
+intros r1 r2 H H0.
+transitivity (r1 * / r2); auto.
+transitivity (0 * / r2); auto.
+Qed.
+
+
+Theorem cross_product_eq : forall r1 r2 r3 r4,
+ ~ r2 == 0 -> ~ r4 == 0 -> r1 * r4 == r3 * r2 -> r1 / r2 == r3 / r4.
+intros.
+transitivity (r1 / r2 * (r4 / r4)).
+ rewrite rdiv_r_r in |- *; trivial.
+ symmetry in |- *.
+ apply (ARmul_1_r Rsth ARth).
+ rewrite rdiv4 in |- *; trivial.
+ rewrite H1 in |- *.
+ rewrite (ARmul_comm ARth r2 r4) in |- *.
+ rewrite <- rdiv4 in |- *; trivial.
+ rewrite rdiv_r_r in |- * by trivial.
+ apply (ARmul_1_r Rsth ARth).
+Qed.
+
+(***************************************************************************
+
+ Some equality test
+
+ ***************************************************************************)
+
+Fixpoint positive_eq (p1 p2 : positive) {struct p1} : bool :=
+ match p1, p2 with
+ xH, xH => true
+ | xO p3, xO p4 => positive_eq p3 p4
+ | xI p3, xI p4 => positive_eq p3 p4
+ | _, _ => false
+ end.
+
+Theorem positive_eq_correct:
+ forall p1 p2, if positive_eq p1 p2 then p1 = p2 else p1 <> p2.
+intros p1; elim p1;
+ (try (intros p2; case p2; simpl; auto; intros; discriminate)).
+intros p3 rec p2; case p2; simpl; auto; (try (intros; discriminate)); intros p4.
+generalize (rec p4); case (positive_eq p3 p4); auto.
+intros H1; apply f_equal with ( f := xI ); auto.
+intros H1 H2; case H1; injection H2; auto.
+intros p3 rec p2; case p2; simpl; auto; (try (intros; discriminate)); intros p4.
+generalize (rec p4); case (positive_eq p3 p4); auto.
+intros H1; apply f_equal with ( f := xO ); auto.
+intros H1 H2; case H1; injection H2; auto.
+Qed.
+
+Definition N_eq n1 n2 :=
+ match n1, n2 with
+ | N0, N0 => true
+ | Npos p1, Npos p2 => positive_eq p1 p2
+ | _, _ => false
+ end.
+
+Lemma N_eq_correct : forall n1 n2, if N_eq n1 n2 then n1 = n2 else n1 <> n2.
+Proof.
+ intros [ |p1] [ |p2];simpl;trivial;try(intro H;discriminate H;fail).
+ assert (H:=positive_eq_correct p1 p2);destruct (positive_eq p1 p2);
+ [rewrite H;trivial | intro H1;injection H1;subst;apply H;trivial].
+Qed.
+
+(* equality test *)
+Fixpoint PExpr_eq (e1 e2 : PExpr C) {struct e1} : bool :=
+ match e1, e2 with
+ PEc c1, PEc c2 => ceqb c1 c2
+ | PEX p1, PEX p2 => positive_eq p1 p2
+ | PEadd e3 e5, PEadd e4 e6 => if PExpr_eq e3 e4 then PExpr_eq e5 e6 else false
+ | PEsub e3 e5, PEsub e4 e6 => if PExpr_eq e3 e4 then PExpr_eq e5 e6 else false
+ | PEmul e3 e5, PEmul e4 e6 => if PExpr_eq e3 e4 then PExpr_eq e5 e6 else false
+ | PEopp e3, PEopp e4 => PExpr_eq e3 e4
+ | PEpow e3 n3, PEpow e4 n4 => if N_eq n3 n4 then PExpr_eq e3 e4 else false
+ | _, _ => false
+ end.
+
+Add Morphism (pow_pos rmul) with signature req ==> eq ==> req as pow_morph.
+intros x y H p;induction p as [p IH| p IH|];simpl;auto;ring[IH].
+Qed.
+
+Add Morphism (pow_N rI rmul) with signature req ==> eq ==> req as pow_N_morph.
+intros x y H [|p];simpl;auto. apply pow_morph;trivial.
+Qed.
+(*
+Lemma rpow_morph : forall x y n, x == y ->rpow x (Cp_phi n) == rpow y (Cp_phi n).
+Proof.
+ intros; repeat rewrite pow_th.(rpow_pow_N).
+ destruct n;simpl. apply eq_refl.
+ induction p;simpl;try rewrite IHp;try rewrite H; apply eq_refl.
+Qed.
+*)
+Theorem PExpr_eq_semi_correct:
+ forall l e1 e2, PExpr_eq e1 e2 = true -> NPEeval l e1 == NPEeval l e2.
+intros l e1; elim e1.
+intros c1; intros e2; elim e2; simpl; (try (intros; discriminate)).
+intros c2; apply (morph_eq CRmorph).
+intros p1; intros e2; elim e2; simpl; (try (intros; discriminate)).
+intros p2; generalize (positive_eq_correct p1 p2); case (positive_eq p1 p2);
+ (try (intros; discriminate)); intros H; rewrite H; auto.
+intros e3 rec1 e5 rec2 e2; case e2; simpl; (try (intros; discriminate)).
+intros e4 e6; generalize (rec1 e4); case (PExpr_eq e3 e4);
+ (try (intros; discriminate)); generalize (rec2 e6); case (PExpr_eq e5 e6);
+ (try (intros; discriminate)); auto.
+intros e3 rec1 e5 rec2 e2; case e2; simpl; (try (intros; discriminate)).
+intros e4 e6; generalize (rec1 e4); case (PExpr_eq e3 e4);
+ (try (intros; discriminate)); generalize (rec2 e6); case (PExpr_eq e5 e6);
+ (try (intros; discriminate)); auto.
+intros e3 rec1 e5 rec2 e2; case e2; simpl; (try (intros; discriminate)).
+intros e4 e6; generalize (rec1 e4); case (PExpr_eq e3 e4);
+ (try (intros; discriminate)); generalize (rec2 e6); case (PExpr_eq e5 e6);
+ (try (intros; discriminate)); auto.
+intros e3 rec e2; (case e2; simpl; (try (intros; discriminate))).
+intros e4; generalize (rec e4); case (PExpr_eq e3 e4);
+ (try (intros; discriminate)); auto.
+intros e3 rec n3 e2;(case e2;simpl;(try (intros;discriminate))).
+intros e4 n4;generalize (N_eq_correct n3 n4);destruct (N_eq n3 n4);
+intros;try discriminate.
+repeat rewrite pow_th.(rpow_pow_N);rewrite H;rewrite (rec _ H0);auto.
+Qed.
+
+(* add *)
+Definition NPEadd e1 e2 :=
+ match e1, e2 with
+ PEc c1, PEc c2 => PEc (cadd c1 c2)
+ | PEc c, _ => if ceqb c cO then e2 else PEadd e1 e2
+ | _, PEc c => if ceqb c cO then e1 else PEadd e1 e2
+ (* Peut t'on factoriser ici ??? *)
+ | _, _ => PEadd e1 e2
+ end.
+
+Theorem NPEadd_correct:
+ forall l e1 e2, NPEeval l (NPEadd e1 e2) == NPEeval l (PEadd e1 e2).
+Proof.
+intros l e1 e2.
+destruct e1; destruct e2; simpl in |- *; try reflexivity; try apply ceqb_rect;
+ try (intro eq_c; rewrite eq_c in |- *); simpl in |- *;try apply eq_refl;
+ try (ring [(morph0 CRmorph)]).
+ apply (morph_add CRmorph).
+Qed.
+
+Definition NPEpow x n :=
+ match n with
+ | N0 => PEc cI
+ | Npos p =>
+ if positive_eq p xH then x else
+ match x with
+ | PEc c =>
+ if ceqb c cI then PEc cI else if ceqb c cO then PEc cO else PEc (pow_pos cmul c p)
+ | _ => PEpow x n
+ end
+ end.
+
+Theorem NPEpow_correct : forall l e n,
+ NPEeval l (NPEpow e n) == NPEeval l (PEpow e n).
+Proof.
+ destruct n;simpl.
+ rewrite pow_th.(rpow_pow_N);simpl;auto.
+ generalize (positive_eq_correct p xH).
+ destruct (positive_eq p 1);intros.
+ rewrite H;rewrite pow_th.(rpow_pow_N). trivial.
+ clear H;destruct e;simpl;auto.
+ repeat apply ceqb_rect;simpl;intros;rewrite pow_th.(rpow_pow_N);simpl.
+ symmetry;induction p;simpl;trivial; ring [IHp H CRmorph.(morph1)].
+ symmetry; induction p;simpl;trivial;ring [IHp CRmorph.(morph0)].
+ induction p;simpl;auto;repeat rewrite CRmorph.(morph_mul);ring [IHp].
+Qed.
+
+(* mul *)
+Fixpoint NPEmul (x y : PExpr C) {struct x} : PExpr C :=
+ match x, y with
+ PEc c1, PEc c2 => PEc (cmul c1 c2)
+ | PEc c, _ =>
+ if ceqb c cI then y else if ceqb c cO then PEc cO else PEmul x y
+ | _, PEc c =>
+ if ceqb c cI then x else if ceqb c cO then PEc cO else PEmul x y
+ | PEpow e1 n1, PEpow e2 n2 =>
+ if N_eq n1 n2 then NPEpow (NPEmul e1 e2) n1 else PEmul x y
+ | _, _ => PEmul x y
+ end.
+
+Lemma pow_pos_mul : forall x y p, pow_pos rmul (x * y) p == pow_pos rmul x p * pow_pos rmul y p.
+induction p;simpl;auto;try ring [IHp].
+Qed.
+
+Theorem NPEmul_correct : forall l e1 e2,
+ NPEeval l (NPEmul e1 e2) == NPEeval l (PEmul e1 e2).
+induction e1;destruct e2; simpl in |- *;try reflexivity;
+ repeat apply ceqb_rect;
+ try (intro eq_c; rewrite eq_c in |- *); simpl in |- *; try reflexivity;
+ try ring [(morph0 CRmorph) (morph1 CRmorph)].
+ apply (morph_mul CRmorph).
+assert (H:=N_eq_correct n n0);destruct (N_eq n n0).
+rewrite NPEpow_correct. simpl.
+repeat rewrite pow_th.(rpow_pow_N).
+rewrite IHe1;rewrite <- H;destruct n;simpl;try ring.
+apply pow_pos_mul.
+simpl;auto.
+Qed.
+
+(* sub *)
+Definition NPEsub e1 e2 :=
+ match e1, e2 with
+ PEc c1, PEc c2 => PEc (csub c1 c2)
+ | PEc c, _ => if ceqb c cO then PEopp e2 else PEsub e1 e2
+ | _, PEc c => if ceqb c cO then e1 else PEsub e1 e2
+ (* Peut-on factoriser ici *)
+ | _, _ => PEsub e1 e2
+ end.
+
+Theorem NPEsub_correct:
+ forall l e1 e2, NPEeval l (NPEsub e1 e2) == NPEeval l (PEsub e1 e2).
+intros l e1 e2.
+destruct e1; destruct e2; simpl in |- *; try reflexivity; try apply ceqb_rect;
+ try (intro eq_c; rewrite eq_c in |- *); simpl in |- *;
+ try rewrite (morph0 CRmorph) in |- *; try reflexivity;
+ try (symmetry; apply rsub_0_l); try (symmetry; apply rsub_0_r).
+apply (morph_sub CRmorph).
+Qed.
+
+(* opp *)
+Definition NPEopp e1 :=
+ match e1 with PEc c1 => PEc (copp c1) | _ => PEopp e1 end.
+
+Theorem NPEopp_correct:
+ forall l e1, NPEeval l (NPEopp e1) == NPEeval l (PEopp e1).
+intros l e1; case e1; simpl; auto.
+intros; apply (morph_opp CRmorph).
+Qed.
+
+(* simplification *)
+Fixpoint PExpr_simp (e : PExpr C) : PExpr C :=
+ match e with
+ PEadd e1 e2 => NPEadd (PExpr_simp e1) (PExpr_simp e2)
+ | PEmul e1 e2 => NPEmul (PExpr_simp e1) (PExpr_simp e2)
+ | PEsub e1 e2 => NPEsub (PExpr_simp e1) (PExpr_simp e2)
+ | PEopp e1 => NPEopp (PExpr_simp e1)
+ | PEpow e1 n1 => NPEpow (PExpr_simp e1) n1
+ | _ => e
+ end.
+
+Theorem PExpr_simp_correct:
+ forall l e, NPEeval l (PExpr_simp e) == NPEeval l e.
+intros l e; elim e; simpl; auto.
+intros e1 He1 e2 He2.
+transitivity (NPEeval l (PEadd (PExpr_simp e1) (PExpr_simp e2))); auto.
+apply NPEadd_correct.
+simpl; auto.
+intros e1 He1 e2 He2.
+transitivity (NPEeval l (PEsub (PExpr_simp e1) (PExpr_simp e2))); auto.
+apply NPEsub_correct.
+simpl; auto.
+intros e1 He1 e2 He2.
+transitivity (NPEeval l (PEmul (PExpr_simp e1) (PExpr_simp e2))); auto.
+apply NPEmul_correct.
+simpl; auto.
+intros e1 He1.
+transitivity (NPEeval l (PEopp (PExpr_simp e1))); auto.
+apply NPEopp_correct.
+simpl; auto.
+intros e1 He1 n;simpl.
+rewrite NPEpow_correct;simpl.
+repeat rewrite pow_th.(rpow_pow_N).
+rewrite He1;auto.
+Qed.
+
+
+(****************************************************************************
+
+ Datastructure
+
+ ***************************************************************************)
+
+(* The input: syntax of a field expression *)
+
+Inductive FExpr : Type :=
+ FEc: C -> FExpr
+ | FEX: positive -> FExpr
+ | FEadd: FExpr -> FExpr -> FExpr
+ | FEsub: FExpr -> FExpr -> FExpr
+ | FEmul: FExpr -> FExpr -> FExpr
+ | FEopp: FExpr -> FExpr
+ | FEinv: FExpr -> FExpr
+ | FEdiv: FExpr -> FExpr -> FExpr
+ | FEpow: FExpr -> N -> FExpr .
+
+Fixpoint FEeval (l : list R) (pe : FExpr) {struct pe} : R :=
+ match pe with
+ | FEc c => phi c
+ | FEX x => BinList.nth 0 x l
+ | FEadd x y => FEeval l x + FEeval l y
+ | FEsub x y => FEeval l x - FEeval l y
+ | FEmul x y => FEeval l x * FEeval l y
+ | FEopp x => - FEeval l x
+ | FEinv x => / FEeval l x
+ | FEdiv x y => FEeval l x / FEeval l y
+ | FEpow x n => rpow (FEeval l x) (Cp_phi n)
+ end.
+
+Strategy expand [FEeval].
+
+(* The result of the normalisation *)
+
+Record linear : Type := mk_linear {
+ num : PExpr C;
+ denum : PExpr C;
+ condition : list (PExpr C) }.
+
+(***************************************************************************
+
+ Semantics and properties of side condition
+
+ ***************************************************************************)
+
+Fixpoint PCond (l : list R) (le : list (PExpr C)) {struct le} : Prop :=
+ match le with
+ | nil => True
+ | e1 :: nil => ~ req (NPEeval l e1) rO
+ | e1 :: l1 => ~ req (NPEeval l e1) rO /\ PCond l l1
+ end.
+
+Theorem PCond_cons_inv_l :
+ forall l a l1, PCond l (a::l1) -> ~ NPEeval l a == 0.
+intros l a l1 H.
+destruct l1; simpl in H |- *; trivial.
+destruct H; trivial.
+Qed.
+
+Theorem PCond_cons_inv_r : forall l a l1, PCond l (a :: l1) -> PCond l l1.
+intros l a l1 H.
+destruct l1; simpl in H |- *; trivial.
+destruct H; trivial.
+Qed.
+
+Theorem PCond_app_inv_l: forall l l1 l2, PCond l (l1 ++ l2) -> PCond l l1.
+intros l l1 l2; elim l1; simpl app in |- *.
+ simpl in |- *; auto.
+ destruct l0; simpl in *.
+ destruct l2; firstorder.
+ firstorder.
+Qed.
+
+Theorem PCond_app_inv_r: forall l l1 l2, PCond l (l1 ++ l2) -> PCond l l2.
+intros l l1 l2; elim l1; simpl app; auto.
+intros a l0 H H0; apply H; apply PCond_cons_inv_r with ( 1 := H0 ).
+Qed.
+
+(* An unsatisfiable condition: issued when a division by zero is detected *)
+Definition absurd_PCond := cons (PEc cO) nil.
+
+Lemma absurd_PCond_bottom : forall l, ~ PCond l absurd_PCond.
+unfold absurd_PCond in |- *; simpl in |- *.
+red in |- *; intros.
+apply H.
+apply (morph0 CRmorph).
+Qed.
+
+(***************************************************************************
+
+ Normalisation
+
+ ***************************************************************************)
+
+Fixpoint isIn (e1:PExpr C) (p1:positive)
+ (e2:PExpr C) (p2:positive) {struct e2}: option (N * PExpr C) :=
+ match e2 with
+ | PEmul e3 e4 =>
+ match isIn e1 p1 e3 p2 with
+ | Some (N0, e5) => Some (N0, NPEmul e5 (NPEpow e4 (Npos p2)))
+ | Some (Npos p, e5) =>
+ match isIn e1 p e4 p2 with
+ | Some (n, e6) => Some (n, NPEmul e5 e6)
+ | None => Some (Npos p, NPEmul e5 (NPEpow e4 (Npos p2)))
+ end
+ | None =>
+ match isIn e1 p1 e4 p2 with
+ | Some (n, e5) => Some (n,NPEmul (NPEpow e3 (Npos p2)) e5)
+ | None => None
+ end
+ end
+ | PEpow e3 N0 => None
+ | PEpow e3 (Npos p3) => isIn e1 p1 e3 (Pmult p3 p2)
+ | _ =>
+ if PExpr_eq e1 e2 then
+ match Zminus (Zpos p1) (Zpos p2) with
+ | Zpos p => Some (Npos p, PEc cI)
+ | Z0 => Some (N0, PEc cI)
+ | Zneg p => Some (N0, NPEpow e2 (Npos p))
+ end
+ else None
+ end.
+
+ Definition ZtoN z := match z with Zpos p => Npos p | _ => N0 end.
+ Definition NtoZ n := match n with Npos p => Zpos p | _ => Z0 end.
+
+ Notation pow_pos_plus := (Ring_theory.pow_pos_Pplus _ Rsth Reqe.(Rmul_ext)
+ ARth.(ARmul_comm) ARth.(ARmul_assoc)).
+
+ Lemma isIn_correct_aux : forall l e1 e2 p1 p2,
+ match
+ (if PExpr_eq e1 e2 then
+ match Zminus (Zpos p1) (Zpos p2) with
+ | Zpos p => Some (Npos p, PEc cI)
+ | Z0 => Some (N0, PEc cI)
+ | Zneg p => Some (N0, NPEpow e2 (Npos p))
+ end
+ else None)
+ with
+ | Some(n, e3) =>
+ NPEeval l (PEpow e2 (Npos p2)) ==
+ NPEeval l (PEmul (PEpow e1 (ZtoN (Zpos p1 - NtoZ n))) e3) /\
+ (Zpos p1 > NtoZ n)%Z
+ | _ => True
+ end.
+Proof.
+ intros l e1 e2 p1 p2; generalize (PExpr_eq_semi_correct l e1 e2);
+ case (PExpr_eq e1 e2); simpl; auto; intros H.
+ case_eq ((p1 ?= p2)%positive Eq);intros;simpl.
+ repeat rewrite pow_th.(rpow_pow_N);simpl. split. 2:refine (refl_equal _).
+ rewrite (Pcompare_Eq_eq _ _ H0).
+ rewrite H by trivial. ring [ (morph1 CRmorph)].
+ fold (NPEpow e2 (Npos (p2 - p1))).
+ rewrite NPEpow_correct;simpl.
+ repeat rewrite pow_th.(rpow_pow_N);simpl.
+ rewrite H;trivial. split. 2:refine (refl_equal _).
+ rewrite <- pow_pos_plus; rewrite Pplus_minus;auto. apply ZC2;trivial.
+ repeat rewrite pow_th.(rpow_pow_N);simpl.
+ rewrite H;trivial.
+ change (ZtoN
+ match (p1 ?= p1 - p2)%positive Eq with
+ | Eq => 0
+ | Lt => Zneg (p1 - p2 - p1)
+ | Gt => Zpos (p1 - (p1 - p2))
+ end) with (ZtoN (Zpos p1 - Zpos (p1 -p2))).
+ replace (Zpos (p1 - p2)) with (Zpos p1 - Zpos p2)%Z.
+ split.
+ repeat rewrite Zth.(Rsub_def). rewrite (Ring_theory.Ropp_add Zsth Zeqe Zth).
+ rewrite Zplus_assoc. simpl. rewrite Pcompare_refl. simpl.
+ ring [ (morph1 CRmorph)].
+ assert (Zpos p1 > 0 /\ Zpos p2 > 0)%Z. split;refine (refl_equal _).
+ apply Zplus_gt_reg_l with (Zpos p2).
+ rewrite Zplus_minus. change (Zpos p2 + Zpos p1 > 0 + Zpos p1)%Z.
+ apply Zplus_gt_compat_r. refine (refl_equal _).
+ simpl;rewrite H0;trivial.
+Qed.
+
+Lemma pow_pos_pow_pos : forall x p1 p2, pow_pos rmul (pow_pos rmul x p1) p2 == pow_pos rmul x (p1*p2).
+induction p1;simpl;intros;repeat rewrite pow_pos_mul;repeat rewrite pow_pos_plus;simpl.
+ring [(IHp1 p2)]. ring [(IHp1 p2)]. auto.
+Qed.
+
+
+Theorem isIn_correct: forall l e1 p1 e2 p2,
+ match isIn e1 p1 e2 p2 with
+ | Some(n, e3) =>
+ NPEeval l (PEpow e2 (Npos p2)) ==
+ NPEeval l (PEmul (PEpow e1 (ZtoN (Zpos p1 - NtoZ n))) e3) /\
+ (Zpos p1 > NtoZ n)%Z
+ | _ => True
+ end.
+Proof.
+Opaque NPEpow.
+intros l e1 p1 e2; generalize p1;clear p1;elim e2; intros;
+ try (refine (isIn_correct_aux l e1 _ p1 p2);fail);simpl isIn.
+generalize (H p1 p2);clear H;destruct (isIn e1 p1 p p2). destruct p3.
+destruct n.
+ simpl. rewrite NPEmul_correct. simpl; rewrite NPEpow_correct;simpl.
+ repeat rewrite pow_th.(rpow_pow_N);simpl.
+ rewrite pow_pos_mul;intros (H,H1);split;[ring[H]|trivial].
+ generalize (H0 p4 p2);clear H0;destruct (isIn e1 p4 p0 p2). destruct p5.
+ destruct n;simpl.
+ rewrite NPEmul_correct;repeat rewrite pow_th.(rpow_pow_N);simpl.
+ intros (H1,H2) (H3,H4).
+ unfold Zgt in H2, H4;simpl in H2,H4. rewrite H4 in H3;simpl in H3.
+ rewrite pow_pos_mul. rewrite H1;rewrite H3.
+ assert (pow_pos rmul (NPEeval l e1) (p1 - p4) * NPEeval l p3 *
+ (pow_pos rmul (NPEeval l e1) p4 * NPEeval l p5) ==
+ pow_pos rmul (NPEeval l e1) p4 * pow_pos rmul (NPEeval l e1) (p1 - p4) *
+ NPEeval l p3 *NPEeval l p5) by ring. rewrite H;clear H.
+ rewrite <- pow_pos_plus. rewrite Pplus_minus.
+ split. symmetry;apply ARth.(ARmul_assoc). refine (refl_equal _). trivial.
+ repeat rewrite pow_th.(rpow_pow_N);simpl.
+ intros (H1,H2) (H3,H4).
+ unfold Zgt in H2, H4;simpl in H2,H4. rewrite H4 in H3;simpl in H3.
+ rewrite H2 in H1;simpl in H1.
+ assert (Zpos p1 > Zpos p6)%Z.
+ apply Zgt_trans with (Zpos p4). exact H4. exact H2.
+ unfold Zgt in H;simpl in H;rewrite H.
+ split. 2:exact H.
+ rewrite pow_pos_mul. simpl;rewrite H1;rewrite H3.
+ assert (pow_pos rmul (NPEeval l e1) (p1 - p4) * NPEeval l p3 *
+ (pow_pos rmul (NPEeval l e1) (p4 - p6) * NPEeval l p5) ==
+ pow_pos rmul (NPEeval l e1) (p1 - p4) * pow_pos rmul (NPEeval l e1) (p4 - p6) *
+ NPEeval l p3 * NPEeval l p5) by ring. rewrite H0;clear H0.
+ rewrite <- pow_pos_plus.
+ replace (p1 - p4 + (p4 - p6))%positive with (p1 - p6)%positive.
+ rewrite NPEmul_correct. simpl;ring.
+ assert
+ (Zpos p1 - Zpos p6 = Zpos p1 - Zpos p4 + (Zpos p4 - Zpos p6))%Z.
+ change ((Zpos p1 - Zpos p6)%Z = (Zpos p1 + (- Zpos p4) + (Zpos p4 +(- Zpos p6)))%Z).
+ rewrite <- Zplus_assoc. rewrite (Zplus_assoc (- Zpos p4)).
+ simpl. rewrite Pcompare_refl. simpl. reflexivity.
+ unfold Zminus, Zopp in H0. simpl in H0.
+ rewrite H2 in H0;rewrite H4 in H0;rewrite H in H0. inversion H0;trivial.
+ simpl. repeat rewrite pow_th.(rpow_pow_N).
+ intros H1 (H2,H3). unfold Zgt in H3;simpl in H3. rewrite H3 in H2;rewrite H3.
+ rewrite NPEmul_correct;simpl;rewrite NPEpow_correct;simpl.
+ simpl in H2. rewrite pow_th.(rpow_pow_N);simpl.
+ rewrite pow_pos_mul. split. ring [H2]. exact H3.
+ generalize (H0 p1 p2);clear H0;destruct (isIn e1 p1 p0 p2). destruct p3.
+ destruct n;simpl. rewrite NPEmul_correct;simpl;rewrite NPEpow_correct;simpl.
+ repeat rewrite pow_th.(rpow_pow_N);simpl.
+ intros (H1,H2);split;trivial. rewrite pow_pos_mul;ring [H1].
+ rewrite NPEmul_correct;simpl;rewrite NPEpow_correct;simpl.
+ repeat rewrite pow_th.(rpow_pow_N);simpl. rewrite pow_pos_mul.
+ intros (H1, H2);rewrite H1;split.
+ unfold Zgt in H2;simpl in H2;rewrite H2;rewrite H2 in H1.
+ simpl in H1;ring [H1]. trivial.
+ trivial.
+ destruct n. trivial.
+ generalize (H p1 (p0*p2)%positive);clear H;destruct (isIn e1 p1 p (p0*p2)). destruct p3.
+ destruct n;simpl. repeat rewrite pow_th.(rpow_pow_N). simpl.
+ intros (H1,H2);split. rewrite pow_pos_pow_pos. trivial. trivial.
+ repeat rewrite pow_th.(rpow_pow_N). simpl.
+ intros (H1,H2);split;trivial.
+ rewrite pow_pos_pow_pos;trivial.
+ trivial.
+Qed.
+
+Record rsplit : Type := mk_rsplit {
+ rsplit_left : PExpr C;
+ rsplit_common : PExpr C;
+ rsplit_right : PExpr C}.
+
+(* Stupid name clash *)
+Notation left := rsplit_left.
+Notation right := rsplit_right.
+Notation common := rsplit_common.
+
+Fixpoint split_aux (e1: PExpr C) (p:positive) (e2:PExpr C) {struct e1}: rsplit :=
+ match e1 with
+ | PEmul e3 e4 =>
+ let r1 := split_aux e3 p e2 in
+ let r2 := split_aux e4 p (right r1) in
+ mk_rsplit (NPEmul (left r1) (left r2))
+ (NPEmul (common r1) (common r2))
+ (right r2)
+ | PEpow e3 N0 => mk_rsplit (PEc cI) (PEc cI) e2
+ | PEpow e3 (Npos p3) => split_aux e3 (Pmult p3 p) e2
+ | _ =>
+ match isIn e1 p e2 xH with
+ | Some (N0,e3) => mk_rsplit (PEc cI) (NPEpow e1 (Npos p)) e3
+ | Some (Npos q, e3) => mk_rsplit (NPEpow e1 (Npos q)) (NPEpow e1 (Npos (p - q))) e3
+ | None => mk_rsplit (NPEpow e1 (Npos p)) (PEc cI) e2
+ end
+ end.
+
+Lemma split_aux_correct_1 : forall l e1 p e2,
+ let res := match isIn e1 p e2 xH with
+ | Some (N0,e3) => mk_rsplit (PEc cI) (NPEpow e1 (Npos p)) e3
+ | Some (Npos q, e3) => mk_rsplit (NPEpow e1 (Npos q)) (NPEpow e1 (Npos (p - q))) e3
+ | None => mk_rsplit (NPEpow e1 (Npos p)) (PEc cI) e2
+ end in
+ NPEeval l (PEpow e1 (Npos p)) == NPEeval l (NPEmul (left res) (common res))
+ /\
+ NPEeval l e2 == NPEeval l (NPEmul (right res) (common res)).
+Proof.
+ intros. unfold res;clear res; generalize (isIn_correct l e1 p e2 xH).
+ destruct (isIn e1 p e2 1). destruct p0.
+ Opaque NPEpow NPEmul.
+ destruct n;simpl;
+ (repeat rewrite NPEmul_correct;simpl;
+ repeat rewrite NPEpow_correct;simpl;
+ repeat rewrite pow_th.(rpow_pow_N);simpl).
+ intros (H, Hgt);split;try ring [H CRmorph.(morph1)].
+ intros (H, Hgt). unfold Zgt in Hgt;simpl in Hgt;rewrite Hgt in H.
+ simpl in H;split;try ring [H].
+ rewrite <- pow_pos_plus. rewrite Pplus_minus. reflexivity. trivial.
+ simpl;intros. repeat rewrite NPEmul_correct;simpl.
+ rewrite NPEpow_correct;simpl. split;ring [CRmorph.(morph1)].
+Qed.
+
+Theorem split_aux_correct: forall l e1 p e2,
+ NPEeval l (PEpow e1 (Npos p)) ==
+ NPEeval l (NPEmul (left (split_aux e1 p e2)) (common (split_aux e1 p e2)))
+/\
+ NPEeval l e2 == NPEeval l (NPEmul (right (split_aux e1 p e2))
+ (common (split_aux e1 p e2))).
+Proof.
+intros l; induction e1;intros k e2; try refine (split_aux_correct_1 l _ k e2);simpl.
+generalize (IHe1_1 k e2); clear IHe1_1.
+generalize (IHe1_2 k (rsplit_right (split_aux e1_1 k e2))); clear IHe1_2.
+simpl. repeat (rewrite NPEmul_correct;simpl).
+repeat rewrite pow_th.(rpow_pow_N);simpl.
+intros (H1,H2) (H3,H4);split.
+rewrite pow_pos_mul. rewrite H1;rewrite H3. ring.
+rewrite H4;rewrite H2;ring.
+destruct n;simpl.
+split. repeat rewrite pow_th.(rpow_pow_N);simpl.
+rewrite NPEmul_correct. simpl.
+ induction k;simpl;try ring [CRmorph.(morph1)]; ring [IHk CRmorph.(morph1)].
+ rewrite NPEmul_correct;simpl. ring [CRmorph.(morph1)].
+generalize (IHe1 (p*k)%positive e2);clear IHe1;simpl.
+repeat rewrite NPEmul_correct;simpl.
+repeat rewrite pow_th.(rpow_pow_N);simpl.
+rewrite pow_pos_pow_pos. intros [H1 H2];split;ring [H1 H2].
+Qed.
+
+Definition split e1 e2 := split_aux e1 xH e2.
+
+Theorem split_correct_l: forall l e1 e2,
+ NPEeval l e1 == NPEeval l (NPEmul (left (split e1 e2))
+ (common (split e1 e2))).
+Proof.
+intros l e1 e2; case (split_aux_correct l e1 xH e2);simpl.
+rewrite pow_th.(rpow_pow_N);simpl;auto.
+Qed.
+
+Theorem split_correct_r: forall l e1 e2,
+ NPEeval l e2 == NPEeval l (NPEmul (right (split e1 e2))
+ (common (split e1 e2))).
+Proof.
+intros l e1 e2; case (split_aux_correct l e1 xH e2);simpl;auto.
+Qed.
+
+Fixpoint Fnorm (e : FExpr) : linear :=
+ match e with
+ | FEc c => mk_linear (PEc c) (PEc cI) nil
+ | FEX x => mk_linear (PEX C x) (PEc cI) nil
+ | FEadd e1 e2 =>
+ let x := Fnorm e1 in
+ let y := Fnorm e2 in
+ let s := split (denum x) (denum y) in
+ mk_linear
+ (NPEadd (NPEmul (num x) (right s)) (NPEmul (num y) (left s)))
+ (NPEmul (left s) (NPEmul (right s) (common s)))
+ (condition x ++ condition y)
+
+ | FEsub e1 e2 =>
+ let x := Fnorm e1 in
+ let y := Fnorm e2 in
+ let s := split (denum x) (denum y) in
+ mk_linear
+ (NPEsub (NPEmul (num x) (right s)) (NPEmul (num y) (left s)))
+ (NPEmul (left s) (NPEmul (right s) (common s)))
+ (condition x ++ condition y)
+ | FEmul e1 e2 =>
+ let x := Fnorm e1 in
+ let y := Fnorm e2 in
+ let s1 := split (num x) (denum y) in
+ let s2 := split (num y) (denum x) in
+ mk_linear (NPEmul (left s1) (left s2))
+ (NPEmul (right s2) (right s1))
+ (condition x ++ condition y)
+ | FEopp e1 =>
+ let x := Fnorm e1 in
+ mk_linear (NPEopp (num x)) (denum x) (condition x)
+ | FEinv e1 =>
+ let x := Fnorm e1 in
+ mk_linear (denum x) (num x) (num x :: condition x)
+ | FEdiv e1 e2 =>
+ let x := Fnorm e1 in
+ let y := Fnorm e2 in
+ let s1 := split (num x) (num y) in
+ let s2 := split (denum x) (denum y) in
+ mk_linear (NPEmul (left s1) (right s2))
+ (NPEmul (left s2) (right s1))
+ (num y :: condition x ++ condition y)
+ | FEpow e1 n =>
+ let x := Fnorm e1 in
+ mk_linear (NPEpow (num x) n) (NPEpow (denum x) n) (condition x)
+ end.
+
+
+(* Example *)
+(*
+Eval compute
+ in (Fnorm
+ (FEdiv
+ (FEc cI)
+ (FEadd (FEinv (FEX xH%positive)) (FEinv (FEX (xO xH)%positive))))).
+*)
+
+ Lemma pow_pos_not_0 : forall x, ~x==0 -> forall p, ~pow_pos rmul x p == 0.
+Proof.
+ induction p;simpl.
+ intro Hp;assert (H1 := @rmul_reg_l _ (pow_pos rmul x p * pow_pos rmul x p) 0 H).
+ apply IHp.
+ rewrite (@rmul_reg_l _ (pow_pos rmul x p) 0 IHp).
+ reflexivity.
+ rewrite H1. ring. rewrite Hp;ring.
+ intro Hp;apply IHp. rewrite (@rmul_reg_l _ (pow_pos rmul x p) 0 IHp).
+ reflexivity. rewrite Hp;ring. trivial.
+Qed.
+
+Theorem Pcond_Fnorm:
+ forall l e,
+ PCond l (condition (Fnorm e)) -> ~ NPEeval l (denum (Fnorm e)) == 0.
+intros l e; elim e.
+ simpl in |- *; intros _ _; rewrite (morph1 CRmorph) in |- *; exact rI_neq_rO.
+ simpl in |- *; intros _ _; rewrite (morph1 CRmorph) in |- *; exact rI_neq_rO.
+ intros e1 Hrec1 e2 Hrec2 Hcond.
+ simpl condition in Hcond.
+ simpl denum in |- *.
+ rewrite NPEmul_correct in |- *.
+ simpl in |- *.
+ apply field_is_integral_domain.
+ intros HH; case Hrec1; auto.
+ apply PCond_app_inv_l with (1 := Hcond).
+ rewrite (split_correct_l l (denum (Fnorm e1)) (denum (Fnorm e2))).
+ rewrite NPEmul_correct; simpl; rewrite HH; ring.
+ intros HH; case Hrec2; auto.
+ apply PCond_app_inv_r with (1 := Hcond).
+ rewrite (split_correct_r l (denum (Fnorm e1)) (denum (Fnorm e2))); auto.
+ intros e1 Hrec1 e2 Hrec2 Hcond.
+ simpl condition in Hcond.
+ simpl denum in |- *.
+ rewrite NPEmul_correct in |- *.
+ simpl in |- *.
+ apply field_is_integral_domain.
+ intros HH; case Hrec1; auto.
+ apply PCond_app_inv_l with (1 := Hcond).
+ rewrite (split_correct_l l (denum (Fnorm e1)) (denum (Fnorm e2))).
+ rewrite NPEmul_correct; simpl; rewrite HH; ring.
+ intros HH; case Hrec2; auto.
+ apply PCond_app_inv_r with (1 := Hcond).
+ rewrite (split_correct_r l (denum (Fnorm e1)) (denum (Fnorm e2))); auto.
+ intros e1 Hrec1 e2 Hrec2 Hcond.
+ simpl condition in Hcond.
+ simpl denum in |- *.
+ rewrite NPEmul_correct in |- *.
+ simpl in |- *.
+ apply field_is_integral_domain.
+ intros HH; apply Hrec1.
+ apply PCond_app_inv_l with (1 := Hcond).
+ rewrite (split_correct_r l (num (Fnorm e2)) (denum (Fnorm e1))).
+ rewrite NPEmul_correct; simpl; rewrite HH; ring.
+ intros HH; apply Hrec2.
+ apply PCond_app_inv_r with (1 := Hcond).
+ rewrite (split_correct_r l (num (Fnorm e1)) (denum (Fnorm e2))).
+ rewrite NPEmul_correct; simpl; rewrite HH; ring.
+ intros e1 Hrec1 Hcond.
+ simpl condition in Hcond.
+ simpl denum in |- *.
+ auto.
+ intros e1 Hrec1 Hcond.
+ simpl condition in Hcond.
+ simpl denum in |- *.
+ apply PCond_cons_inv_l with (1:=Hcond).
+ intros e1 Hrec1 e2 Hrec2 Hcond.
+ simpl condition in Hcond.
+ simpl denum in |- *.
+ rewrite NPEmul_correct in |- *.
+ simpl in |- *.
+ apply field_is_integral_domain.
+ intros HH; apply Hrec1.
+ specialize PCond_cons_inv_r with (1:=Hcond); intro Hcond1.
+ apply PCond_app_inv_l with (1 := Hcond1).
+ rewrite (split_correct_l l (denum (Fnorm e1)) (denum (Fnorm e2))).
+ rewrite NPEmul_correct; simpl; rewrite HH; ring.
+ intros HH; apply PCond_cons_inv_l with (1:=Hcond).
+ rewrite (split_correct_r l (num (Fnorm e1)) (num (Fnorm e2))).
+ rewrite NPEmul_correct; simpl; rewrite HH; ring.
+ simpl;intros e1 Hrec1 n Hcond.
+ rewrite NPEpow_correct.
+ simpl;rewrite pow_th.(rpow_pow_N).
+ destruct n;simpl;intros.
+ apply AFth.(AF_1_neq_0). apply pow_pos_not_0;auto.
+Qed.
+Hint Resolve Pcond_Fnorm.
+
+
+(***************************************************************************
+
+ Main theorem
+
+ ***************************************************************************)
+
+Theorem Fnorm_FEeval_PEeval:
+ forall l fe,
+ PCond l (condition (Fnorm fe)) ->
+ FEeval l fe == NPEeval l (num (Fnorm fe)) / NPEeval l (denum (Fnorm fe)).
+Proof.
+intros l fe; elim fe; simpl.
+intros c H; rewrite CRmorph.(morph1); apply rdiv1.
+intros p H; rewrite CRmorph.(morph1); apply rdiv1.
+intros e1 He1 e2 He2 HH.
+assert (HH1: PCond l (condition (Fnorm e1))).
+apply PCond_app_inv_l with ( 1 := HH ).
+assert (HH2: PCond l (condition (Fnorm e2))).
+apply PCond_app_inv_r with ( 1 := HH ).
+rewrite (He1 HH1); rewrite (He2 HH2).
+rewrite NPEadd_correct; simpl.
+repeat rewrite NPEmul_correct; simpl.
+generalize (split_correct_l l (denum (Fnorm e1)) (denum (Fnorm e2)))
+ (split_correct_r l (denum (Fnorm e1)) (denum (Fnorm e2))).
+repeat rewrite NPEmul_correct; simpl.
+intros U1 U2; rewrite U1; rewrite U2.
+apply rdiv2b; auto.
+ rewrite <- U1; auto.
+ rewrite <- U2; auto.
+
+intros e1 He1 e2 He2 HH.
+assert (HH1: PCond l (condition (Fnorm e1))).
+apply PCond_app_inv_l with ( 1 := HH ).
+assert (HH2: PCond l (condition (Fnorm e2))).
+apply PCond_app_inv_r with ( 1 := HH ).
+rewrite (He1 HH1); rewrite (He2 HH2).
+rewrite NPEsub_correct; simpl.
+repeat rewrite NPEmul_correct; simpl.
+generalize (split_correct_l l (denum (Fnorm e1)) (denum (Fnorm e2)))
+ (split_correct_r l (denum (Fnorm e1)) (denum (Fnorm e2))).
+repeat rewrite NPEmul_correct; simpl.
+intros U1 U2; rewrite U1; rewrite U2.
+apply rdiv3b; auto.
+ rewrite <- U1; auto.
+ rewrite <- U2; auto.
+
+intros e1 He1 e2 He2 HH.
+assert (HH1: PCond l (condition (Fnorm e1))).
+apply PCond_app_inv_l with ( 1 := HH ).
+assert (HH2: PCond l (condition (Fnorm e2))).
+apply PCond_app_inv_r with ( 1 := HH ).
+rewrite (He1 HH1); rewrite (He2 HH2).
+repeat rewrite NPEmul_correct; simpl.
+generalize (split_correct_l l (num (Fnorm e1)) (denum (Fnorm e2)))
+ (split_correct_r l (num (Fnorm e1)) (denum (Fnorm e2)))
+ (split_correct_l l (num (Fnorm e2)) (denum (Fnorm e1)))
+ (split_correct_r l (num (Fnorm e2)) (denum (Fnorm e1))).
+repeat rewrite NPEmul_correct; simpl.
+intros U1 U2 U3 U4; rewrite U1; rewrite U2; rewrite U3;
+ rewrite U4; simpl.
+apply rdiv4b; auto.
+ rewrite <- U4; auto.
+ rewrite <- U2; auto.
+
+intros e1 He1 HH.
+rewrite NPEopp_correct; simpl; rewrite (He1 HH); apply rdiv5; auto.
+
+intros e1 He1 HH.
+assert (HH1: PCond l (condition (Fnorm e1))).
+apply PCond_cons_inv_r with ( 1 := HH ).
+rewrite (He1 HH1); apply rdiv6; auto.
+apply PCond_cons_inv_l with ( 1 := HH ).
+
+intros e1 He1 e2 He2 HH.
+assert (HH1: PCond l (condition (Fnorm e1))).
+apply PCond_app_inv_l with (condition (Fnorm e2)).
+apply PCond_cons_inv_r with ( 1 := HH ).
+assert (HH2: PCond l (condition (Fnorm e2))).
+apply PCond_app_inv_r with (condition (Fnorm e1)).
+apply PCond_cons_inv_r with ( 1 := HH ).
+rewrite (He1 HH1); rewrite (He2 HH2).
+repeat rewrite NPEmul_correct;simpl.
+generalize (split_correct_l l (num (Fnorm e1)) (num (Fnorm e2)))
+ (split_correct_r l (num (Fnorm e1)) (num (Fnorm e2)))
+ (split_correct_l l (denum (Fnorm e1)) (denum (Fnorm e2)))
+ (split_correct_r l (denum (Fnorm e1)) (denum (Fnorm e2))).
+repeat rewrite NPEmul_correct; simpl.
+intros U1 U2 U3 U4; rewrite U1; rewrite U2; rewrite U3;
+ rewrite U4; simpl.
+apply rdiv7b; auto.
+ rewrite <- U3; auto.
+ rewrite <- U2; auto.
+apply PCond_cons_inv_l with ( 1 := HH ).
+ rewrite <- U4; auto.
+
+intros e1 He1 n Hcond;assert (He1' := He1 Hcond);clear He1.
+repeat rewrite NPEpow_correct;simpl;repeat rewrite pow_th.(rpow_pow_N).
+rewrite He1';clear He1'.
+destruct n;simpl. apply rdiv1.
+generalize (NPEeval l (num (Fnorm e1))) (NPEeval l (denum (Fnorm e1)))
+ (Pcond_Fnorm _ _ Hcond).
+intros r r0 Hdiff;induction p;simpl.
+repeat (rewrite <- rdiv4;trivial).
+rewrite IHp. reflexivity.
+apply pow_pos_not_0;trivial.
+apply pow_pos_not_0;trivial.
+intro Hp. apply (pow_pos_not_0 Hdiff p).
+rewrite (@rmul_reg_l (pow_pos rmul r0 p) (pow_pos rmul r0 p) 0).
+ reflexivity. apply pow_pos_not_0;trivial. ring [Hp].
+rewrite <- rdiv4;trivial.
+rewrite IHp;reflexivity.
+apply pow_pos_not_0;trivial. apply pow_pos_not_0;trivial.
+reflexivity.
+Qed.
+
+Theorem Fnorm_crossproduct:
+ forall l fe1 fe2,
+ let nfe1 := Fnorm fe1 in
+ let nfe2 := Fnorm fe2 in
+ NPEeval l (PEmul (num nfe1) (denum nfe2)) ==
+ NPEeval l (PEmul (num nfe2) (denum nfe1)) ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ FEeval l fe1 == FEeval l fe2.
+intros l fe1 fe2 nfe1 nfe2 Hcrossprod Hcond; subst nfe1 nfe2.
+rewrite Fnorm_FEeval_PEeval in |- * by
+ apply PCond_app_inv_l with (1 := Hcond).
+ rewrite Fnorm_FEeval_PEeval in |- * by
+ apply PCond_app_inv_r with (1 := Hcond).
+ apply cross_product_eq; trivial.
+ apply Pcond_Fnorm.
+ apply PCond_app_inv_l with (1 := Hcond).
+ apply Pcond_Fnorm.
+ apply PCond_app_inv_r with (1 := Hcond).
+Qed.
+
+(* Correctness lemmas of reflexive tactics *)
+Notation Ninterp_PElist := (interp_PElist rO radd rmul rsub ropp req phi Cp_phi rpow).
+Notation Nmk_monpol_list := (mk_monpol_list cO cI cadd cmul csub copp ceqb cdiv).
+
+Theorem Fnorm_correct:
+ forall n l lpe fe,
+ Ninterp_PElist l lpe ->
+ Peq ceqb (Nnorm n (Nmk_monpol_list lpe) (num (Fnorm fe))) (Pc cO) = true ->
+ PCond l (condition (Fnorm fe)) -> FEeval l fe == 0.
+intros n l lpe fe Hlpe H H1;
+ apply eq_trans with (1 := Fnorm_FEeval_PEeval l fe H1).
+apply rdiv8; auto.
+transitivity (NPEeval l (PEc cO)); auto.
+rewrite (norm_subst_ok Rsth Reqe ARth CRmorph pow_th cdiv_th n l lpe);auto.
+change (NPEeval l (PEc cO)) with (Pphi 0 radd rmul phi l (Pc cO)).
+apply (Peq_ok Rsth Reqe CRmorph);auto.
+simpl. apply (morph0 CRmorph); auto.
+Qed.
+
+(* simplify a field expression into a fraction *)
+(* TODO: simplify when den is constant... *)
+Definition display_linear l num den :=
+ NPphi_dev l num / NPphi_dev l den.
+
+Definition display_pow_linear l num den :=
+ NPphi_pow l num / NPphi_pow l den.
+
+Theorem Field_rw_correct :
+ forall n lpe l,
+ Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall fe nfe, Fnorm fe = nfe ->
+ PCond l (condition nfe) ->
+ FEeval l fe == display_linear l (Nnorm n lmp (num nfe)) (Nnorm n lmp (denum nfe)).
+Proof.
+ intros n lpe l Hlpe lmp lmp_eq fe nfe eq_nfe H; subst nfe lmp.
+ apply eq_trans with (1 := Fnorm_FEeval_PEeval _ _ H).
+ unfold display_linear; apply SRdiv_ext;
+ eapply (ring_rw_correct Rsth Reqe ARth CRmorph);eauto.
+Qed.
+
+Theorem Field_rw_pow_correct :
+ forall n lpe l,
+ Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall fe nfe, Fnorm fe = nfe ->
+ PCond l (condition nfe) ->
+ FEeval l fe == display_pow_linear l (Nnorm n lmp (num nfe)) (Nnorm n lmp (denum nfe)).
+Proof.
+ intros n lpe l Hlpe lmp lmp_eq fe nfe eq_nfe H; subst nfe lmp.
+ apply eq_trans with (1 := Fnorm_FEeval_PEeval _ _ H).
+ unfold display_pow_linear; apply SRdiv_ext;
+ eapply (ring_rw_pow_correct Rsth Reqe ARth CRmorph);eauto.
+Qed.
+
+Theorem Field_correct :
+ forall n l lpe fe1 fe2, Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall nfe1, Fnorm fe1 = nfe1 ->
+ forall nfe2, Fnorm fe2 = nfe2 ->
+ Peq ceqb (Nnorm n lmp (PEmul (num nfe1) (denum nfe2)))
+ (Nnorm n lmp (PEmul (num nfe2) (denum nfe1))) = true ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ FEeval l fe1 == FEeval l fe2.
+Proof.
+intros n l lpe fe1 fe2 Hlpe lmp eq_lmp nfe1 eq1 nfe2 eq2 Hnorm Hcond; subst nfe1 nfe2 lmp.
+apply Fnorm_crossproduct; trivial.
+eapply (ring_correct Rsth Reqe ARth CRmorph); eauto.
+Qed.
+
+(* simplify a field equation : generate the crossproduct and simplify
+ polynomials *)
+Theorem Field_simplify_eq_old_correct :
+ forall l fe1 fe2 nfe1 nfe2,
+ Fnorm fe1 = nfe1 ->
+ Fnorm fe2 = nfe2 ->
+ NPphi_dev l (Nnorm O nil (PEmul (num nfe1) (denum nfe2))) ==
+ NPphi_dev l (Nnorm O nil (PEmul (num nfe2) (denum nfe1))) ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ FEeval l fe1 == FEeval l fe2.
+Proof.
+intros l fe1 fe2 nfe1 nfe2 eq1 eq2 Hcrossprod Hcond; subst nfe1 nfe2.
+apply Fnorm_crossproduct; trivial.
+match goal with
+ [ |- NPEeval l ?x == NPEeval l ?y] =>
+ rewrite (ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec
+ O nil l I (refl_equal nil) x (refl_equal (Nnorm O nil x)));
+ rewrite (ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec
+ O nil l I (refl_equal nil) y (refl_equal (Nnorm O nil y)))
+ end.
+trivial.
+Qed.
+
+Theorem Field_simplify_eq_correct :
+ forall n l lpe fe1 fe2,
+ Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall nfe1, Fnorm fe1 = nfe1 ->
+ forall nfe2, Fnorm fe2 = nfe2 ->
+ forall den, split (denum nfe1) (denum nfe2) = den ->
+ NPphi_dev l (Nnorm n lmp (PEmul (num nfe1) (right den))) ==
+ NPphi_dev l (Nnorm n lmp (PEmul (num nfe2) (left den))) ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ FEeval l fe1 == FEeval l fe2.
+Proof.
+intros n l lpe fe1 fe2 Hlpe lmp Hlmp nfe1 eq1 nfe2 eq2 den eq3 Hcrossprod Hcond;
+ subst nfe1 nfe2 den lmp.
+apply Fnorm_crossproduct; trivial.
+simpl in |- *.
+rewrite (split_correct_l l (denum (Fnorm fe1)) (denum (Fnorm fe2))) in |- *.
+rewrite (split_correct_r l (denum (Fnorm fe1)) (denum (Fnorm fe2))) in |- *.
+rewrite NPEmul_correct in |- *.
+rewrite NPEmul_correct in |- *.
+simpl in |- *.
+repeat rewrite (ARmul_assoc ARth) in |- *.
+rewrite <-(
+ let x := PEmul (num (Fnorm fe1))
+ (rsplit_right (split (denum (Fnorm fe1)) (denum (Fnorm fe2)))) in
+ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
+ Hlpe (refl_equal (Nmk_monpol_list lpe))
+ x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+rewrite <-(
+ let x := (PEmul (num (Fnorm fe2))
+ (rsplit_left
+ (split (denum (Fnorm fe1)) (denum (Fnorm fe2))))) in
+ ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
+ Hlpe (refl_equal (Nmk_monpol_list lpe))
+ x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+simpl in Hcrossprod.
+rewrite Hcrossprod in |- *.
+reflexivity.
+Qed.
+
+Theorem Field_simplify_eq_pow_correct :
+ forall n l lpe fe1 fe2,
+ Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall nfe1, Fnorm fe1 = nfe1 ->
+ forall nfe2, Fnorm fe2 = nfe2 ->
+ forall den, split (denum nfe1) (denum nfe2) = den ->
+ NPphi_pow l (Nnorm n lmp (PEmul (num nfe1) (right den))) ==
+ NPphi_pow l (Nnorm n lmp (PEmul (num nfe2) (left den))) ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ FEeval l fe1 == FEeval l fe2.
+Proof.
+intros n l lpe fe1 fe2 Hlpe lmp Hlmp nfe1 eq1 nfe2 eq2 den eq3 Hcrossprod Hcond;
+ subst nfe1 nfe2 den lmp.
+apply Fnorm_crossproduct; trivial.
+simpl in |- *.
+rewrite (split_correct_l l (denum (Fnorm fe1)) (denum (Fnorm fe2))) in |- *.
+rewrite (split_correct_r l (denum (Fnorm fe1)) (denum (Fnorm fe2))) in |- *.
+rewrite NPEmul_correct in |- *.
+rewrite NPEmul_correct in |- *.
+simpl in |- *.
+repeat rewrite (ARmul_assoc ARth) in |- *.
+rewrite <-(
+ let x := PEmul (num (Fnorm fe1))
+ (rsplit_right (split (denum (Fnorm fe1)) (denum (Fnorm fe2)))) in
+ring_rw_pow_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
+ Hlpe (refl_equal (Nmk_monpol_list lpe))
+ x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+rewrite <-(
+ let x := (PEmul (num (Fnorm fe2))
+ (rsplit_left
+ (split (denum (Fnorm fe1)) (denum (Fnorm fe2))))) in
+ ring_rw_pow_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
+ Hlpe (refl_equal (Nmk_monpol_list lpe))
+ x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+simpl in Hcrossprod.
+rewrite Hcrossprod in |- *.
+reflexivity.
+Qed.
+
+Theorem Field_simplify_eq_pow_in_correct :
+ forall n l lpe fe1 fe2,
+ Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall nfe1, Fnorm fe1 = nfe1 ->
+ forall nfe2, Fnorm fe2 = nfe2 ->
+ forall den, split (denum nfe1) (denum nfe2) = den ->
+ forall np1, Nnorm n lmp (PEmul (num nfe1) (right den)) = np1 ->
+ forall np2, Nnorm n lmp (PEmul (num nfe2) (left den)) = np2 ->
+ FEeval l fe1 == FEeval l fe2 ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ NPphi_pow l np1 ==
+ NPphi_pow l np2.
+Proof.
+ intros. subst nfe1 nfe2 lmp np1 np2.
+ repeat rewrite (Pphi_pow_ok Rsth Reqe ARth CRmorph pow_th get_sign_spec).
+ repeat (rewrite <- (norm_subst_ok Rsth Reqe ARth CRmorph pow_th);trivial). simpl.
+ assert (N1 := Pcond_Fnorm _ _ (PCond_app_inv_l _ _ _ H7)).
+ assert (N2 := Pcond_Fnorm _ _ (PCond_app_inv_r _ _ _ H7)).
+ apply (@rmul_reg_l (NPEeval l (rsplit_common den))).
+ intro Heq;apply N1.
+ rewrite (split_correct_l l (denum (Fnorm fe1)) (denum (Fnorm fe2))).
+ rewrite H3. rewrite NPEmul_correct. simpl. ring [Heq].
+ repeat rewrite (ARth.(ARmul_comm) (NPEeval l (rsplit_common den))).
+ repeat rewrite <- ARth.(ARmul_assoc).
+ change (NPEeval l (rsplit_right den) * NPEeval l (rsplit_common den)) with
+ (NPEeval l (PEmul (rsplit_right den) (rsplit_common den))).
+ change (NPEeval l (rsplit_left den) * NPEeval l (rsplit_common den)) with
+ (NPEeval l (PEmul (rsplit_left den) (rsplit_common den))).
+ repeat rewrite <- NPEmul_correct. rewrite <- H3. rewrite <- split_correct_l.
+ rewrite <- split_correct_r.
+ apply (@rmul_reg_l (/NPEeval l (denum (Fnorm fe2)))).
+ intro Heq; apply AFth.(AF_1_neq_0).
+ rewrite <- (@AFinv_l AFth (NPEeval l (denum (Fnorm fe2))));trivial.
+ ring [Heq]. rewrite (ARth.(ARmul_comm) (/ NPEeval l (denum (Fnorm fe2)))).
+ repeat rewrite <- (ARth.(ARmul_assoc)).
+ rewrite <- (AFth.(AFdiv_def)). rewrite rdiv_r_r by trivial.
+ apply (@rmul_reg_l (/NPEeval l (denum (Fnorm fe1)))).
+ intro Heq; apply AFth.(AF_1_neq_0).
+ rewrite <- (@AFinv_l AFth (NPEeval l (denum (Fnorm fe1))));trivial.
+ ring [Heq]. repeat rewrite (ARth.(ARmul_comm) (/ NPEeval l (denum (Fnorm fe1)))).
+ repeat rewrite <- (ARth.(ARmul_assoc)).
+ repeat rewrite <- (AFth.(AFdiv_def)). rewrite rdiv_r_r by trivial.
+ rewrite (AFth.(AFdiv_def)). ring_simplify. unfold SRopp.
+ rewrite (ARth.(ARmul_comm) (/ NPEeval l (denum (Fnorm fe2)))).
+ repeat rewrite <- (AFth.(AFdiv_def)).
+ repeat rewrite <- Fnorm_FEeval_PEeval ; trivial.
+ apply (PCond_app_inv_r _ _ _ H7). apply (PCond_app_inv_l _ _ _ H7).
+Qed.
+
+Theorem Field_simplify_eq_in_correct :
+forall n l lpe fe1 fe2,
+ Ninterp_PElist l lpe ->
+ forall lmp, Nmk_monpol_list lpe = lmp ->
+ forall nfe1, Fnorm fe1 = nfe1 ->
+ forall nfe2, Fnorm fe2 = nfe2 ->
+ forall den, split (denum nfe1) (denum nfe2) = den ->
+ forall np1, Nnorm n lmp (PEmul (num nfe1) (right den)) = np1 ->
+ forall np2, Nnorm n lmp (PEmul (num nfe2) (left den)) = np2 ->
+ FEeval l fe1 == FEeval l fe2 ->
+ PCond l (condition nfe1 ++ condition nfe2) ->
+ NPphi_dev l np1 ==
+ NPphi_dev l np2.
+Proof.
+ intros. subst nfe1 nfe2 lmp np1 np2.
+ repeat rewrite (Pphi_dev_ok Rsth Reqe ARth CRmorph get_sign_spec).
+ repeat (rewrite <- (norm_subst_ok Rsth Reqe ARth CRmorph pow_th);trivial). simpl.
+ assert (N1 := Pcond_Fnorm _ _ (PCond_app_inv_l _ _ _ H7)).
+ assert (N2 := Pcond_Fnorm _ _ (PCond_app_inv_r _ _ _ H7)).
+ apply (@rmul_reg_l (NPEeval l (rsplit_common den))).
+ intro Heq;apply N1.
+ rewrite (split_correct_l l (denum (Fnorm fe1)) (denum (Fnorm fe2))).
+ rewrite H3. rewrite NPEmul_correct. simpl. ring [Heq].
+ repeat rewrite (ARth.(ARmul_comm) (NPEeval l (rsplit_common den))).
+ repeat rewrite <- ARth.(ARmul_assoc).
+ change (NPEeval l (rsplit_right den) * NPEeval l (rsplit_common den)) with
+ (NPEeval l (PEmul (rsplit_right den) (rsplit_common den))).
+ change (NPEeval l (rsplit_left den) * NPEeval l (rsplit_common den)) with
+ (NPEeval l (PEmul (rsplit_left den) (rsplit_common den))).
+ repeat rewrite <- NPEmul_correct;rewrite <- H3. rewrite <- split_correct_l.
+ rewrite <- split_correct_r.
+ apply (@rmul_reg_l (/NPEeval l (denum (Fnorm fe2)))).
+ intro Heq; apply AFth.(AF_1_neq_0).
+ rewrite <- (@AFinv_l AFth (NPEeval l (denum (Fnorm fe2))));trivial.
+ ring [Heq]. rewrite (ARth.(ARmul_comm) (/ NPEeval l (denum (Fnorm fe2)))).
+ repeat rewrite <- (ARth.(ARmul_assoc)).
+ rewrite <- (AFth.(AFdiv_def)). rewrite rdiv_r_r by trivial.
+ apply (@rmul_reg_l (/NPEeval l (denum (Fnorm fe1)))).
+ intro Heq; apply AFth.(AF_1_neq_0).
+ rewrite <- (@AFinv_l AFth (NPEeval l (denum (Fnorm fe1))));trivial.
+ ring [Heq]. repeat rewrite (ARth.(ARmul_comm) (/ NPEeval l (denum (Fnorm fe1)))).
+ repeat rewrite <- (ARth.(ARmul_assoc)).
+ repeat rewrite <- (AFth.(AFdiv_def)). rewrite rdiv_r_r by trivial.
+ rewrite (AFth.(AFdiv_def)). ring_simplify. unfold SRopp.
+ rewrite (ARth.(ARmul_comm) (/ NPEeval l (denum (Fnorm fe2)))).
+ repeat rewrite <- (AFth.(AFdiv_def)).
+ repeat rewrite <- Fnorm_FEeval_PEeval;trivial.
+ apply (PCond_app_inv_r _ _ _ H7). apply (PCond_app_inv_l _ _ _ H7).
+Qed.
+
+
+Section Fcons_impl.
+
+Variable Fcons : PExpr C -> list (PExpr C) -> list (PExpr C).
+
+Hypothesis PCond_fcons_inv : forall l a l1,
+ PCond l (Fcons a l1) -> ~ NPEeval l a == 0 /\ PCond l l1.
+
+Fixpoint Fapp (l m:list (PExpr C)) {struct l} : list (PExpr C) :=
+ match l with
+ | nil => m
+ | cons a l1 => Fcons a (Fapp l1 m)
+ end.
+
+Lemma fcons_correct : forall l l1,
+ PCond l (Fapp l1 nil) -> PCond l l1.
+induction l1; simpl in |- *; intros.
+ trivial.
+ elim PCond_fcons_inv with (1 := H); intros.
+ destruct l1; auto.
+Qed.
+
+End Fcons_impl.
+
+Section Fcons_simpl.
+
+(* Some general simpifications of the condition: eliminate duplicates,
+ split multiplications *)
+
+Fixpoint Fcons (e:PExpr C) (l:list (PExpr C)) {struct l} : list (PExpr C) :=
+ match l with
+ nil => cons e nil
+ | cons a l1 => if PExpr_eq e a then l else cons a (Fcons e l1)
+ end.
+
+Theorem PFcons_fcons_inv:
+ forall l a l1, PCond l (Fcons a l1) -> ~ NPEeval l a == 0 /\ PCond l l1.
+intros l a l1; elim l1; simpl Fcons; auto.
+simpl; auto.
+intros a0 l0.
+generalize (PExpr_eq_semi_correct l a a0); case (PExpr_eq a a0).
+intros H H0 H1; split; auto.
+rewrite H; auto.
+generalize (PCond_cons_inv_l _ _ _ H1); simpl; auto.
+intros H H0 H1;
+ assert (Hp: ~ NPEeval l a0 == 0 /\ (~ NPEeval l a == 0 /\ PCond l l0)).
+split.
+generalize (PCond_cons_inv_l _ _ _ H1); simpl; auto.
+apply H0.
+generalize (PCond_cons_inv_r _ _ _ H1); simpl; auto.
+generalize Hp; case l0; simpl; intuition.
+Qed.
+
+(* equality of normal forms rather than syntactic equality *)
+Fixpoint Fcons0 (e:PExpr C) (l:list (PExpr C)) {struct l} : list (PExpr C) :=
+ match l with
+ nil => cons e nil
+ | cons a l1 =>
+ if Peq ceqb (Nnorm O nil e) (Nnorm O nil a) then l
+ else cons a (Fcons0 e l1)
+ end.
+
+Theorem PFcons0_fcons_inv:
+ forall l a l1, PCond l (Fcons0 a l1) -> ~ NPEeval l a == 0 /\ PCond l l1.
+intros l a l1; elim l1; simpl Fcons0; auto.
+simpl; auto.
+intros a0 l0.
+generalize (ring_correct Rsth Reqe ARth CRmorph pow_th cdiv_th O l nil a a0). simpl.
+ case (Peq ceqb (Nnorm O nil a) (Nnorm O nil a0)).
+intros H H0 H1; split; auto.
+rewrite H; auto.
+generalize (PCond_cons_inv_l _ _ _ H1); simpl; auto.
+intros H H0 H1;
+ assert (Hp: ~ NPEeval l a0 == 0 /\ (~ NPEeval l a == 0 /\ PCond l l0)).
+split.
+generalize (PCond_cons_inv_l _ _ _ H1); simpl; auto.
+apply H0.
+generalize (PCond_cons_inv_r _ _ _ H1); simpl; auto.
+clear get_sign get_sign_spec.
+generalize Hp; case l0; simpl; intuition.
+Qed.
+
+(* split factorized denominators *)
+Fixpoint Fcons00 (e:PExpr C) (l:list (PExpr C)) {struct e} : list (PExpr C) :=
+ match e with
+ PEmul e1 e2 => Fcons00 e1 (Fcons00 e2 l)
+ | PEpow e1 _ => Fcons00 e1 l
+ | _ => Fcons0 e l
+ end.
+
+Theorem PFcons00_fcons_inv:
+ forall l a l1, PCond l (Fcons00 a l1) -> ~ NPEeval l a == 0 /\ PCond l l1.
+intros l a; elim a; try (intros; apply PFcons0_fcons_inv; auto; fail).
+ intros p H p0 H0 l1 H1.
+ simpl in H1.
+ case (H _ H1); intros H2 H3.
+ case (H0 _ H3); intros H4 H5; split; auto.
+ simpl in |- *.
+ apply field_is_integral_domain; trivial.
+ simpl;intros. rewrite pow_th.(rpow_pow_N).
+ destruct (H _ H0);split;auto.
+ destruct n;simpl. apply AFth.(AF_1_neq_0).
+ apply pow_pos_not_0;trivial.
+Qed.
+
+Definition Pcond_simpl_gen :=
+ fcons_correct _ PFcons00_fcons_inv.
+
+
+(* Specific case when the equality test of coefs is complete w.r.t. the
+ field equality: non-zero coefs can be eliminated, and opposite can
+ be simplified (if -1 <> 0) *)
+
+Hypothesis ceqb_complete : forall c1 c2, phi c1 == phi c2 -> ceqb c1 c2 = true.
+
+Lemma ceqb_rect_complete : forall c1 c2 (A:Type) (x y:A) (P:A->Type),
+ (phi c1 == phi c2 -> P x) ->
+ (~ phi c1 == phi c2 -> P y) ->
+ P (if ceqb c1 c2 then x else y).
+Proof.
+intros.
+generalize (fun h => X (morph_eq CRmorph c1 c2 h)).
+generalize (@ceqb_complete c1 c2).
+case (c1 ?=! c2); auto; intros.
+apply X0.
+red in |- *; intro.
+absurd (false = true); auto; discriminate.
+Qed.
+
+Fixpoint Fcons1 (e:PExpr C) (l:list (PExpr C)) {struct e} : list (PExpr C) :=
+ match e with
+ PEmul e1 e2 => Fcons1 e1 (Fcons1 e2 l)
+ | PEpow e _ => Fcons1 e l
+ | PEopp e => if ceqb (copp cI) cO then absurd_PCond else Fcons1 e l
+ | PEc c => if ceqb c cO then absurd_PCond else l
+ | _ => Fcons0 e l
+ end.
+
+Theorem PFcons1_fcons_inv:
+ forall l a l1, PCond l (Fcons1 a l1) -> ~ NPEeval l a == 0 /\ PCond l l1.
+intros l a; elim a; try (intros; apply PFcons0_fcons_inv; auto; fail).
+ simpl in |- *; intros c l1.
+ apply ceqb_rect_complete; intros.
+ elim (@absurd_PCond_bottom l H0).
+ split; trivial.
+ rewrite <- (morph0 CRmorph) in |- *; trivial.
+ intros p H p0 H0 l1 H1.
+ simpl in H1.
+ case (H _ H1); intros H2 H3.
+ case (H0 _ H3); intros H4 H5; split; auto.
+ simpl in |- *.
+ apply field_is_integral_domain; trivial.
+ simpl in |- *; intros p H l1.
+ apply ceqb_rect_complete; intros.
+ elim (@absurd_PCond_bottom l H1).
+ destruct (H _ H1).
+ split; trivial.
+ apply ropp_neq_0; trivial.
+ rewrite (morph_opp CRmorph) in H0.
+ rewrite (morph1 CRmorph) in H0.
+ rewrite (morph0 CRmorph) in H0.
+ trivial.
+ intros;simpl. destruct (H _ H0);split;trivial.
+ rewrite pow_th.(rpow_pow_N). destruct n;simpl.
+ apply AFth.(AF_1_neq_0). apply pow_pos_not_0;trivial.
+Qed.
+
+Definition Fcons2 e l := Fcons1 (PExpr_simp e) l.
+
+Theorem PFcons2_fcons_inv:
+ forall l a l1, PCond l (Fcons2 a l1) -> ~ NPEeval l a == 0 /\ PCond l l1.
+unfold Fcons2 in |- *; intros l a l1 H; split;
+ case (PFcons1_fcons_inv l (PExpr_simp a) l1); auto.
+intros H1 H2 H3; case H1.
+transitivity (NPEeval l a); trivial.
+apply PExpr_simp_correct.
+Qed.
+
+Definition Pcond_simpl_complete :=
+ fcons_correct _ PFcons2_fcons_inv.
+
+End Fcons_simpl.
+
+End AlmostField.
+
+Section FieldAndSemiField.
+
+ Record field_theory : Prop := mk_field {
+ F_R : ring_theory rO rI radd rmul rsub ropp req;
+ F_1_neq_0 : ~ 1 == 0;
+ Fdiv_def : forall p q, p / q == p * / q;
+ Finv_l : forall p, ~ p == 0 -> / p * p == 1
+ }.
+
+ Definition F2AF f :=
+ mk_afield
+ (Rth_ARth Rsth Reqe f.(F_R)) f.(F_1_neq_0) f.(Fdiv_def) f.(Finv_l).
+
+ Record semi_field_theory : Prop := mk_sfield {
+ SF_SR : semi_ring_theory rO rI radd rmul req;
+ SF_1_neq_0 : ~ 1 == 0;
+ SFdiv_def : forall p q, p / q == p * / q;
+ SFinv_l : forall p, ~ p == 0 -> / p * p == 1
+ }.
+
+End FieldAndSemiField.
+
+End MakeFieldPol.
+
+ Definition SF2AF R (rO rI:R) radd rmul rdiv rinv req Rsth
+ (sf:semi_field_theory rO rI radd rmul rdiv rinv req) :=
+ mk_afield _ _
+ (SRth_ARth Rsth sf.(SF_SR))
+ sf.(SF_1_neq_0)
+ sf.(SFdiv_def)
+ sf.(SFinv_l).
+
+
+Section Complete.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable (rdiv : R -> R -> R) (rinv : R -> R).
+ Variable req : R -> R -> Prop.
+ 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 " := (rdiv x y). Notation "/ x" := (rinv x).
+ Notation "x == y" := (req x y) (at level 70, no associativity).
+ Variable Rsth : Setoid_Theory R req.
+ Add Setoid R req Rsth as R_setoid3.
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Add Morphism radd : radd_ext3. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul : rmul_ext3. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp : ropp_ext3. exact (Ropp_ext Reqe). Qed.
+
+Section AlmostField.
+
+ Variable AFth : almost_field_theory rO rI radd rmul rsub ropp rdiv rinv req.
+ Let ARth := AFth.(AF_AR).
+ Let rI_neq_rO := AFth.(AF_1_neq_0).
+ Let rdiv_def := AFth.(AFdiv_def).
+ Let rinv_l := AFth.(AFinv_l).
+
+Hypothesis S_inj : forall x y, 1+x==1+y -> x==y.
+
+Hypothesis gen_phiPOS_not_0 : forall p, ~ gen_phiPOS1 rI radd rmul p == 0.
+
+Lemma add_inj_r : forall p x y,
+ gen_phiPOS1 rI radd rmul p + x == gen_phiPOS1 rI radd rmul p + y -> x==y.
+intros p x y.
+elim p using Pind; simpl in |- *; intros.
+ apply S_inj; trivial.
+ apply H.
+ apply S_inj.
+ repeat rewrite (ARadd_assoc ARth) in |- *.
+ rewrite <- (ARgen_phiPOS_Psucc Rsth Reqe ARth) in |- *; trivial.
+Qed.
+
+Lemma gen_phiPOS_inj : forall x y,
+ gen_phiPOS rI radd rmul x == gen_phiPOS rI radd rmul y ->
+ x = y.
+intros x y.
+repeat rewrite <- (same_gen Rsth Reqe ARth) in |- *.
+ElimPcompare x y; intro.
+ intros.
+ apply Pcompare_Eq_eq; trivial.
+ intro.
+ elim gen_phiPOS_not_0 with (y - x)%positive.
+ apply add_inj_r with x.
+ symmetry in |- *.
+ rewrite (ARadd_0_r Rsth ARth) in |- *.
+ rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth) in |- *.
+ rewrite Pplus_minus in |- *; trivial.
+ change Eq with (CompOpp Eq) in |- *.
+ rewrite <- Pcompare_antisym in |- *; trivial.
+ rewrite H in |- *; trivial.
+ intro.
+ elim gen_phiPOS_not_0 with (x - y)%positive.
+ apply add_inj_r with y.
+ rewrite (ARadd_0_r Rsth ARth) in |- *.
+ rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth) in |- *.
+ rewrite Pplus_minus in |- *; trivial.
+Qed.
+
+
+Lemma gen_phiN_inj : forall x y,
+ gen_phiN rO rI radd rmul x == gen_phiN rO rI radd rmul y ->
+ x = y.
+destruct x; destruct y; simpl in |- *; intros; trivial.
+ elim gen_phiPOS_not_0 with p.
+ symmetry in |- *.
+ rewrite (same_gen Rsth Reqe ARth) in |- *; trivial.
+ elim gen_phiPOS_not_0 with p.
+ rewrite (same_gen Rsth Reqe ARth) in |- *; trivial.
+ rewrite gen_phiPOS_inj with (1 := H) in |- *; trivial.
+Qed.
+
+Lemma gen_phiN_complete : forall x y,
+ gen_phiN rO rI radd rmul x == gen_phiN rO rI radd rmul y ->
+ Neq_bool x y = true.
+intros.
+ replace y with x.
+ unfold Neq_bool in |- *.
+ rewrite Ncompare_refl in |- *; trivial.
+ apply gen_phiN_inj; trivial.
+Qed.
+
+End AlmostField.
+
+Section Field.
+
+ Variable Fth : field_theory rO rI radd rmul rsub ropp rdiv rinv req.
+ Let Rth := Fth.(F_R).
+ Let rI_neq_rO := Fth.(F_1_neq_0).
+ Let rdiv_def := Fth.(Fdiv_def).
+ Let rinv_l := Fth.(Finv_l).
+ Let AFth := F2AF Rsth Reqe Fth.
+ Let ARth := Rth_ARth Rsth Reqe Rth.
+
+Lemma ring_S_inj : forall x y, 1+x==1+y -> x==y.
+intros.
+transitivity (x + (1 + - (1))).
+ rewrite (Ropp_def Rth) in |- *.
+ symmetry in |- *.
+ apply (ARadd_0_r Rsth ARth).
+ transitivity (y + (1 + - (1))).
+ repeat rewrite <- (ARplus_assoc ARth) in |- *.
+ repeat rewrite (ARadd_assoc ARth) in |- *.
+ apply (Radd_ext Reqe).
+ repeat rewrite <- (ARadd_comm ARth 1) in |- *.
+ trivial.
+ reflexivity.
+ rewrite (Ropp_def Rth) in |- *.
+ apply (ARadd_0_r Rsth ARth).
+Qed.
+
+
+ Hypothesis gen_phiPOS_not_0 : forall p, ~ gen_phiPOS1 rI radd rmul p == 0.
+
+Let gen_phiPOS_inject :=
+ gen_phiPOS_inj AFth ring_S_inj gen_phiPOS_not_0.
+
+Lemma gen_phiPOS_discr_sgn : forall x y,
+ ~ gen_phiPOS rI radd rmul x == - gen_phiPOS rI radd rmul y.
+red in |- *; intros.
+apply gen_phiPOS_not_0 with (y + x)%positive.
+rewrite (ARgen_phiPOS_add Rsth Reqe ARth) in |- *.
+transitivity (gen_phiPOS1 1 radd rmul y + - gen_phiPOS1 1 radd rmul y).
+ apply (Radd_ext Reqe); trivial.
+ reflexivity.
+ rewrite (same_gen Rsth Reqe ARth) in |- *.
+ rewrite (same_gen Rsth Reqe ARth) in |- *.
+ trivial.
+ apply (Ropp_def Rth).
+Qed.
+
+Lemma gen_phiZ_inj : forall x y,
+ gen_phiZ rO rI radd rmul ropp x == gen_phiZ rO rI radd rmul ropp y ->
+ x = y.
+destruct x; destruct y; simpl in |- *; intros.
+ trivial.
+ elim gen_phiPOS_not_0 with p.
+ rewrite (same_gen Rsth Reqe ARth) in |- *.
+ symmetry in |- *; trivial.
+ elim gen_phiPOS_not_0 with p.
+ rewrite (same_gen Rsth Reqe ARth) in |- *.
+ rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p)) in |- *.
+ rewrite <- H in |- *.
+ apply (ARopp_zero Rsth Reqe ARth).
+ elim gen_phiPOS_not_0 with p.
+ rewrite (same_gen Rsth Reqe ARth) in |- *.
+ trivial.
+ rewrite gen_phiPOS_inject with (1 := H) in |- *; trivial.
+ elim gen_phiPOS_discr_sgn with (1 := H).
+ elim gen_phiPOS_not_0 with p.
+ rewrite (same_gen Rsth Reqe ARth) in |- *.
+ rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p)) in |- *.
+ rewrite H in |- *.
+ apply (ARopp_zero Rsth Reqe ARth).
+ elim gen_phiPOS_discr_sgn with p0 p.
+ symmetry in |- *; trivial.
+ replace p0 with p; trivial.
+ apply gen_phiPOS_inject.
+ rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p)) in |- *.
+ rewrite <- (Ropp_opp Rsth Reqe Rth (gen_phiPOS 1 radd rmul p0)) in |- *.
+ rewrite H in |- *; trivial.
+ reflexivity.
+Qed.
+
+Lemma gen_phiZ_complete : forall x y,
+ gen_phiZ rO rI radd rmul ropp x == gen_phiZ rO rI radd rmul ropp y ->
+ Zeq_bool x y = true.
+intros.
+ replace y with x.
+ unfold Zeq_bool in |- *.
+ rewrite Zcompare_refl in |- *; trivial.
+ apply gen_phiZ_inj; trivial.
+Qed.
+
+End Field.
+
+End Complete.
diff --git a/plugins/setoid_ring/InitialRing.v b/plugins/setoid_ring/InitialRing.v
new file mode 100644
index 00000000..b5384f80
--- /dev/null
+++ b/plugins/setoid_ring/InitialRing.v
@@ -0,0 +1,908 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import ZArith_base.
+Require Import Zpow_def.
+Require Import BinInt.
+Require Import BinNat.
+Require Import Setoid.
+Require Import Ring_theory.
+Require Import Ring_polynom.
+Require Import ZOdiv_def.
+Import List.
+
+Set Implicit Arguments.
+
+Import RingSyntax.
+
+(* An object to return when an expression is not recognized as a constant *)
+Definition NotConstant := false.
+
+(** Z is a ring and a setoid*)
+
+Lemma Zsth : Setoid_Theory Z (@eq Z).
+Proof (Eqsth Z).
+
+Lemma Zeqe : ring_eq_ext Zplus Zmult Zopp (@eq Z).
+Proof (Eq_ext Zplus Zmult Zopp).
+
+Lemma Zth : ring_theory Z0 (Zpos xH) Zplus Zmult Zminus Zopp (@eq Z).
+Proof.
+ constructor. exact Zplus_0_l. exact Zplus_comm. exact Zplus_assoc.
+ exact Zmult_1_l. exact Zmult_comm. exact Zmult_assoc.
+ exact Zmult_plus_distr_l. trivial. exact Zminus_diag.
+Qed.
+
+(** Two generic morphisms from Z to (abrbitrary) rings, *)
+(**second one is more convenient for proofs but they are ext. equal*)
+Section ZMORPHISM.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable req : R -> R -> Prop.
+ 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).
+ Variable Rsth : Setoid_Theory R req.
+ Add Setoid R req Rsth as R_setoid3.
+ Ltac rrefl := gen_reflexivity Rsth.
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Add Morphism radd : radd_ext3. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul : rmul_ext3. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp : ropp_ext3. exact (Ropp_ext Reqe). Qed.
+
+ Fixpoint gen_phiPOS1 (p:positive) : R :=
+ match p with
+ | xH => 1
+ | xO p => (1 + 1) * (gen_phiPOS1 p)
+ | xI p => 1 + ((1 + 1) * (gen_phiPOS1 p))
+ end.
+
+ Fixpoint gen_phiPOS (p:positive) : R :=
+ match p with
+ | xH => 1
+ | xO xH => (1 + 1)
+ | xO p => (1 + 1) * (gen_phiPOS p)
+ | xI xH => 1 + (1 +1)
+ | xI p => 1 + ((1 + 1) * (gen_phiPOS p))
+ end.
+
+ Definition gen_phiZ1 z :=
+ match z with
+ | Zpos p => gen_phiPOS1 p
+ | Z0 => 0
+ | Zneg p => -(gen_phiPOS1 p)
+ end.
+
+ Definition gen_phiZ z :=
+ match z with
+ | Zpos p => gen_phiPOS p
+ | Z0 => 0
+ | Zneg p => -(gen_phiPOS p)
+ end.
+ Notation "[ x ]" := (gen_phiZ x).
+
+ Definition get_signZ z :=
+ match z with
+ | Zneg p => Some (Zpos p)
+ | _ => None
+ end.
+
+ Lemma get_signZ_th : sign_theory Zopp Zeq_bool get_signZ.
+ Proof.
+ constructor.
+ destruct c;intros;try discriminate.
+ injection H;clear H;intros H1;subst c'.
+ simpl. unfold Zeq_bool. rewrite Zcompare_refl. trivial.
+ Qed.
+
+
+ Section ALMOST_RING.
+ Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
+ Add Morphism rsub : rsub_ext3. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Ltac norm := gen_srewrite Rsth Reqe ARth.
+ Ltac add_push := gen_add_push radd Rsth Reqe ARth.
+
+ Lemma same_gen : forall x, gen_phiPOS1 x == gen_phiPOS x.
+ Proof.
+ induction x;simpl.
+ rewrite IHx;destruct x;simpl;norm.
+ rewrite IHx;destruct x;simpl;norm.
+ rrefl.
+ Qed.
+
+ Lemma ARgen_phiPOS_Psucc : forall x,
+ gen_phiPOS1 (Psucc x) == 1 + (gen_phiPOS1 x).
+ Proof.
+ induction x;simpl;norm.
+ rewrite IHx;norm.
+ add_push 1;rrefl.
+ Qed.
+
+ Lemma ARgen_phiPOS_add : forall x y,
+ gen_phiPOS1 (x + y) == (gen_phiPOS1 x) + (gen_phiPOS1 y).
+ Proof.
+ induction x;destruct y;simpl;norm.
+ rewrite Pplus_carry_spec.
+ rewrite ARgen_phiPOS_Psucc.
+ rewrite IHx;norm.
+ add_push (gen_phiPOS1 y);add_push 1;rrefl.
+ rewrite IHx;norm;add_push (gen_phiPOS1 y);rrefl.
+ rewrite ARgen_phiPOS_Psucc;norm;add_push 1;rrefl.
+ rewrite IHx;norm;add_push(gen_phiPOS1 y); add_push 1;rrefl.
+ rewrite IHx;norm;add_push(gen_phiPOS1 y);rrefl.
+ add_push 1;rrefl.
+ rewrite ARgen_phiPOS_Psucc;norm;add_push 1;rrefl.
+ Qed.
+
+ Lemma ARgen_phiPOS_mult :
+ forall x y, gen_phiPOS1 (x * y) == gen_phiPOS1 x * gen_phiPOS1 y.
+ Proof.
+ induction x;intros;simpl;norm.
+ rewrite ARgen_phiPOS_add;simpl;rewrite IHx;norm.
+ rewrite IHx;rrefl.
+ Qed.
+
+ End ALMOST_RING.
+
+ Variable Rth : ring_theory 0 1 radd rmul rsub ropp req.
+ Let ARth := Rth_ARth Rsth Reqe Rth.
+ Add Morphism rsub : rsub_ext4. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Ltac norm := gen_srewrite Rsth Reqe ARth.
+ Ltac add_push := gen_add_push radd Rsth Reqe ARth.
+
+(*morphisms are extensionaly equal*)
+ Lemma same_genZ : forall x, [x] == gen_phiZ1 x.
+ Proof.
+ destruct x;simpl; try rewrite (same_gen ARth);rrefl.
+ Qed.
+
+ Lemma gen_Zeqb_ok : forall x y,
+ Zeq_bool x y = true -> [x] == [y].
+ Proof.
+ intros x y H.
+ assert (H1 := Zeq_bool_eq x y H);unfold IDphi in H1.
+ rewrite H1;rrefl.
+ Qed.
+
+ Lemma gen_phiZ1_add_pos_neg : forall x y,
+ gen_phiZ1
+ match (x ?= y)%positive Eq with
+ | Eq => Z0
+ | Lt => Zneg (y - x)
+ | Gt => Zpos (x - y)
+ end
+ == gen_phiPOS1 x + -gen_phiPOS1 y.
+ Proof.
+ intros x y.
+ assert (H:= (Pcompare_Eq_eq x y)); assert (H0 := Pminus_mask_Gt x y).
+ generalize (Pminus_mask_Gt y x).
+ replace Eq with (CompOpp Eq);[intro H1;simpl|trivial].
+ rewrite <- Pcompare_antisym in H1.
+ destruct ((x ?= y)%positive Eq).
+ rewrite H;trivial. rewrite (Ropp_def Rth);rrefl.
+ destruct H1 as [h [Heq1 [Heq2 Hor]]];trivial.
+ unfold Pminus; rewrite Heq1;rewrite <- Heq2.
+ rewrite (ARgen_phiPOS_add ARth);simpl;norm.
+ rewrite (Ropp_def Rth);norm.
+ destruct H0 as [h [Heq1 [Heq2 Hor]]];trivial.
+ unfold Pminus; rewrite Heq1;rewrite <- Heq2.
+ rewrite (ARgen_phiPOS_add ARth);simpl;norm.
+ add_push (gen_phiPOS1 h);rewrite (Ropp_def Rth); norm.
+ Qed.
+
+ Lemma match_compOpp : forall x (B:Type) (be bl bg:B),
+ match CompOpp x with Eq => be | Lt => bl | Gt => bg end
+ = match x with Eq => be | Lt => bg | Gt => bl end.
+ Proof. destruct x;simpl;intros;trivial. Qed.
+
+ Lemma gen_phiZ_add : forall x y, [x + y] == [x] + [y].
+ Proof.
+ intros x y; repeat rewrite same_genZ; generalize x y;clear x y.
+ induction x;destruct y;simpl;norm.
+ apply (ARgen_phiPOS_add ARth).
+ apply gen_phiZ1_add_pos_neg.
+ replace Eq with (CompOpp Eq);trivial.
+ rewrite <- Pcompare_antisym;simpl.
+ rewrite match_compOpp.
+ rewrite (Radd_comm Rth).
+ apply gen_phiZ1_add_pos_neg.
+ rewrite (ARgen_phiPOS_add ARth); norm.
+ Qed.
+
+ Lemma gen_phiZ_mul : forall x y, [x * y] == [x] * [y].
+ Proof.
+ intros x y;repeat rewrite same_genZ.
+ destruct x;destruct y;simpl;norm;
+ rewrite (ARgen_phiPOS_mult ARth);try (norm;fail).
+ rewrite (Ropp_opp Rsth Reqe Rth);rrefl.
+ Qed.
+
+ Lemma gen_phiZ_ext : forall x y : Z, x = y -> [x] == [y].
+ Proof. intros;subst;rrefl. Qed.
+
+(*proof that [.] satisfies morphism specifications*)
+ Lemma gen_phiZ_morph :
+ ring_morph 0 1 radd rmul rsub ropp req Z0 (Zpos xH)
+ Zplus Zmult Zminus Zopp Zeq_bool gen_phiZ.
+ Proof.
+ assert ( SRmorph : semi_morph 0 1 radd rmul req Z0 (Zpos xH)
+ Zplus Zmult Zeq_bool gen_phiZ).
+ apply mkRmorph;simpl;try rrefl.
+ apply gen_phiZ_add. apply gen_phiZ_mul. apply gen_Zeqb_ok.
+ apply (Smorph_morph Rsth Reqe Rth Zth SRmorph gen_phiZ_ext).
+ Qed.
+
+End ZMORPHISM.
+
+(** N is a semi-ring and a setoid*)
+Lemma Nsth : Setoid_Theory N (@eq N).
+Proof (Eqsth N).
+
+Lemma Nseqe : sring_eq_ext Nplus Nmult (@eq N).
+Proof (Eq_s_ext Nplus Nmult).
+
+Lemma Nth : semi_ring_theory N0 (Npos xH) Nplus Nmult (@eq N).
+Proof.
+ constructor. exact Nplus_0_l. exact Nplus_comm. exact Nplus_assoc.
+ exact Nmult_1_l. exact Nmult_0_l. exact Nmult_comm. exact Nmult_assoc.
+ exact Nmult_plus_distr_r.
+Qed.
+
+Definition Nsub := SRsub Nplus.
+Definition Nopp := (@SRopp N).
+
+Lemma Neqe : ring_eq_ext Nplus Nmult Nopp (@eq N).
+Proof (SReqe_Reqe Nseqe).
+
+Lemma Nath :
+ almost_ring_theory N0 (Npos xH) Nplus Nmult Nsub Nopp (@eq N).
+Proof (SRth_ARth Nsth Nth).
+
+Definition Neq_bool (x y:N) :=
+ match Ncompare x y with
+ | Eq => true
+ | _ => false
+ end.
+
+Lemma Neq_bool_ok : forall x y, Neq_bool x y = true -> x = y.
+ Proof.
+ intros x y;unfold Neq_bool.
+ assert (H:=Ncompare_Eq_eq x y);
+ destruct (Ncompare x y);intros;try discriminate.
+ rewrite H;trivial.
+ Qed.
+
+Lemma Neq_bool_complete : forall x y, Neq_bool x y = true -> x = y.
+ Proof.
+ intros x y;unfold Neq_bool.
+ assert (H:=Ncompare_Eq_eq x y);
+ destruct (Ncompare x y);intros;try discriminate.
+ rewrite H;trivial.
+ Qed.
+
+(**Same as above : definition of two,extensionaly equal, generic morphisms *)
+(**from N to any semi-ring*)
+Section NMORPHISM.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul: R->R->R).
+ Variable req : R -> R -> Prop.
+ Notation "0" := rO. Notation "1" := rI.
+ Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y).
+ Variable Rsth : Setoid_Theory R req.
+ Add Setoid R req Rsth as R_setoid4.
+ Ltac rrefl := gen_reflexivity Rsth.
+ Variable SReqe : sring_eq_ext radd rmul req.
+ Variable SRth : semi_ring_theory 0 1 radd rmul req.
+ Let ARth := SRth_ARth Rsth SRth.
+ Let Reqe := SReqe_Reqe SReqe.
+ Let ropp := (@SRopp R).
+ Let rsub := (@SRsub R radd).
+ Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
+ Notation "x == y" := (req x y).
+ Add Morphism radd : radd_ext4. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul : rmul_ext4. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp : ropp_ext4. exact (Ropp_ext Reqe). Qed.
+ Add Morphism rsub : rsub_ext5. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Ltac norm := gen_srewrite Rsth Reqe ARth.
+
+ Definition gen_phiN1 x :=
+ match x with
+ | N0 => 0
+ | Npos x => gen_phiPOS1 1 radd rmul x
+ end.
+
+ Definition gen_phiN x :=
+ match x with
+ | N0 => 0
+ | Npos x => gen_phiPOS 1 radd rmul x
+ end.
+ Notation "[ x ]" := (gen_phiN x).
+
+ Lemma same_genN : forall x, [x] == gen_phiN1 x.
+ Proof.
+ destruct x;simpl. rrefl.
+ rewrite (same_gen Rsth Reqe ARth);rrefl.
+ Qed.
+
+ Lemma gen_phiN_add : forall x y, [x + y] == [x] + [y].
+ Proof.
+ intros x y;repeat rewrite same_genN.
+ destruct x;destruct y;simpl;norm.
+ apply (ARgen_phiPOS_add Rsth Reqe ARth).
+ Qed.
+
+ Lemma gen_phiN_mult : forall x y, [x * y] == [x] * [y].
+ Proof.
+ intros x y;repeat rewrite same_genN.
+ destruct x;destruct y;simpl;norm.
+ apply (ARgen_phiPOS_mult Rsth Reqe ARth).
+ Qed.
+
+ Lemma gen_phiN_sub : forall x y, [Nsub x y] == [x] - [y].
+ Proof. exact gen_phiN_add. Qed.
+
+(*gen_phiN satisfies morphism specifications*)
+ Lemma gen_phiN_morph : ring_morph 0 1 radd rmul rsub ropp req
+ N0 (Npos xH) Nplus Nmult Nsub Nopp Neq_bool gen_phiN.
+ Proof.
+ constructor;intros;simpl; try rrefl.
+ apply gen_phiN_add. apply gen_phiN_sub. apply gen_phiN_mult.
+ rewrite (Neq_bool_ok x y);trivial. rrefl.
+ Qed.
+
+End NMORPHISM.
+
+(* Words on N : initial structure for almost-rings. *)
+Definition Nword := list N.
+Definition NwO : Nword := nil.
+Definition NwI : Nword := 1%N :: nil.
+
+Definition Nwcons n (w : Nword) : Nword :=
+ match w, n with
+ | nil, 0%N => nil
+ | _, _ => n :: w
+ end.
+
+Fixpoint Nwadd (w1 w2 : Nword) {struct w1} : Nword :=
+ match w1, w2 with
+ | n1::w1', n2:: w2' => (n1+n2)%N :: Nwadd w1' w2'
+ | nil, _ => w2
+ | _, nil => w1
+ end.
+
+Definition Nwopp (w:Nword) : Nword := Nwcons 0%N w.
+
+Definition Nwsub w1 w2 := Nwadd w1 (Nwopp w2).
+
+Fixpoint Nwscal (n : N) (w : Nword) {struct w} : Nword :=
+ match w with
+ | m :: w' => (n*m)%N :: Nwscal n w'
+ | nil => nil
+ end.
+
+Fixpoint Nwmul (w1 w2 : Nword) {struct w1} : Nword :=
+ match w1 with
+ | 0%N::w1' => Nwopp (Nwmul w1' w2)
+ | n1::w1' => Nwsub (Nwscal n1 w2) (Nwmul w1' w2)
+ | nil => nil
+ end.
+Fixpoint Nw_is0 (w : Nword) : bool :=
+ match w with
+ | nil => true
+ | 0%N :: w' => Nw_is0 w'
+ | _ => false
+ end.
+
+Fixpoint Nweq_bool (w1 w2 : Nword) {struct w1} : bool :=
+ match w1, w2 with
+ | n1::w1', n2::w2' =>
+ if Neq_bool n1 n2 then Nweq_bool w1' w2' else false
+ | nil, _ => Nw_is0 w2
+ | _, nil => Nw_is0 w1
+ end.
+
+Section NWORDMORPHISM.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable req : R -> R -> Prop.
+ 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).
+ Variable Rsth : Setoid_Theory R req.
+ Add Setoid R req Rsth as R_setoid5.
+ Ltac rrefl := gen_reflexivity Rsth.
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Add Morphism radd : radd_ext5. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul : rmul_ext5. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp : ropp_ext5. exact (Ropp_ext Reqe). Qed.
+
+ Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
+ Add Morphism rsub : rsub_ext7. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Ltac norm := gen_srewrite Rsth Reqe ARth.
+ Ltac add_push := gen_add_push radd Rsth Reqe ARth.
+
+ Fixpoint gen_phiNword (w : Nword) : R :=
+ match w with
+ | nil => 0
+ | n :: nil => gen_phiN rO rI radd rmul n
+ | N0 :: w' => - gen_phiNword w'
+ | n::w' => gen_phiN rO rI radd rmul n - gen_phiNword w'
+ end.
+
+ Lemma gen_phiNword0_ok : forall w, Nw_is0 w = true -> gen_phiNword w == 0.
+Proof.
+induction w; simpl in |- *; intros; auto.
+ reflexivity.
+
+ destruct a.
+ destruct w.
+ reflexivity.
+
+ rewrite IHw in |- *; trivial.
+ apply (ARopp_zero Rsth Reqe ARth).
+
+ discriminate.
+Qed.
+
+ Lemma gen_phiNword_cons : forall w n,
+ gen_phiNword (n::w) == gen_phiN rO rI radd rmul n - gen_phiNword w.
+induction w.
+ destruct n; simpl in |- *; norm.
+
+ intros.
+ destruct n; norm.
+Qed.
+
+ Lemma gen_phiNword_Nwcons : forall w n,
+ gen_phiNword (Nwcons n w) == gen_phiN rO rI radd rmul n - gen_phiNword w.
+destruct w; intros.
+ destruct n; norm.
+
+ unfold Nwcons in |- *.
+ rewrite gen_phiNword_cons in |- *.
+ reflexivity.
+Qed.
+
+ Lemma gen_phiNword_ok : forall w1 w2,
+ Nweq_bool w1 w2 = true -> gen_phiNword w1 == gen_phiNword w2.
+induction w1; intros.
+ simpl in |- *.
+ rewrite (gen_phiNword0_ok _ H) in |- *.
+ reflexivity.
+
+ rewrite gen_phiNword_cons in |- *.
+ destruct w2.
+ simpl in H.
+ destruct a; try discriminate.
+ rewrite (gen_phiNword0_ok _ H) in |- *.
+ norm.
+
+ simpl in H.
+ rewrite gen_phiNword_cons in |- *.
+ case_eq (Neq_bool a n); intros.
+ rewrite H0 in H.
+ rewrite <- (Neq_bool_ok _ _ H0) in |- *.
+ rewrite (IHw1 _ H) in |- *.
+ reflexivity.
+
+ rewrite H0 in H; discriminate H.
+Qed.
+
+
+Lemma Nwadd_ok : forall x y,
+ gen_phiNword (Nwadd x y) == gen_phiNword x + gen_phiNword y.
+induction x; intros.
+ simpl in |- *.
+ norm.
+
+ destruct y.
+ simpl Nwadd; norm.
+
+ simpl Nwadd in |- *.
+ repeat rewrite gen_phiNword_cons in |- *.
+ rewrite (fun sreq => gen_phiN_add Rsth sreq (ARth_SRth ARth)) in |- * by
+ (destruct Reqe; constructor; trivial).
+
+ rewrite IHx in |- *.
+ norm.
+ add_push (- gen_phiNword x); reflexivity.
+Qed.
+
+Lemma Nwopp_ok : forall x, gen_phiNword (Nwopp x) == - gen_phiNword x.
+simpl in |- *.
+unfold Nwopp in |- *; simpl in |- *.
+intros.
+rewrite gen_phiNword_Nwcons in |- *; norm.
+Qed.
+
+Lemma Nwscal_ok : forall n x,
+ gen_phiNword (Nwscal n x) == gen_phiN rO rI radd rmul n * gen_phiNword x.
+induction x; intros.
+ norm.
+
+ simpl Nwscal in |- *.
+ repeat rewrite gen_phiNword_cons in |- *.
+ rewrite (fun sreq => gen_phiN_mult Rsth sreq (ARth_SRth ARth)) in |- *
+ by (destruct Reqe; constructor; trivial).
+
+ rewrite IHx in |- *.
+ norm.
+Qed.
+
+Lemma Nwmul_ok : forall x y,
+ gen_phiNword (Nwmul x y) == gen_phiNword x * gen_phiNword y.
+induction x; intros.
+ norm.
+
+ destruct a.
+ simpl Nwmul in |- *.
+ rewrite Nwopp_ok in |- *.
+ rewrite IHx in |- *.
+ rewrite gen_phiNword_cons in |- *.
+ norm.
+
+ simpl Nwmul in |- *.
+ unfold Nwsub in |- *.
+ rewrite Nwadd_ok in |- *.
+ rewrite Nwscal_ok in |- *.
+ rewrite Nwopp_ok in |- *.
+ rewrite IHx in |- *.
+ rewrite gen_phiNword_cons in |- *.
+ norm.
+Qed.
+
+(* Proof that [.] satisfies morphism specifications *)
+ Lemma gen_phiNword_morph :
+ ring_morph 0 1 radd rmul rsub ropp req
+ NwO NwI Nwadd Nwmul Nwsub Nwopp Nweq_bool gen_phiNword.
+constructor.
+ reflexivity.
+
+ reflexivity.
+
+ exact Nwadd_ok.
+
+ intros.
+ unfold Nwsub in |- *.
+ rewrite Nwadd_ok in |- *.
+ rewrite Nwopp_ok in |- *.
+ norm.
+
+ exact Nwmul_ok.
+
+ exact Nwopp_ok.
+
+ exact gen_phiNword_ok.
+Qed.
+
+End NWORDMORPHISM.
+
+Section GEN_DIV.
+
+ Variables (R : Type) (rO : R) (rI : R) (radd : R -> R -> R)
+ (rmul : R -> R -> R) (rsub : R -> R -> R) (ropp : R -> R)
+ (req : R -> R -> Prop) (C : Type) (cO : C) (cI : C)
+ (cadd : C -> C -> C) (cmul : C -> C -> C) (csub : C -> C -> C)
+ (copp : C -> C) (ceqb : C -> C -> bool) (phi : C -> R).
+ Variable Rsth : Setoid_Theory R req.
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req.
+ Variable morph : ring_morph rO rI radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi.
+
+ (* 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.
+ Ltac rsimpl := gen_srewrite Rsth Reqe ARth.
+
+ Definition triv_div x y :=
+ if ceqb x y then (cI, cO)
+ else (cO, x).
+
+ Ltac Esimpl :=repeat (progress (
+ match goal with
+ | |- context [phi cO] => rewrite (morph0 morph)
+ | |- context [phi cI] => rewrite (morph1 morph)
+ | |- context [phi (cadd ?x ?y)] => rewrite ((morph_add morph) x y)
+ | |- context [phi (cmul ?x ?y)] => rewrite ((morph_mul morph) x y)
+ | |- context [phi (csub ?x ?y)] => rewrite ((morph_sub morph) x y)
+ | |- context [phi (copp ?x)] => rewrite ((morph_opp morph) x)
+ end)).
+
+ Lemma triv_div_th : Ring_theory.div_theory req cadd cmul phi triv_div.
+ Proof.
+ constructor.
+ intros a b;unfold triv_div.
+ assert (X:= morph.(morph_eq) a b);destruct (ceqb a b).
+ Esimpl.
+ rewrite X; trivial.
+ rsimpl.
+ Esimpl; rsimpl.
+Qed.
+
+ Variable zphi : Z -> R.
+
+ Lemma Ztriv_div_th : div_theory req Zplus Zmult zphi ZOdiv_eucl.
+ Proof.
+ constructor.
+ intros; generalize (ZOdiv_eucl_correct a b); case ZOdiv_eucl; intros; subst.
+ rewrite Zmult_comm; rsimpl.
+ Qed.
+
+ Variable nphi : N -> R.
+
+ Lemma Ntriv_div_th : div_theory req Nplus Nmult nphi Ndiv_eucl.
+ constructor.
+ intros; generalize (Ndiv_eucl_correct a b); case Ndiv_eucl; intros; subst.
+ rewrite Nmult_comm; rsimpl.
+ Qed.
+
+End GEN_DIV.
+
+ (* syntaxification of constants in an abstract ring:
+ the inverse of gen_phiPOS *)
+ Ltac inv_gen_phi_pos rI add mul t :=
+ let rec inv_cst t :=
+ match t with
+ rI => constr:1%positive
+ | (add rI rI) => constr:2%positive
+ | (add rI (add rI rI)) => constr:3%positive
+ | (mul (add rI rI) ?p) => (* 2p *)
+ match inv_cst p with
+ NotConstant => constr:NotConstant
+ | 1%positive => constr:NotConstant (* 2*1 is not convertible to 2 *)
+ | ?p => constr:(xO p)
+ end
+ | (add rI (mul (add rI rI) ?p)) => (* 1+2p *)
+ match inv_cst p with
+ NotConstant => constr:NotConstant
+ | 1%positive => constr:NotConstant
+ | ?p => constr:(xI p)
+ end
+ | _ => constr:NotConstant
+ end in
+ inv_cst t.
+
+(* The (partial) inverse of gen_phiNword *)
+ Ltac inv_gen_phiNword rO rI add mul opp t :=
+ match t with
+ rO => constr:NwO
+ | _ =>
+ match inv_gen_phi_pos rI add mul t with
+ NotConstant => constr:NotConstant
+ | ?p => constr:(Npos p::nil)
+ end
+ end.
+
+
+(* The inverse of gen_phiN *)
+ Ltac inv_gen_phiN rO rI add mul t :=
+ match t with
+ rO => constr:0%N
+ | _ =>
+ match inv_gen_phi_pos rI add mul t with
+ NotConstant => constr:NotConstant
+ | ?p => constr:(Npos p)
+ end
+ end.
+
+(* The inverse of gen_phiZ *)
+ Ltac inv_gen_phiZ rO rI add mul opp t :=
+ match t with
+ rO => constr:0%Z
+ | (opp ?p) =>
+ match inv_gen_phi_pos rI add mul p with
+ NotConstant => constr:NotConstant
+ | ?p => constr:(Zneg p)
+ end
+ | _ =>
+ match inv_gen_phi_pos rI add mul t with
+ NotConstant => constr:NotConstant
+ | ?p => constr:(Zpos p)
+ end
+ end.
+
+(* A simple tactic recognizing only 0 and 1. The inv_gen_phiX above
+ are only optimisations that directly returns the reifid constant
+ instead of resorting to the constant propagation of the simplification
+ algorithm. *)
+Ltac inv_gen_phi rO rI cO cI t :=
+ match t with
+ | rO => cO
+ | rI => cI
+ end.
+
+(* A simple tactic recognizing no constant *)
+ Ltac inv_morph_nothing t := constr:NotConstant.
+
+Ltac coerce_to_almost_ring set ext rspec :=
+ match type of rspec with
+ | ring_theory _ _ _ _ _ _ _ => constr:(Rth_ARth set ext rspec)
+ | semi_ring_theory _ _ _ _ _ => constr:(SRth_ARth set rspec)
+ | almost_ring_theory _ _ _ _ _ _ _ => rspec
+ | _ => fail 1 "not a valid ring theory"
+ end.
+
+Ltac coerce_to_ring_ext ext :=
+ match type of ext with
+ | ring_eq_ext _ _ _ _ => ext
+ | sring_eq_ext _ _ _ => constr:(SReqe_Reqe ext)
+ | _ => fail 1 "not a valid ring_eq_ext theory"
+ end.
+
+Ltac abstract_ring_morphism set ext rspec :=
+ match type of rspec with
+ | ring_theory _ _ _ _ _ _ _ => constr:(gen_phiZ_morph set ext rspec)
+ | semi_ring_theory _ _ _ _ _ => constr:(gen_phiN_morph set ext rspec)
+ | almost_ring_theory _ _ _ _ _ _ _ =>
+ constr:(gen_phiNword_morph set ext rspec)
+ | _ => fail 1 "bad ring structure"
+ end.
+
+Record hypo : Type := mkhypo {
+ hypo_type : Type;
+ hypo_proof : hypo_type
+ }.
+
+Ltac gen_ring_pow set arth pspec :=
+ match pspec with
+ | None =>
+ match type of arth with
+ | @almost_ring_theory ?R ?rO ?rI ?radd ?rmul ?rsub ?ropp ?req =>
+ constr:(mkhypo (@pow_N_th R rI rmul req set))
+ | _ => fail 1 "gen_ring_pow"
+ end
+ | Some ?t => constr:(t)
+ end.
+
+Ltac gen_ring_sign morph sspec :=
+ match sspec with
+ | None =>
+ match type of morph with
+ | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
+ Z ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceqb ?phi =>
+ constr:(@mkhypo (sign_theory copp ceqb get_signZ) get_signZ_th)
+ | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
+ ?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceqb ?phi =>
+ constr:(mkhypo (@get_sign_None_th C copp ceqb))
+ | _ => fail 2 "ring anomaly : default_sign_spec"
+ end
+ | Some ?t => constr:(t)
+ end.
+
+Ltac default_div_spec set reqe arth morph :=
+ match type of morph with
+ | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
+ Z ?c0 ?c1 Zplus Zmult ?csub ?copp ?ceq_b ?phi =>
+ constr:(mkhypo (Ztriv_div_th set phi))
+ | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
+ N ?c0 ?c1 Nplus Nmult ?csub ?copp ?ceq_b ?phi =>
+ constr:(mkhypo (Ntriv_div_th set phi))
+ | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
+ ?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceq_b ?phi =>
+ constr:(mkhypo (triv_div_th set reqe arth morph))
+ | _ => fail 1 "ring anomaly : default_sign_spec"
+ end.
+
+Ltac gen_ring_div set reqe arth morph dspec :=
+ match dspec with
+ | None => default_div_spec set reqe arth morph
+ | Some ?t => constr:(t)
+ end.
+
+Ltac ring_elements set ext rspec pspec sspec dspec rk :=
+ let arth := coerce_to_almost_ring set ext rspec in
+ let ext_r := coerce_to_ring_ext ext in
+ let morph :=
+ match rk with
+ | Abstract => abstract_ring_morphism set ext rspec
+ | @Computational ?reqb_ok =>
+ match type of arth with
+ | almost_ring_theory ?rO ?rI ?add ?mul ?sub ?opp _ =>
+ constr:(IDmorph rO rI add mul sub opp set _ reqb_ok)
+ | _ => fail 2 "ring anomaly"
+ end
+ | @Morphism ?m =>
+ match type of m with
+ | ring_morph _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ => m
+ | @semi_morph _ _ _ _ _ _ _ _ _ _ _ _ _ =>
+ constr:(SRmorph_Rmorph set m)
+ | _ => fail 2 "ring anomaly"
+ end
+ | _ => fail 1 "ill-formed ring kind"
+ end in
+ let p_spec := gen_ring_pow set arth pspec in
+ let s_spec := gen_ring_sign morph sspec in
+ let d_spec := gen_ring_div set ext_r arth morph dspec in
+ fun f => f arth ext_r morph p_spec s_spec d_spec.
+
+(* Given a ring structure and the kind of morphism,
+ returns 2 lemmas (one for ring, and one for ring_simplify). *)
+
+ Ltac ring_lemmas set ext rspec pspec sspec dspec rk :=
+ let gen_lemma2 :=
+ match pspec with
+ | None => constr:(ring_rw_correct)
+ | Some _ => constr:(ring_rw_pow_correct)
+ end in
+ ring_elements set ext rspec pspec sspec dspec rk
+ ltac:(fun arth ext_r morph p_spec s_spec d_spec =>
+ match type of morph with
+ | @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
+ ?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceq_b ?phi =>
+ let gen_lemma2_0 :=
+ constr:(gen_lemma2 R r0 rI radd rmul rsub ropp req set ext_r arth
+ C c0 c1 cadd cmul csub copp ceq_b phi morph) in
+ match p_spec with
+ | @mkhypo (power_theory _ _ _ ?Cp_phi ?rpow) ?pp_spec =>
+ let gen_lemma2_1 := constr:(gen_lemma2_0 _ Cp_phi rpow pp_spec) in
+ match d_spec with
+ | @mkhypo (div_theory _ _ _ _ ?cdiv) ?dd_spec =>
+ let gen_lemma2_2 := constr:(gen_lemma2_1 cdiv dd_spec) in
+ match s_spec with
+ | @mkhypo (sign_theory _ _ ?get_sign) ?ss_spec =>
+ let lemma2 := constr:(gen_lemma2_2 get_sign ss_spec) in
+ let lemma1 :=
+ constr:(ring_correct set ext_r arth morph pp_spec dd_spec) in
+ fun f => f arth ext_r morph lemma1 lemma2
+ | _ => fail 4 "ring: bad sign specification"
+ end
+ | _ => fail 3 "ring: bad coefficiant division specification"
+ end
+ | _ => fail 2 "ring: bad power specification"
+ end
+ | _ => fail 1 "ring internal error: ring_lemmas, please report"
+ end).
+
+(* Tactic for constant *)
+Ltac isnatcst t :=
+ match t with
+ O => constr:true
+ | S ?p => isnatcst p
+ | _ => constr:false
+ end.
+
+Ltac isPcst t :=
+ match t with
+ | xI ?p => isPcst p
+ | xO ?p => isPcst p
+ | xH => constr:true
+ (* nat -> positive *)
+ | P_of_succ_nat ?n => isnatcst n
+ | _ => constr:false
+ end.
+
+Ltac isNcst t :=
+ match t with
+ N0 => constr:true
+ | Npos ?p => isPcst p
+ | _ => constr:false
+ end.
+
+Ltac isZcst t :=
+ match t with
+ Z0 => constr:true
+ | Zpos ?p => isPcst p
+ | Zneg ?p => isPcst p
+ (* injection nat -> Z *)
+ | Z_of_nat ?n => isnatcst n
+ (* injection N -> Z *)
+ | Z_of_N ?n => isNcst n
+ (* *)
+ | _ => constr:false
+ end.
+
+
+
+
+
diff --git a/plugins/setoid_ring/NArithRing.v b/plugins/setoid_ring/NArithRing.v
new file mode 100644
index 00000000..0ba519fd
--- /dev/null
+++ b/plugins/setoid_ring/NArithRing.v
@@ -0,0 +1,21 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Export Ring.
+Require Import BinPos BinNat.
+Import InitialRing.
+
+Set Implicit Arguments.
+
+Ltac Ncst t :=
+ match isNcst t with
+ true => t
+ | _ => constr:NotConstant
+ end.
+
+Add Ring Nr : Nth (decidable Neq_bool_ok, constants [Ncst]).
diff --git a/plugins/setoid_ring/RealField.v b/plugins/setoid_ring/RealField.v
new file mode 100644
index 00000000..56473adb
--- /dev/null
+++ b/plugins/setoid_ring/RealField.v
@@ -0,0 +1,134 @@
+Require Import Nnat.
+Require Import ArithRing.
+Require Export Ring Field.
+Require Import Rdefinitions.
+Require Import Rpow_def.
+Require Import Raxioms.
+
+Open Local Scope R_scope.
+
+Lemma RTheory : ring_theory 0 1 Rplus Rmult Rminus Ropp (eq (A:=R)).
+Proof.
+constructor.
+ intro; apply Rplus_0_l.
+ exact Rplus_comm.
+ symmetry in |- *; apply Rplus_assoc.
+ intro; apply Rmult_1_l.
+ exact Rmult_comm.
+ symmetry in |- *; apply Rmult_assoc.
+ intros m n p.
+ rewrite Rmult_comm in |- *.
+ rewrite (Rmult_comm n p) in |- *.
+ rewrite (Rmult_comm m p) in |- *.
+ apply Rmult_plus_distr_l.
+ reflexivity.
+ exact Rplus_opp_r.
+Qed.
+
+Lemma Rfield : field_theory 0 1 Rplus Rmult Rminus Ropp Rdiv Rinv (eq(A:=R)).
+Proof.
+constructor.
+ exact RTheory.
+ exact R1_neq_R0.
+ reflexivity.
+ exact Rinv_l.
+Qed.
+
+Lemma Rlt_n_Sn : forall x, x < x + 1.
+Proof.
+intro.
+elim archimed with x; intros.
+destruct H0.
+ apply Rlt_trans with (IZR (up x)); trivial.
+ replace (IZR (up x)) with (x + (IZR (up x) - x))%R.
+ apply Rplus_lt_compat_l; trivial.
+ unfold Rminus in |- *.
+ rewrite (Rplus_comm (IZR (up x)) (- x)) in |- *.
+ rewrite <- Rplus_assoc in |- *.
+ rewrite Rplus_opp_r in |- *.
+ apply Rplus_0_l.
+ elim H0.
+ unfold Rminus in |- *.
+ rewrite (Rplus_comm (IZR (up x)) (- x)) in |- *.
+ rewrite <- Rplus_assoc in |- *.
+ rewrite Rplus_opp_r in |- *.
+ rewrite Rplus_0_l in |- *; trivial.
+Qed.
+
+Notation Rset := (Eqsth R).
+Notation Rext := (Eq_ext Rplus Rmult Ropp).
+
+Lemma Rlt_0_2 : 0 < 2.
+apply Rlt_trans with (0 + 1).
+ apply Rlt_n_Sn.
+ rewrite Rplus_comm in |- *.
+ apply Rplus_lt_compat_l.
+ replace 1 with (0 + 1).
+ apply Rlt_n_Sn.
+ apply Rplus_0_l.
+Qed.
+
+Lemma Rgen_phiPOS : forall x, InitialRing.gen_phiPOS1 1 Rplus Rmult x > 0.
+unfold Rgt in |- *.
+induction x; simpl in |- *; intros.
+ apply Rlt_trans with (1 + 0).
+ rewrite Rplus_comm in |- *.
+ apply Rlt_n_Sn.
+ apply Rplus_lt_compat_l.
+ rewrite <- (Rmul_0_l Rset Rext RTheory 2) in |- *.
+ rewrite Rmult_comm in |- *.
+ apply Rmult_lt_compat_l.
+ apply Rlt_0_2.
+ trivial.
+ rewrite <- (Rmul_0_l Rset Rext RTheory 2) in |- *.
+ rewrite Rmult_comm in |- *.
+ apply Rmult_lt_compat_l.
+ apply Rlt_0_2.
+ trivial.
+ replace 1 with (0 + 1).
+ apply Rlt_n_Sn.
+ apply Rplus_0_l.
+Qed.
+
+
+Lemma Rgen_phiPOS_not_0 :
+ forall x, InitialRing.gen_phiPOS1 1 Rplus Rmult x <> 0.
+red in |- *; intros.
+specialize (Rgen_phiPOS x).
+rewrite H in |- *; intro.
+apply (Rlt_asym 0 0); trivial.
+Qed.
+
+Lemma Zeq_bool_complete : forall x y,
+ InitialRing.gen_phiZ 0%R 1%R Rplus Rmult Ropp x =
+ InitialRing.gen_phiZ 0%R 1%R Rplus Rmult Ropp y ->
+ Zeq_bool x y = true.
+Proof gen_phiZ_complete Rset Rext Rfield Rgen_phiPOS_not_0.
+
+Lemma Rdef_pow_add : forall (x:R) (n m:nat), pow x (n + m) = pow x n * pow x m.
+Proof.
+ intros x n; elim n; simpl in |- *; auto with real.
+ intros n0 H' m; rewrite H'; auto with real.
+Qed.
+
+Lemma R_power_theory : power_theory 1%R Rmult (eq (A:=R)) nat_of_N pow.
+Proof.
+ constructor. destruct n. reflexivity.
+ simpl. induction p;simpl.
+ rewrite ZL6. rewrite Rdef_pow_add;rewrite IHp. reflexivity.
+ unfold nat_of_P;simpl;rewrite ZL6;rewrite Rdef_pow_add;rewrite IHp;trivial.
+ rewrite Rmult_comm;apply Rmult_1_l.
+Qed.
+
+Ltac Rpow_tac t :=
+ match isnatcst t with
+ | false => constr:(InitialRing.NotConstant)
+ | _ => constr:(N_of_nat t)
+ end.
+
+Add Field RField : Rfield
+ (completeness Zeq_bool_complete, power_tac R_power_theory [Rpow_tac]).
+
+
+
+
diff --git a/plugins/setoid_ring/Ring.v b/plugins/setoid_ring/Ring.v
new file mode 100644
index 00000000..d01b1625
--- /dev/null
+++ b/plugins/setoid_ring/Ring.v
@@ -0,0 +1,44 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Bool.
+Require Export Ring_theory.
+Require Export Ring_base.
+Require Export InitialRing.
+Require Export Ring_tac.
+
+Lemma BoolTheory :
+ ring_theory false true xorb andb xorb (fun b:bool => b) (eq(A:=bool)).
+split; simpl in |- *.
+destruct x; reflexivity.
+destruct x; destruct y; reflexivity.
+destruct x; destruct y; destruct z; reflexivity.
+reflexivity.
+destruct x; destruct y; reflexivity.
+destruct x; destruct y; reflexivity.
+destruct x; destruct y; destruct z; reflexivity.
+reflexivity.
+destruct x; reflexivity.
+Qed.
+
+Definition bool_eq (b1 b2:bool) :=
+ if b1 then b2 else negb b2.
+
+Lemma bool_eq_ok : forall b1 b2, bool_eq b1 b2 = true -> b1 = b2.
+destruct b1; destruct b2; auto.
+Qed.
+
+Ltac bool_cst t :=
+ let t := eval hnf in t in
+ match t with
+ true => constr:true
+ | false => constr:false
+ | _ => constr:NotConstant
+ end.
+
+Add Ring bool_ring : BoolTheory (decidable bool_eq_ok, constants [bool_cst]).
diff --git a/plugins/setoid_ring/Ring_base.v b/plugins/setoid_ring/Ring_base.v
new file mode 100644
index 00000000..fd9dd8d0
--- /dev/null
+++ b/plugins/setoid_ring/Ring_base.v
@@ -0,0 +1,17 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This module gathers the necessary base to build an instance of the
+ ring tactic. Abstract rings need more theory, depending on
+ ZArith_base. *)
+
+Require Import Quote.
+Declare ML Module "newring_plugin".
+Require Export Ring_theory.
+Require Export Ring_tac.
+Require Import InitialRing.
diff --git a/plugins/setoid_ring/Ring_equiv.v b/plugins/setoid_ring/Ring_equiv.v
new file mode 100644
index 00000000..945f6c68
--- /dev/null
+++ b/plugins/setoid_ring/Ring_equiv.v
@@ -0,0 +1,74 @@
+Require Import Setoid_ring_theory.
+Require Import LegacyRing_theory.
+Require Import Ring_theory.
+
+Set Implicit Arguments.
+
+Section Old2New.
+
+Variable A : Type.
+
+Variable Aplus : A -> A -> A.
+Variable Amult : A -> A -> A.
+Variable Aone : A.
+Variable Azero : A.
+Variable Aopp : A -> A.
+Variable Aeq : A -> A -> bool.
+Variable R : Ring_Theory Aplus Amult Aone Azero Aopp Aeq.
+
+Let Aminus := fun x y => Aplus x (Aopp y).
+
+Lemma ring_equiv1 :
+ ring_theory Azero Aone Aplus Amult Aminus Aopp (eq (A:=A)).
+Proof.
+destruct R.
+split; eauto.
+Qed.
+
+End Old2New.
+
+Section New2OldRing.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable Rth : ring_theory rO rI radd rmul rsub ropp (eq (A:=R)).
+
+ Variable reqb : R -> R -> bool.
+ Variable reqb_ok : forall x y, reqb x y = true -> x = y.
+
+ Lemma ring_equiv2 :
+ Ring_Theory radd rmul rI rO ropp reqb.
+Proof.
+elim Rth; intros; constructor; eauto.
+intros.
+apply reqb_ok.
+destruct (reqb x y); trivial; intros.
+elim H.
+Qed.
+
+ Definition default_eqb : R -> R -> bool := fun x y => false.
+ Lemma default_eqb_ok : forall x y, default_eqb x y = true -> x = y.
+Proof.
+discriminate 1.
+Qed.
+
+End New2OldRing.
+
+Section New2OldSemiRing.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul: R->R->R).
+ Variable SRth : semi_ring_theory rO rI radd rmul (eq (A:=R)).
+
+ Variable reqb : R -> R -> bool.
+ Variable reqb_ok : forall x y, reqb x y = true -> x = y.
+
+ Lemma sring_equiv2 :
+ Semi_Ring_Theory radd rmul rI rO reqb.
+Proof.
+elim SRth; intros; constructor; eauto.
+intros.
+apply reqb_ok.
+destruct (reqb x y); trivial; intros.
+elim H.
+Qed.
+
+End New2OldSemiRing.
diff --git a/plugins/setoid_ring/Ring_polynom.v b/plugins/setoid_ring/Ring_polynom.v
new file mode 100644
index 00000000..faa83ded
--- /dev/null
+++ b/plugins/setoid_ring/Ring_polynom.v
@@ -0,0 +1,1781 @@
+(************************************************************************)
+(* V * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Set Implicit Arguments.
+Require Import Setoid.
+Require Import BinList.
+Require Import BinPos.
+Require Import BinNat.
+Require Import BinInt.
+Require Export Ring_theory.
+
+Open Local Scope positive_scope.
+Import RingSyntax.
+
+Section MakeRingPol.
+
+ (* Ring elements *)
+ Variable R:Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R->R).
+ Variable req : R -> R -> Prop.
+
+ (* Ring properties *)
+ Variable Rsth : Setoid_Theory R req.
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req.
+
+ (* Coefficients *)
+ Variable C: Type.
+ Variable (cO cI: C) (cadd cmul csub : C->C->C) (copp : C->C).
+ Variable ceqb : C->C->bool.
+ Variable phi : C -> R.
+ Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req
+ cO cI cadd cmul csub copp ceqb phi.
+
+ (* Power coefficients *)
+ Variable Cpow : Set.
+ Variable Cp_phi : N -> Cpow.
+ Variable rpow : R -> Cpow -> R.
+ Variable pow_th : power_theory rI rmul req Cp_phi rpow.
+
+ (* division is ok *)
+ Variable cdiv: C -> C -> C * C.
+ Variable div_th: div_theory req cadd cmul phi cdiv.
+
+
+ (* 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).
+
+ (* 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).
+
+ (* 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.
+ 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.
+
+ (* Definition of multivariable polynomials with coefficients in C :
+ Type [Pol] represents [X1 ... Xn].
+ The representation is Horner's where a [n] variable polynomial
+ (C[X1..Xn]) is seen as a polynomial on [X1] which coefficients
+ are polynomials with [n-1] variables (C[X2..Xn]).
+ There are several optimisations to make the repr compacter:
+ - [Pc c] is the constant polynomial of value c
+ == c*X1^0*..*Xn^0
+ - [Pinj j Q] is a polynomial constant w.r.t the [j] first variables.
+ variable indices are shifted of j in Q.
+ == X1^0 *..* Xj^0 * Q{X1 <- Xj+1;..; Xn-j <- Xn}
+ - [PX P i Q] is an optimised Horner form of P*X^i + Q
+ with P not the null polynomial
+ == P * X1^i + Q{X1 <- X2; ..; Xn-1 <- Xn}
+
+ In addition:
+ - polynomials of the form (PX (PX P i (Pc 0)) j Q) are forbidden
+ since they can be represented by the simpler form (PX P (i+j) Q)
+ - (Pinj i (Pinj j P)) is (Pinj (i+j) P)
+ - (Pinj i (Pc c)) is (Pc c)
+ *)
+
+ Inductive Pol : Type :=
+ | Pc : C -> Pol
+ | Pinj : positive -> Pol -> Pol
+ | PX : Pol -> positive -> Pol -> Pol.
+
+ Definition P0 := Pc cO.
+ Definition P1 := Pc cI.
+
+ Fixpoint Peq (P P' : Pol) {struct P'} : bool :=
+ match P, P' with
+ | Pc c, Pc c' => c ?=! c'
+ | Pinj j Q, Pinj j' Q' =>
+ match Pcompare j j' Eq with
+ | Eq => Peq Q Q'
+ | _ => false
+ end
+ | PX P i Q, PX P' i' Q' =>
+ match Pcompare i i' Eq with
+ | Eq => if Peq P P' then Peq Q Q' else false
+ | _ => false
+ end
+ | _, _ => false
+ end.
+
+ Notation " P ?== P' " := (Peq P P').
+
+ Definition mkPinj j P :=
+ match P with
+ | Pc _ => P
+ | Pinj j' Q => Pinj ((j + j'):positive) Q
+ | _ => Pinj j P
+ end.
+
+ Definition mkPinj_pred j P:=
+ match j with
+ | xH => P
+ | xO j => Pinj (Pdouble_minus_one j) P
+ | xI j => Pinj (xO j) P
+ end.
+
+ Definition mkPX P i Q :=
+ match P with
+ | Pc c => if c ?=! cO then mkPinj xH Q else PX P i Q
+ | Pinj _ _ => PX P i Q
+ | PX P' i' Q' => if Q' ?== P0 then PX P' (i' + i) Q else PX P i Q
+ end.
+
+ Definition mkXi i := PX P1 i P0.
+
+ Definition mkX := mkXi 1.
+
+ (** Opposite of addition *)
+
+ Fixpoint Popp (P:Pol) : Pol :=
+ match P with
+ | Pc c => Pc (-! c)
+ | Pinj j Q => Pinj j (Popp Q)
+ | PX P i Q => PX (Popp P) i (Popp Q)
+ end.
+
+ Notation "-- P" := (Popp P).
+
+ (** Addition et subtraction *)
+
+ Fixpoint PaddC (P:Pol) (c:C) {struct P} : 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 :=
+ match P with
+ | Pc c1 => Pc (c1 -! c)
+ | Pinj j Q => Pinj j (PsubC Q c)
+ | PX P i Q => PX P i (PsubC Q c)
+ end.
+
+ Section PopI.
+
+ Variable Pop : Pol -> Pol -> Pol.
+ Variable Q : Pol.
+
+ Fixpoint PaddI (j:positive) (P:Pol){struct P} : Pol :=
+ match P with
+ | Pc c => mkPinj j (PaddC Q c)
+ | Pinj j' Q' =>
+ match ZPminus 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')
+ end
+ | 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')
+ | xI j => PX P i (PaddI (xO j) Q')
+ end
+ end.
+
+ Fixpoint PsubI (j:positive) (P:Pol){struct P} : Pol :=
+ match P with
+ | Pc c => mkPinj j (PaddC (--Q) c)
+ | Pinj j' Q' =>
+ match ZPminus 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')
+ end
+ | 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')
+ | xI j => PX P i (PsubI (xO j) Q')
+ end
+ end.
+
+ Variable P' : Pol.
+
+ Fixpoint PaddX (i':positive) (P:Pol) {struct P} : 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')
+ | xI j => PX P' i' (Pinj (xO j) Q')
+ end
+ | PX P i Q' =>
+ match ZPminus 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 :=
+ 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')
+ | xI j => PX (--P') i' (Pinj (xO j) Q')
+ end
+ | PX P i Q' =>
+ match ZPminus 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'
+ end
+ end.
+
+
+ End PopI.
+
+ Fixpoint Padd (P P': Pol) {struct P'} : Pol :=
+ match P' with
+ | Pc c' => PaddC P c'
+ | Pinj j' Q' => PaddI Padd Q' j' P
+ | PX P' i' Q' =>
+ match P with
+ | Pc c => PX P' i' (PaddC Q' c)
+ | 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')
+ | xI j => PX P' i' (Padd (Pinj (xO j) Q) Q')
+ end
+ | PX P i Q =>
+ match ZPminus 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').
+
+ Fixpoint Psub (P P': Pol) {struct P'} : Pol :=
+ match P' with
+ | Pc c' => PsubC P c'
+ | Pinj j' Q' => PsubI Psub Q' j' P
+ | PX P' i' Q' =>
+ match P with
+ | Pc c => PX (--P') i' (*(--(PsubC Q' c))*) (PaddC (--Q') c)
+ | 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')
+ | xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q')
+ end
+ | PX P i Q =>
+ match ZPminus 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').
+
+ (** Multiplication *)
+
+ Fixpoint PmulC_aux (P:Pol) (c:C) {struct P} : Pol :=
+ match P with
+ | Pc c' => Pc (c' *! c)
+ | Pinj j Q => mkPinj j (PmulC_aux Q c)
+ | PX P i Q => mkPX (PmulC_aux P c) i (PmulC_aux Q c)
+ end.
+
+ Definition PmulC P c :=
+ if c ?=! cO then P0 else
+ if c ?=! cI then P else PmulC_aux P c.
+
+ Section PmulI.
+ Variable Pmul : Pol -> Pol -> Pol.
+ Variable Q : Pol.
+ Fixpoint PmulI (j:positive) (P:Pol) {struct P} : Pol :=
+ match P with
+ | Pc c => mkPinj j (PmulC Q c)
+ | Pinj j' Q' =>
+ match ZPminus 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')
+ end
+ | 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')
+ | 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
+ | Pc c => PmulC P c
+ | Pinj j' Q' => PmulI Pmul Q' j' P
+ | PX P' i' Q' =>
+ match P with
+ | Pc c => PmulC P'' c
+ | Pinj j Q =>
+ let QQ' :=
+ match j with
+ | xH => Pmul Q Q'
+ | xO j => Pmul (Pinj (Pdouble_minus_one j) Q) Q'
+ | xI j => Pmul (Pinj (xO j) Q) Q'
+ end in
+ mkPX (Pmul P P') i' QQ'
+ | PX P i Q=>
+ let QQ' := Pmul Q Q' in
+ let PQ' := PmulI Pmul Q' xH P in
+ let QP' := Pmul (mkPinj xH Q) P' in
+ let PP' := Pmul P P' in
+ (mkPX (mkPX PP' i P0 ++ QP') i' P0) ++ mkPX PQ' i QQ'
+ 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').
+
+ Fixpoint Psquare (P:Pol) : Pol :=
+ match P with
+ | Pc c => Pc (c *! c)
+ | Pinj j Q => Pinj j (Psquare Q)
+ | PX P i Q =>
+ let twoPQ := Pmul P (mkPinj xH (PmulC Q (cI +! cI))) in
+ let Q2 := Psquare Q in
+ let P2 := Psquare P in
+ mkPX (mkPX P2 i P0 ++ twoPQ) i Q2
+ end.
+
+ (** Monomial **)
+
+ Inductive Mon: Set :=
+ 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.
+
+ Definition mkVmon i M :=
+ match M with
+ | mon0 => vmon i mon0
+ | zmon j m => vmon i (zmon_pred j m)
+ | vmon i' m => vmon (i+i') m
+ end.
+
+ Fixpoint CFactor (P: Pol) (c: C) {struct P}: Pol * Pol :=
+ match P with
+ | Pc c1 => let (q,r) := cdiv c1 c in (Pc r, Pc q)
+ | Pinj j1 P1 =>
+ let (R,S) := CFactor P1 c in
+ (mkPinj j1 R, mkPinj j1 S)
+ | 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)
+ end.
+
+ 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
+ | Pc _, _ => (P, Pc cO)
+ | Pinj j1 P1, zmon j2 M1 =>
+ match (j1 ?= j2) Eq with
+ Eq => let (R,S) := MFactor P1 c M1 in
+ (mkPinj j1 R, mkPinj j1 S)
+ | Lt => let (R,S) := MFactor P1 c (zmon (j2 - j1) M1) in
+ (mkPinj j1 R, mkPinj j1 S)
+ | Gt => (P, Pc cO)
+ end
+ | Pinj _ _, vmon _ _ => (P, Pc cO)
+ | PX P1 i Q1, zmon j M1 =>
+ let M2 := zmon_pred j M1 in
+ let (R1, S1) := MFactor P1 c M in
+ let (R2, S2) := MFactor Q1 c M2 in
+ (mkPX R1 i R2, mkPX S1 i S2)
+ | PX P1 i Q1, vmon j M1 =>
+ match (i ?= j) Eq with
+ Eq => let (R1,S1) := MFactor P1 c (mkZmon xH M1) in
+ (mkPX R1 i Q1, S1)
+ | Lt => let (R1,S1) := MFactor P1 c (vmon (j - i) M1) in
+ (mkPX R1 i Q1, S1)
+ | Gt => let (R1,S1) := MFactor P1 c (mkZmon xH M1) in
+ (mkPX R1 i Q1, mkPX S1 (i-j) (Pc cO))
+ end
+ end.
+
+ Definition POneSubst (P1: Pol) (cM1: C * Mon) (P2: Pol): option Pol :=
+ let (c,M1) := cM1 in
+ let (Q1,R1) := MFactor P1 c M1 in
+ match R1 with
+ (Pc c) => if c ?=! cO then None
+ else Some (Padd Q1 (Pmul P2 R1))
+ | _ => Some (Padd Q1 (Pmul P2 R1))
+ end.
+
+ Fixpoint PNSubst1 (P1: Pol) (cM1: C * Mon) (P2: Pol) (n: nat) {struct n}: Pol :=
+ match POneSubst P1 cM1 P2 with
+ Some P3 => match n with S n1 => PNSubst1 P3 cM1 P2 n1 | _ => P3 end
+ | _ => P1
+ end.
+
+ Definition PNSubst (P1: Pol) (cM1: C * Mon) (P2: Pol) (n: nat): option Pol :=
+ match POneSubst P1 cM1 P2 with
+ Some P3 => match n with S n1 => Some (PNSubst1 P3 cM1 P2 n1) | _ => None end
+ | _ => None
+ end.
+
+ Fixpoint PSubstL1 (P1: Pol) (LM1: list ((C * Mon) * Pol)) (n: nat) {struct LM1}:
+ 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 :=
+ match LM1 with
+ cons (M1,P2) LM2 =>
+ match PNSubst P1 M1 P2 n with
+ Some P3 => Some (PSubstL1 P3 LM2 n)
+ | None => PSubstL P1 LM2 n
+ end
+ | _ => None
+ end.
+
+ Fixpoint PNSubstL (P1: Pol) (LM1: list ((C * Mon) * Pol)) (m n: nat) {struct m}: Pol :=
+ match PSubstL P1 LM1 n with
+ Some P3 => match m with S m1 => PNSubstL P3 LM1 m1 n | _ => P3 end
+ | _ => P1
+ end.
+
+ (** Evaluation of a polynomial towards R *)
+
+ Fixpoint Pphi(l:list R) (P:Pol) {struct P} : 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)
+ end.
+
+ Reserved Notation "P @ l " (at level 10, no associativity).
+ Notation "P @ l " := (Pphi l P).
+ (** 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
+ end.
+ 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.
+ Qed.
+
+ Lemma Peq_ok : forall P P',
+ (P ?== P') = true -> forall l, P@l == P'@ l.
+ Proof.
+ induction P;destruct P';simpl;intros;try discriminate;trivial.
+ apply (morph_eq CRmorph);trivial.
+ assert (H1 := Pcompare_Eq_eq p p0); destruct ((p ?= p0)%positive Eq);
+ try discriminate H.
+ rewrite (IHP P' H); rewrite H1;trivial;rrefl.
+ assert (H1 := Pcompare_Eq_eq p p0); destruct ((p ?= p0)%positive Eq);
+ 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.
+ Qed.
+
+ Lemma Pphi0 : forall l, P0@l == 0.
+ Proof.
+ intros;simpl;apply (morph0 CRmorph).
+ Qed.
+
+ Lemma Pphi1 : forall l, P1@l == 1.
+ Proof.
+ intros;simpl;apply (morph1 CRmorph).
+ Qed.
+
+ Lemma mkPinj_ok : forall 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.
+ Qed.
+
+ Let pow_pos_Pplus :=
+ pow_pos_Pplus rmul Rsth Reqe.(Rmul_ext) ARth.(ARmul_comm) ARth.(ARmul_assoc).
+
+ 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).
+ 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.
+ 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].
+ Proof.
+ induction P;simpl;intros;Esimpl;trivial.
+ rewrite IHP2;rsimpl.
+ Qed.
+
+ Lemma PsubC_ok : forall c P l, (PsubC P c)@l == P@l - [c].
+ Proof.
+ induction P;simpl;intros.
+ Esimpl.
+ rewrite IHP;rsimpl.
+ rewrite IHP2;rsimpl.
+ Qed.
+
+ Lemma PmulC_aux_ok : forall c P l, (PmulC_aux P c)@l == P@l * [c].
+ Proof.
+ induction P;simpl;intros;Esimpl;trivial.
+ rewrite IHP1;rewrite IHP2;rsimpl.
+ mul_push ([c]);rrefl.
+ Qed.
+
+ Lemma PmulC_ok : forall c P l, (PmulC 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.
+ Qed.
+
+ Lemma Popp_ok : forall P l, (--P)@l == - P@l.
+ Proof.
+ induction P;simpl;intros.
+ Esimpl.
+ apply IHP.
+ rewrite IHP1;rewrite IHP2;rsimpl.
+ 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.
+ 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.
+ Qed.
+
+ Lemma Psub_ok : forall P' P l, (P -- P')@l == P@l - P'@l.
+ 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.
+ 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).
+ 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.
+ 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).
+ 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.
+ Qed.
+
+ Lemma Pmul_aux_ok : forall P' P l,(Pmul_aux 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.
+ Qed.
+*)
+
+(* Proof for the symmetric version *)
+ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * 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.
+ Qed.
+
+(*
+Lemma Pmul_ok : forall 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.
+ Qed.
+*)
+
+ Lemma Psquare_ok : forall 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.
+ 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.
+ Qed.
+
+ Lemma zmon_pred_ok : forall M j l,
+ Mphi (tail l) (zmon_pred j M) == Mphi l (zmon j M).
+ Proof.
+ destruct j; simpl;intros auto; rsimpl.
+ rewrite mkZmon_ok;rsimpl.
+ rewrite mkZmon_ok;simpl. rewrite jump_Pdouble_minus_one; rsimpl.
+ Qed.
+
+ Lemma mkVmon_ok : forall M i l, Mphi l (mkVmon i M) == Mphi l M*pow_pos rmul (hd 0 l) i.
+ Proof.
+ destruct M;simpl;intros;rsimpl.
+ rewrite zmon_pred_ok;simpl;rsimpl.
+ rewrite Pplus_comm;rewrite pow_pos_Pplus;rsimpl.
+ Qed.
+
+ Lemma Mcphi_ok: forall P c l,
+ let (Q,R) := CFactor P c in
+ P@l == Q@l + (phi 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.
+ 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).
+ 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) Eq); intros He; simpl.
+ rewrite (Pcompare_Eq_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) Eq); intros He; simpl.
+ rewrite (Pcompare_Eq_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.
+ 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.
+ 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.
+ 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.
+ 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.
+ 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.
+ 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.
+ 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.
+
+ Lemma PSubstL1_ok: forall 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.
+ Qed.
+
+ Lemma PSubstL_ok: forall 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.
+ Qed.
+
+ Lemma PNSubstL_ok: forall 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.
+ Qed.
+
+ (** Definition of polynomial expressions *)
+
+ Inductive PExpr : Type :=
+ | PEc : C -> PExpr
+ | PEX : positive -> PExpr
+ | PEadd : PExpr -> PExpr -> PExpr
+ | PEsub : PExpr -> PExpr -> PExpr
+ | PEmul : PExpr -> PExpr -> PExpr
+ | PEopp : PExpr -> PExpr
+ | PEpow : PExpr -> N -> PExpr.
+
+ (** evaluation of polynomial expressions towards R *)
+ Definition mk_X j := mkPinj_pred j mkX.
+
+ (** evaluation of polynomial expressions towards R *)
+
+ Fixpoint PEeval (l:list R) (pe:PExpr) {struct pe} : R :=
+ match pe with
+ | PEc c => phi c
+ | PEX j => nth 0 j l
+ | PEadd pe1 pe2 => (PEeval l pe1) + (PEeval l pe2)
+ | PEsub pe1 pe2 => (PEeval l pe1) - (PEeval l pe2)
+ | PEmul pe1 pe2 => (PEeval l pe1) * (PEeval l pe2)
+ | PEopp pe1 => - (PEeval l pe1)
+ | PEpow pe1 n => rpow (PEeval l pe1) (Cp_phi n)
+ end.
+
+Strategy expand [PEeval].
+
+ (** Correctness proofs *)
+
+ Lemma mkX_ok : forall 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.
+ 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. *)
+
+Section POWER.
+ Variable subst_l : Pol -> Pol.
+ Fixpoint Ppow_pos (res P:Pol) (p:positive){struct p} : Pol :=
+ match p with
+ | xH => subst_l (Pmul 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)
+ end.
+
+ Definition Ppow_N P n :=
+ match n with
+ | N0 => P1
+ | 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.
+ 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.
+ 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.
+
+ End POWER.
+
+ (** Normalization and rewriting *)
+
+ Section NORM_SUBST_REC.
+ Variable n : nat.
+ Variable lmp:list (C*Mon*Pol).
+ Let subst_l P := PNSubstL P lmp n n.
+ Let Pmul_subst P1 P2 := subst_l (Pmul P1 P2).
+ Let Ppow_subst := Ppow_N subst_l.
+
+ Fixpoint norm_aux (pe:PExpr) : Pol :=
+ 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)
+ | 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 :=
+ 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
+ end.
+
+ Lemma norm_subst_spec :
+ forall l pe, MPcond lmp l ->
+ PEeval l pe == (norm_subst pe)@l.
+ 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.
+ Qed.
+*)
+ Lemma norm_aux_spec :
+ forall l pe, MPcond lmp l ->
+ 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.
+ Qed.
+
+ Lemma norm_subst_spec :
+ forall l pe, MPcond lmp l ->
+ PEeval l pe == (norm_subst pe)@l.
+ Proof.
+ intros;unfold norm_subst.
+ unfold subst_l;rewrite <- PNSubstL_ok;trivial. apply norm_aux_spec. trivial.
+ Qed.
+
+ End NORM_SUBST_REC.
+
+ Fixpoint interp_PElist (l:list R) (lpe:list (PExpr*PExpr)) {struct lpe} : Prop :=
+ match lpe with
+ | nil => True
+ | (me,pe)::lpe =>
+ match lpe with
+ | nil => PEeval l me == PEeval l pe
+ | _ => PEeval l me == PEeval l pe /\ interp_PElist l lpe
+ end
+ end.
+
+ Fixpoint mon_of_pol (P:Pol) : option (C * Mon) :=
+ match P with
+ | Pc c => if (c ?=! cO) then None else Some (c, mon0)
+ | Pinj j P =>
+ match mon_of_pol P with
+ | None => None
+ | Some (c,m) => Some (c, mkZmon j m)
+ end
+ | PX P i Q =>
+ if Peq Q P0 then
+ match mon_of_pol P with
+ | None => None
+ | Some (c,m) => Some (c, mkVmon i m)
+ end
+ else None
+ end.
+
+ Fixpoint mk_monpol_list (lpe:list (PExpr * PExpr)) : list (C*Mon*Pol) :=
+ match lpe with
+ | nil => nil
+ | (me,pe)::lpe =>
+ match mon_of_pol (norm_subst 0 nil me) with
+ | None => mk_monpol_list lpe
+ | Some m => (m,norm_subst 0 nil pe):: mk_monpol_list lpe
+ end
+ end.
+
+ Lemma mon_of_pol_ok : forall P m, mon_of_pol P = Some m ->
+ forall l, [fst m] * Mphi l (snd m) == P@l.
+ Proof.
+ induction P;simpl;intros;Esimpl.
+ assert (H1 := (morph_eq CRmorph) c cO).
+ destruct (c ?=! cO).
+ discriminate.
+ inversion H;trivial;Esimpl.
+ generalize H;clear H;case_eq (mon_of_pol P).
+ intros (c1,P2) H0 H1; inversion H1; Esimpl.
+ generalize (IHP (c1, P2) H0 (jump p l)).
+ rewrite mkZmon_ok;simpl;auto.
+ intros; discriminate.
+ generalize H;clear H;change match P3 with
+ | Pc c => c ?=! cO
+ | Pinj _ _ => false
+ | PX _ _ _ => false
+ end with (P3 ?== P0).
+ assert (H := Peq_ok P3 P0).
+ destruct (P3 ?== P0).
+ case_eq (mon_of_pol P2);try intros (cc, pp); intros.
+ inversion H1.
+ simpl.
+ rewrite mkVmon_ok;simpl.
+ rewrite H;trivial;Esimpl.
+ generalize (IHP1 _ H0); simpl; intros HH; rewrite HH; rsimpl.
+ discriminate.
+ intros;discriminate.
+ Qed.
+
+ Lemma interp_PElist_ok : forall l lpe,
+ interp_PElist l lpe -> MPcond (mk_monpol_list lpe) l.
+ Proof.
+ induction lpe;simpl. trivial.
+ destruct a;simpl;intros.
+ assert (HH:=mon_of_pol_ok (norm_subst 0 nil p));
+ destruct (mon_of_pol (norm_subst 0 nil p)).
+ split.
+ rewrite <- norm_subst_spec by exact I.
+ destruct lpe;try destruct H;rewrite <- H;
+ rewrite (norm_subst_spec 0 nil); try exact I;apply HH;trivial.
+ apply IHlpe. destruct lpe;simpl;trivial. destruct H. exact H0.
+ apply IHlpe. destruct lpe;simpl;trivial. destruct H. exact H0.
+ Qed.
+
+ Lemma norm_subst_ok : forall n l lpe pe,
+ interp_PElist l lpe ->
+ PEeval l pe == (norm_subst n (mk_monpol_list lpe) pe)@l.
+ Proof.
+ intros;apply norm_subst_spec. apply interp_PElist_ok;trivial.
+ Qed.
+
+ Lemma ring_correct : forall n l lpe pe1 pe2,
+ interp_PElist l lpe ->
+ (let lmp := mk_monpol_list lpe in
+ norm_subst n lmp pe1 ?== norm_subst n lmp pe2) = true ->
+ PEeval l pe1 == PEeval l pe2.
+ Proof.
+ simpl;intros.
+ do 2 (rewrite (norm_subst_ok n l lpe);trivial).
+ apply Peq_ok;trivial.
+ Qed.
+
+
+
+ (** Generic evaluation of polynomial towards R avoiding parenthesis *)
+ Variable get_sign : C -> option C.
+ Variable get_sign_spec : sign_theory copp ceqb get_sign.
+
+
+ Section EVALUATION.
+
+ (* [mkpow x p] = x^p *)
+ Variable mkpow : R -> positive -> R.
+ (* [mkpow x p] = -(x^p) *)
+ Variable mkopp_pow : R -> positive -> R.
+ (* [mkmult_pow r x p] = r * x^p *)
+ Variable mkmult_pow : R -> R -> positive -> R.
+
+ Fixpoint mkmult_rec (r:R) (lm:list (R*positive)) {struct lm}: R :=
+ match lm with
+ | nil => r
+ | cons (x,p) t => mkmult_rec (mkmult_pow r x p) t
+ end.
+
+ Definition mkmult1 lm :=
+ match lm with
+ | nil => 1
+ | cons (x,p) t => mkmult_rec (mkpow x p) t
+ end.
+
+ Definition mkmultm1 lm :=
+ match lm with
+ | nil => ropp rI
+ | cons (x,p) t => mkmult_rec (mkopp_pow x p) t
+ end.
+
+ Definition mkmult_c_pos c lm :=
+ if c ?=! cI then mkmult1 (rev' lm)
+ else mkmult_rec [c] (rev' lm).
+
+ Definition mkmult_c c lm :=
+ match get_sign c with
+ | None => mkmult_c_pos c lm
+ | Some c' =>
+ if c' ?=! cI then mkmultm1 (rev' lm)
+ else mkmult_rec [c] (rev' lm)
+ end.
+
+ Definition mkadd_mult rP c lm :=
+ match get_sign c with
+ | None => rP + mkmult_c_pos c lm
+ | Some c' => rP - mkmult_c_pos c' lm
+ end.
+
+ Definition add_pow_list (r:R) n l :=
+ match n with
+ | N0 => l
+ | Npos p => (r,p)::l
+ end.
+
+ Fixpoint add_mult_dev
+ (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
+ 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)
+ | PX P i Q =>
+ let rP := add_mult_dev rP P fv (Nplus (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)
+ 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)
+ | PX P i Q =>
+ let rP := mult_dev P fv (Nplus (Npos i) n) lm in
+ if Q ?== P0 then rP
+ else
+ let lmq := add_pow_list (hd 0 fv) n lm in
+ add_mult_dev rP Q (tail fv) N0 lmq
+ end.
+
+ Definition Pphi_avoid fv P := mult_dev P fv N0 nil.
+
+ Fixpoint r_list_pow (l:list (R*positive)) : R :=
+ match l with
+ | nil => rI
+ | cons (r,p) l => pow_pos rmul r p * r_list_pow l
+ end.
+
+ Hypothesis mkpow_spec : forall r p, mkpow r p == pow_pos rmul r p.
+ Hypothesis mkopp_pow_spec : forall r p, mkopp_pow r p == - (pow_pos rmul r p).
+ Hypothesis mkmult_pow_spec : forall r x p, mkmult_pow r x p == r * pow_pos rmul x p.
+
+ Lemma mkmult_rec_ok : forall lm r, mkmult_rec r lm == r * r_list_pow lm.
+ Proof.
+ induction lm;intros;simpl;Esimpl.
+ destruct a as (x,p);Esimpl.
+ rewrite IHlm. rewrite mkmult_pow_spec. Esimpl.
+ Qed.
+
+ Lemma mkmult1_ok : forall lm, mkmult1 lm == r_list_pow lm.
+ Proof.
+ destruct lm;simpl;Esimpl.
+ destruct p. rewrite mkmult_rec_ok;rewrite mkpow_spec;Esimpl.
+ Qed.
+
+ Lemma mkmultm1_ok : forall lm, mkmultm1 lm == - r_list_pow lm.
+ Proof.
+ destruct lm;simpl;Esimpl.
+ destruct p;rewrite mkmult_rec_ok. rewrite mkopp_pow_spec;Esimpl.
+ Qed.
+
+ Lemma r_list_pow_rev : forall l, r_list_pow (rev' l) == r_list_pow l.
+ Proof.
+ assert
+ (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.
+ intros;unfold rev'. rewrite H;simpl;Esimpl.
+ Qed.
+
+ Lemma mkmult_c_pos_ok : forall c lm, mkmult_c_pos c lm == [c]* r_list_pow lm.
+ Proof.
+ intros;unfold mkmult_c_pos;simpl.
+ assert (H := (morph_eq CRmorph) c cI).
+ rewrite <- r_list_pow_rev; destruct (c ?=! cI).
+ rewrite H;trivial;Esimpl.
+ apply mkmult1_ok. apply mkmult_rec_ok.
+ Qed.
+
+ Lemma mkmult_c_ok : forall c lm, mkmult_c c lm == [c] * r_list_pow lm.
+ Proof.
+ intros;unfold mkmult_c;simpl.
+ case_eq (get_sign c);intros.
+ assert (H1 := (morph_eq CRmorph) c0 cI).
+ destruct (c0 ?=! cI).
+ rewrite (CRmorph.(morph_eq) _ _ (get_sign_spec.(sign_spec) _ H)). Esimpl. rewrite H1;trivial.
+ rewrite <- r_list_pow_rev;trivial;Esimpl.
+ apply mkmultm1_ok.
+ rewrite <- r_list_pow_rev; apply mkmult_rec_ok.
+ apply mkmult_c_pos_ok.
+Qed.
+
+ Lemma mkadd_mult_ok : forall rP c lm, mkadd_mult rP c lm == rP + [c]*r_list_pow lm.
+ Proof.
+ intros;unfold mkadd_mult.
+ case_eq (get_sign c);intros.
+ rewrite (CRmorph.(morph_eq) _ _ (get_sign_spec.(sign_spec) _ H));Esimpl.
+ rewrite mkmult_c_pos_ok;Esimpl.
+ rewrite mkmult_c_pos_ok;Esimpl.
+ Qed.
+
+ Lemma add_pow_list_ok :
+ forall r n l, r_list_pow (add_pow_list r n l) == pow_N rI rmul r n * r_list_pow l.
+ Proof.
+ destruct n;simpl;intros;Esimpl.
+ 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.
+ Proof.
+ induction P;simpl;intros.
+ 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
+ | PX _ _ _ => false
+ end) with (Peq P3 P0).
+ change match n with
+ | N0 => Npos p
+ | Npos q => Npos (p + q)
+ end with (Nplus (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.
+ 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.
+ Proof.
+ induction P;simpl;intros;Esimpl.
+ rewrite mkmult_c_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
+ | PX _ _ _ => false
+ end) with (Peq P3 P0).
+ change match n with
+ | N0 => Npos p
+ | Npos q => Npos (p + q)
+ end with (Nplus (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 add_mult_dev_ok. rewrite IHP1; rewrite add_pow_list_ok.
+ destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl.
+ Qed.
+
+ Lemma Pphi_avoid_ok : forall P fv, Pphi_avoid fv P == P@fv.
+ Proof.
+ unfold Pphi_avoid;intros;rewrite mult_dev_ok;simpl;Esimpl.
+ Qed.
+
+ End EVALUATION.
+
+ Definition Pphi_pow :=
+ let mkpow x p :=
+ match p with xH => x | _ => rpow x (Cp_phi (Npos p)) end in
+ let mkopp_pow x p := ropp (mkpow x p) in
+ 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),
+ 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.
+
+ 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.
+ Qed.
+
+ Lemma ring_rw_pow_correct : forall n lH l,
+ interp_PElist l lH ->
+ forall lmp, mk_monpol_list lH = lmp ->
+ forall pe npe, norm_subst n lmp pe = npe ->
+ 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.
+ apply norm_subst_ok. trivial.
+ Qed.
+
+ Fixpoint mkmult_pow (r x:R) (p: positive) {struct p} : R :=
+ match p with
+ | xH => r*x
+ | xO p => mkmult_pow (mkmult_pow r x p) x p
+ | xI p => mkmult_pow (mkmult_pow (r*x) x p) x p
+ end.
+
+ Definition mkpow x p :=
+ match p with
+ | xH => x
+ | xO p => mkmult_pow x x (Pdouble_minus_one 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)
+ | 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.
+ Proof.
+ induction p;intros;simpl;Esimpl.
+ repeat rewrite IHp;Esimpl.
+ repeat rewrite IHp;Esimpl.
+ Qed.
+
+ Lemma mkpow_ok : forall p x, mkpow x p == pow_pos rmul 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.
+ Qed.
+
+ Lemma mkopp_pow_ok : forall p x, mkopp_pow x p == - pow_pos rmul 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.
+ 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.
+ Qed.
+
+ Lemma ring_rw_correct : forall n lH l,
+ interp_PElist l lH ->
+ forall lmp, mk_monpol_list lH = lmp ->
+ forall pe npe, norm_subst n lmp pe = npe ->
+ PEeval l pe == Pphi_dev l npe.
+ Proof.
+ intros n lH l H1 lmp Heq1 pe npe Heq2.
+ rewrite Pphi_dev_ok. rewrite <- Heq2;rewrite <- Heq1.
+ apply norm_subst_ok. trivial.
+ Qed.
+
+
+End MakeRingPol.
+
diff --git a/plugins/setoid_ring/Ring_tac.v b/plugins/setoid_ring/Ring_tac.v
new file mode 100644
index 00000000..d33e9a82
--- /dev/null
+++ b/plugins/setoid_ring/Ring_tac.v
@@ -0,0 +1,434 @@
+Set Implicit Arguments.
+Require Import Setoid.
+Require Import BinPos.
+Require Import Ring_polynom.
+Require Import BinList.
+Require Import InitialRing.
+Require Import Quote.
+Declare ML Module "newring_plugin".
+
+
+(* adds a definition t' on the normal form of t and an hypothesis id
+ stating that t = t' (tries to produces a proof as small as possible) *)
+Ltac compute_assertion eqn t' t :=
+ let nft := eval vm_compute in t in
+ pose (t' := nft);
+ assert (eqn : t = t');
+ [vm_cast_no_check (refl_equal t')|idtac].
+
+Ltac relation_carrier req :=
+ let ty := type of req in
+ match eval hnf in ty with
+ ?R -> _ => R
+ | _ => fail 1000 "Equality has no relation type"
+ end.
+
+Ltac Get_goal := match goal with [|- ?G] => G end.
+
+(********************************************************************)
+(* Tacticals to build reflexive tactics *)
+
+Ltac OnEquation req :=
+ match goal with
+ | |- req ?lhs ?rhs => (fun f => f lhs rhs)
+ | _ => (fun _ => fail "Goal is not an equation (of expected equality)")
+ end.
+
+Ltac OnEquationHyp req h :=
+ match type of h with
+ | req ?lhs ?rhs => fun f => f lhs rhs
+ | _ => (fun _ => fail "Hypothesis is not an equation (of expected equality)")
+ end.
+
+(* Note: auxiliary subgoals in reverse order *)
+Ltac OnMainSubgoal H ty :=
+ match ty with
+ | _ -> ?ty' =>
+ let subtac := OnMainSubgoal H ty' in
+ fun kont => lapply H; [clear H; intro H; subtac kont | idtac]
+ | _ => (fun kont => kont())
+ end.
+
+(* A generic pattern to have reflexive tactics do some computation:
+ lemmas of the form [forall x', x=x' -> P(x')] are understood as:
+ compute the normal form of x, instantiate x' with it, prove
+ hypothesis x=x' with vm_compute and reflexivity, and pass the
+ instantiated lemma to the continuation.
+ *)
+Ltac ProveLemmaHyp lemma :=
+ match type of lemma with
+ forall x', ?x = x' -> _ =>
+ (fun kont =>
+ let x' := fresh "res" in
+ let H := fresh "res_eq" in
+ compute_assertion H x' x;
+ let lemma' := constr:(lemma x' H) in
+ kont lemma';
+ (clear H||idtac"ProveLemmaHyp: cleanup failed");
+ subst x')
+ | _ => (fun _ => fail "ProveLemmaHyp: lemma not of the expected form")
+ end.
+
+Ltac ProveLemmaHyps lemma :=
+ match type of lemma with
+ forall x', ?x = x' -> _ =>
+ (fun kont =>
+ let x' := fresh "res" in
+ let H := fresh "res_eq" in
+ compute_assertion H x' x;
+ let lemma' := constr:(lemma x' H) in
+ ProveLemmaHyps lemma' kont;
+ (clear H||idtac"ProveLemmaHyps: cleanup failed");
+ subst x')
+ | _ => (fun kont => kont lemma)
+ end.
+
+(*
+Ltac ProveLemmaHyps lemma := (* expects a continuation *)
+ let try_step := ProveLemmaHyp lemma in
+ (fun kont =>
+ try_step ltac:(fun lemma' => ProveLemmaHyps lemma' kont) ||
+ kont lemma).
+*)
+Ltac ApplyLemmaThen lemma expr kont :=
+ let lem := constr:(lemma expr) in
+ ProveLemmaHyp lem ltac:(fun lem' =>
+ let Heq := fresh "thm" in
+ assert (Heq:=lem');
+ OnMainSubgoal Heq ltac:(type of Heq) ltac:(fun _ => kont Heq);
+ (clear Heq||idtac"ApplyLemmaThen: cleanup failed")).
+(*
+Ltac ApplyLemmaThenAndCont lemma expr tac CONT_tac cont_arg :=
+ let pe :=
+ match type of (lemma expr) with
+ forall pe', ?pe = pe' -> _ => pe
+ | _ => fail 1 "ApplyLemmaThenAndCont: cannot find norm expression"
+ end in
+ let pe' := fresh "expr_nf" in
+ let nf_pe := fresh "pe_eq" in
+ compute_assertion nf_pe pe' pe;
+ let Heq := fresh "thm" in
+ (assert (Heq:=lemma pe pe' H) || fail "anomaly: failed to apply lemma");
+ clear nf_pe;
+ OnMainSubgoal Heq ltac:(type of Heq)
+ ltac:(try tac Heq; clear Heq pe';CONT_tac cont_arg)).
+*)
+Ltac ApplyLemmaThenAndCont lemma expr tac CONT_tac :=
+ ApplyLemmaThen lemma expr
+ ltac:(fun lemma' => try tac lemma'; CONT_tac()).
+
+(* General scheme of reflexive tactics using of correctness lemma
+ that involves normalisation of one expression
+ - [FV_tac term fv] is a tactic that adds the atomic expressions
+ of [term] into [fv]
+ - [SYN_tac term fv] reifies [term] given the list of atomic expressions
+ - [LEMMA_tac fv kont] computes the correctness lemma and passes it to
+ continuation kont
+ - [MAIN_tac H] process H which is the conclusion of the correctness lemma
+ instantiated with each reified term
+ - [fv] is the initial value of atomic expressions (to be completed by
+ the reification of the terms
+ - [terms] the list (a constr of type list) of terms to reify and process.
+ *)
+Ltac ReflexiveRewriteTactic
+ FV_tac SYN_tac LEMMA_tac MAIN_tac fv terms :=
+ (* extend the atom list *)
+ let fv := list_fold_left FV_tac fv terms in
+ let RW_tac lemma :=
+ let fcons term CONT_tac :=
+ let expr := SYN_tac term fv in
+ let main H :=
+ match type of H with
+ | (?req _ ?rhs) => change (req term rhs) in H
+ end;
+ MAIN_tac H in
+ (ApplyLemmaThenAndCont lemma expr main CONT_tac) in
+ (* rewrite steps *)
+ lazy_list_fold_right fcons ltac:(fun _=>idtac) terms in
+ LEMMA_tac fv RW_tac.
+
+(********************************************************)
+
+Ltac FV_hypo_tac mkFV req lH :=
+ let R := relation_carrier req in
+ let FV_hypo_l_tac h :=
+ match h with @mkhypo (req ?pe _) _ => mkFV pe end in
+ let FV_hypo_r_tac h :=
+ match h with @mkhypo (req _ ?pe) _ => mkFV pe end in
+ let fv := list_fold_right FV_hypo_l_tac (@nil R) lH in
+ list_fold_right FV_hypo_r_tac fv lH.
+
+Ltac mkHyp_tac C req Reify lH :=
+ let mkHyp h res :=
+ match h with
+ | @mkhypo (req ?r1 ?r2) _ =>
+ let pe1 := Reify r1 in
+ let pe2 := Reify r2 in
+ constr:(cons (pe1,pe2) res)
+ | _ => fail 1 "hypothesis is not a ring equality"
+ end in
+ list_fold_right mkHyp (@nil (PExpr C * PExpr C)) lH.
+
+Ltac proofHyp_tac lH :=
+ let get_proof h :=
+ match h with
+ | @mkhypo _ ?p => p
+ end in
+ let rec bh l :=
+ match l with
+ | nil => constr:(I)
+ | cons ?h nil => get_proof h
+ | cons ?h ?tl =>
+ let l := get_proof h in
+ let r := bh tl in
+ constr:(conj l r)
+ end in
+ bh lH.
+
+Ltac get_MonPol lemma :=
+ match type of lemma with
+ | context [(mk_monpol_list ?cO ?cI ?cadd ?cmul ?csub ?copp ?cdiv ?ceqb _)] =>
+ constr:(mk_monpol_list cO cI cadd cmul csub copp cdiv ceqb)
+ | _ => fail 1 "ring/field anomaly: bad correctness lemma (get_MonPol)"
+ end.
+
+(********************************************************)
+
+(* Building the atom list of a ring expression *)
+Ltac FV Cst CstPow add mul sub opp pow t fv :=
+ let rec TFV t fv :=
+ let f :=
+ match Cst t with
+ | NotConstant =>
+ match t with
+ | (add ?t1 ?t2) => fun _ => TFV t2 ltac:(TFV t1 fv)
+ | (mul ?t1 ?t2) => fun _ => TFV t2 ltac:(TFV t1 fv)
+ | (sub ?t1 ?t2) => fun _ => TFV t2 ltac:(TFV t1 fv)
+ | (opp ?t1) => fun _ => TFV t1 fv
+ | (pow ?t1 ?n) =>
+ match CstPow n with
+ | InitialRing.NotConstant => fun _ => AddFvTail t fv
+ | _ => fun _ => TFV t1 fv
+ end
+ | _ => fun _ => AddFvTail t fv
+ end
+ | _ => fun _ => fv
+ end in
+ f()
+ in TFV t fv.
+
+ (* syntaxification of ring expressions *)
+Ltac mkPolexpr C Cst CstPow radd rmul rsub ropp rpow t fv :=
+ let rec mkP t :=
+ let f :=
+ match Cst t with
+ | InitialRing.NotConstant =>
+ match t with
+ | (radd ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(PEadd e1 e2)
+ | (rmul ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(PEmul e1 e2)
+ | (rsub ?t1 ?t2) =>
+ fun _ =>
+ let e1 := mkP t1 in
+ let e2 := mkP t2 in constr:(PEsub e1 e2)
+ | (ropp ?t1) =>
+ fun _ =>
+ let e1 := mkP t1 in constr:(PEopp e1)
+ | (rpow ?t1 ?n) =>
+ match CstPow n with
+ | InitialRing.NotConstant =>
+ fun _ => let p := Find_at t fv in constr:(PEX C p)
+ | ?c => fun _ => let e1 := mkP t1 in constr:(PEpow e1 c)
+ end
+ | _ =>
+ fun _ => let p := Find_at t fv in constr:(PEX C p)
+ end
+ | ?c => fun _ => constr:(@PEc C c)
+ end in
+ f ()
+ in mkP t.
+
+(* packaging the ring structure *)
+
+Ltac PackRing F req sth ext morph arth cst_tac pow_tac lemma1 lemma2 pre post :=
+ let RNG :=
+ match type of lemma1 with
+ | context
+ [@PEeval ?R ?rO ?add ?mul ?sub ?opp ?C ?phi ?Cpow ?powphi ?pow _ _] =>
+ (fun proj => proj
+ cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2)
+ | _ => fail 1 "field anomaly: bad correctness lemma (parse)"
+ end in
+ F RNG.
+
+Ltac get_Carrier RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ R).
+
+Ltac get_Eq RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ req).
+
+Ltac get_Pre RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ pre).
+
+Ltac get_Post RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ post).
+
+Ltac get_NormLemma RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ lemma1).
+
+Ltac get_SimplifyLemma RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ lemma2).
+
+Ltac get_RingFV RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ FV cst_tac pow_tac add mul sub opp pow).
+
+Ltac get_RingMeta RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ mkPolexpr C cst_tac pow_tac add mul sub opp pow).
+
+Ltac get_RingHypTac RNG :=
+ RNG ltac:(fun cst_tac pow_tac pre post
+ R req add mul sub opp C Cpow powphi pow lemma1 lemma2 =>
+ let mkPol := mkPolexpr C cst_tac pow_tac add mul sub opp pow in
+ fun fv lH => mkHyp_tac C req ltac:(fun t => mkPol t fv) lH).
+
+(* ring tactics *)
+
+Definition ring_subst_niter := (10*10*10)%nat.
+
+Ltac Ring RNG lemma lH :=
+ let req := get_Eq RNG in
+ OnEquation req ltac:(fun lhs rhs =>
+ let mkFV := get_RingFV RNG in
+ let mkPol := get_RingMeta RNG in
+ let mkHyp := get_RingHypTac RNG in
+ let fv := FV_hypo_tac mkFV ltac:(get_Eq RNG) lH in
+ let fv := mkFV lhs fv in
+ let fv := mkFV rhs fv in
+ check_fv fv;
+ let pe1 := mkPol lhs fv in
+ let pe2 := mkPol rhs fv in
+ let lpe := mkHyp fv lH in
+ let vlpe := fresh "hyp_list" in
+ let vfv := fresh "fv_list" in
+ pose (vlpe := lpe);
+ pose (vfv := fv);
+ (apply (lemma vfv vlpe pe1 pe2)
+ || fail "typing error while applying ring");
+ [ ((let prh := proofHyp_tac lH in exact prh)
+ || idtac "can not automatically proof hypothesis :";
+ idtac " maybe a left member of a hypothesis is not a monomial")
+ | vm_compute;
+ (exact (refl_equal true) || fail "not a valid ring equation")]).
+
+Ltac Ring_norm_gen f RNG lemma lH rl :=
+ let mkFV := get_RingFV RNG in
+ let mkPol := get_RingMeta RNG in
+ let mkHyp := get_RingHypTac RNG in
+ let mk_monpol := get_MonPol lemma in
+ let fv := FV_hypo_tac mkFV ltac:(get_Eq RNG) lH in
+ let lemma_tac fv kont :=
+ let lpe := mkHyp fv lH in
+ let vlpe := fresh "list_hyp" in
+ let vlmp := fresh "list_hyp_norm" in
+ let vlmp_eq := fresh "list_hyp_norm_eq" in
+ let prh := proofHyp_tac lH in
+ pose (vlpe := lpe);
+ compute_assertion vlmp_eq vlmp (mk_monpol vlpe);
+ let H := fresh "ring_lemma" in
+ (assert (H := lemma vlpe fv prh vlmp vlmp_eq)
+ || fail "type error when build the rewriting lemma");
+ clear vlmp_eq;
+ kont H;
+ (clear H||idtac"Ring_norm_gen: cleanup failed");
+ subst vlpe vlmp in
+ let simpl_ring H := (protect_fv "ring" in H; f H) in
+ ReflexiveRewriteTactic mkFV mkPol lemma_tac simpl_ring fv rl.
+
+Ltac Ring_gen RNG lH rl :=
+ let lemma := get_NormLemma RNG in
+ get_Pre RNG ();
+ Ring RNG (lemma ring_subst_niter) lH.
+
+Tactic Notation (at level 0) "ring" :=
+ let G := Get_goal in
+ ring_lookup (PackRing Ring_gen) [] G.
+
+Tactic Notation (at level 0) "ring" "[" constr_list(lH) "]" :=
+ let G := Get_goal in
+ ring_lookup (PackRing Ring_gen) [lH] G.
+
+(* Simplification *)
+
+Ltac Ring_simplify_gen f RNG lH rl :=
+ let lemma := get_SimplifyLemma RNG in
+ let l := fresh "to_rewrite" in
+ pose (l:= rl);
+ generalize (refl_equal l);
+ unfold l at 2;
+ get_Pre RNG ();
+ let rl :=
+ match goal with
+ | [|- l = ?RL -> _ ] => RL
+ | _ => fail 1 "ring_simplify anomaly: bad goal after pre"
+ end in
+ let Heq := fresh "Heq" in
+ intros Heq;clear Heq l;
+ Ring_norm_gen f RNG (lemma ring_subst_niter) lH rl;
+ get_Post RNG ().
+
+Ltac Ring_simplify := Ring_simplify_gen ltac:(fun H => rewrite H).
+
+Tactic Notation (at level 0) "ring_simplify" constr_list(rl) :=
+ let G := Get_goal in
+ ring_lookup (PackRing Ring_simplify) [] rl G.
+
+Tactic Notation (at level 0)
+ "ring_simplify" "[" constr_list(lH) "]" constr_list(rl) :=
+ let G := Get_goal in
+ ring_lookup (PackRing Ring_simplify) [lH] rl G.
+
+(* MON DIEU QUE C'EST MOCHE !!!!!!!!!!!!! *)
+
+Tactic Notation "ring_simplify" constr_list(rl) "in" hyp(H):=
+ let G := Get_goal in
+ let t := type of H in
+ let g := fresh "goal" in
+ set (g:= G);
+ generalize H;clear H;
+ ring_lookup (PackRing Ring_simplify) [] rl t;
+ intro H;
+ unfold g;clear g.
+
+Tactic Notation
+ "ring_simplify" "["constr_list(lH)"]" constr_list(rl) "in" hyp(H):=
+ let G := Get_goal in
+ let t := type of H in
+ let g := fresh "goal" in
+ set (g:= G);
+ generalize H;clear H;
+ ring_lookup (PackRing Ring_simplify) [lH] rl t;
+ intro H;
+ unfold g;clear g.
+
diff --git a/plugins/setoid_ring/Ring_theory.v b/plugins/setoid_ring/Ring_theory.v
new file mode 100644
index 00000000..b3250a51
--- /dev/null
+++ b/plugins/setoid_ring/Ring_theory.v
@@ -0,0 +1,608 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Setoid.
+Require Import BinPos.
+Require Import BinNat.
+
+Set Implicit Arguments.
+
+Module RingSyntax.
+Reserved Notation "x ?=! y" (at level 70, no associativity).
+Reserved Notation "x +! y " (at level 50, left associativity).
+Reserved Notation "x -! y" (at level 50, left associativity).
+Reserved Notation "x *! y" (at level 40, left associativity).
+Reserved Notation "-! x" (at level 35, right associativity).
+
+Reserved Notation "[ x ]" (at level 0).
+
+Reserved Notation "x ?== y" (at level 70, no associativity).
+Reserved Notation "x -- y" (at level 50, left associativity).
+Reserved Notation "x ** y" (at level 40, left associativity).
+Reserved Notation "-- x" (at level 35, right associativity).
+
+Reserved Notation "x == y" (at level 70, no associativity).
+End RingSyntax.
+Import RingSyntax.
+
+Section Power.
+ Variable R:Type.
+ Variable rI : R.
+ Variable rmul : R -> R -> R.
+ Variable req : R -> R -> Prop.
+ Variable Rsth : Setoid_Theory R req.
+ Notation "x * y " := (rmul x y).
+ Notation "x == y" := (req x y).
+
+ Hypothesis mul_ext :
+ forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2.
+ Hypothesis mul_comm : forall x y, x * y == y * x.
+ Hypothesis mul_assoc : forall x y z, x * (y * z) == (x * y) * z.
+ Add Setoid R req Rsth as R_set_Power.
+ Add Morphism rmul : rmul_ext_Power. exact mul_ext. Qed.
+
+
+ Fixpoint pow_pos (x:R) (i:positive) {struct i}: R :=
+ match i with
+ | xH => x
+ | xO i => let p := pow_pos x i in rmul p p
+ | xI i => let p := pow_pos x i in rmul x (rmul p p)
+ end.
+
+ Lemma pow_pos_Psucc : forall x j, pow_pos x (Psucc j) == x * pow_pos x j.
+ Proof.
+ induction j;simpl.
+ rewrite IHj.
+ rewrite (mul_comm x (pow_pos x j *pow_pos x j)).
+ setoid_rewrite (mul_comm x (pow_pos x j)) at 2.
+ repeat rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
+ repeat rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
+ apply (Seq_refl _ _ Rsth).
+ Qed.
+
+ Lemma pow_pos_Pplus : forall x i j, pow_pos x (i + j) == pow_pos x i * pow_pos x j.
+ Proof.
+ intro x;induction i;intros.
+ rewrite xI_succ_xO;rewrite Pplus_one_succ_r.
+ rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
+ repeat rewrite IHi.
+ rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite pow_pos_Psucc.
+ simpl;repeat rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
+ rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
+ repeat rewrite IHi;rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
+ rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite pow_pos_Psucc;
+ simpl. apply (Seq_refl _ _ Rsth).
+ Qed.
+
+ Definition pow_N (x:R) (p:N) :=
+ match p with
+ | N0 => rI
+ | Npos p => pow_pos x p
+ end.
+
+ Definition id_phi_N (x:N) : N := x.
+
+ Lemma pow_N_pow_N : forall x n, pow_N x (id_phi_N n) == pow_N x n.
+ Proof.
+ intros; apply (Seq_refl _ _ Rsth).
+ Qed.
+
+End Power.
+
+Section DEFINITIONS.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable req : R -> R -> Prop.
+ 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).
+
+ (** Semi Ring *)
+ Record semi_ring_theory : Prop := mk_srt {
+ SRadd_0_l : forall n, 0 + n == n;
+ SRadd_comm : forall n m, n + m == m + n ;
+ SRadd_assoc : forall n m p, n + (m + p) == (n + m) + p;
+ SRmul_1_l : forall n, 1*n == n;
+ SRmul_0_l : forall n, 0*n == 0;
+ SRmul_comm : forall n m, n*m == m*n;
+ SRmul_assoc : forall n m p, n*(m*p) == (n*m)*p;
+ SRdistr_l : forall n m p, (n + m)*p == n*p + m*p
+ }.
+
+ (** Almost Ring *)
+(*Almost ring are no ring : Ropp_def is missing **)
+ Record almost_ring_theory : Prop := mk_art {
+ ARadd_0_l : forall x, 0 + x == x;
+ ARadd_comm : forall x y, x + y == y + x;
+ ARadd_assoc : forall x y z, x + (y + z) == (x + y) + z;
+ ARmul_1_l : forall x, 1 * x == x;
+ ARmul_0_l : forall x, 0 * x == 0;
+ ARmul_comm : forall x y, x * y == y * x;
+ ARmul_assoc : forall x y z, x * (y * z) == (x * y) * z;
+ ARdistr_l : forall x y z, (x + y) * z == (x * z) + (y * z);
+ ARopp_mul_l : forall x y, -(x * y) == -x * y;
+ ARopp_add : forall x y, -(x + y) == -x + -y;
+ ARsub_def : forall x y, x - y == x + -y
+ }.
+
+ (** Ring *)
+ Record ring_theory : Prop := mk_rt {
+ Radd_0_l : forall x, 0 + x == x;
+ Radd_comm : forall x y, x + y == y + x;
+ Radd_assoc : forall x y z, x + (y + z) == (x + y) + z;
+ Rmul_1_l : forall x, 1 * x == x;
+ Rmul_comm : forall x y, x * y == y * x;
+ Rmul_assoc : forall x y z, x * (y * z) == (x * y) * z;
+ Rdistr_l : forall x y z, (x + y) * z == (x * z) + (y * z);
+ Rsub_def : forall x y, x - y == x + -y;
+ Ropp_def : forall x, x + (- x) == 0
+ }.
+
+ (** Equality is extensional *)
+
+ Record sring_eq_ext : Prop := mk_seqe {
+ (* SRing operators are compatible with equality *)
+ SRadd_ext :
+ forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 + y1 == x2 + y2;
+ SRmul_ext :
+ forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2
+ }.
+
+ Record ring_eq_ext : Prop := mk_reqe {
+ (* Ring operators are compatible with equality *)
+ Radd_ext :
+ forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 + y1 == x2 + y2;
+ Rmul_ext :
+ forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2;
+ Ropp_ext : forall x1 x2, x1 == x2 -> -x1 == -x2
+ }.
+
+ (** Interpretation morphisms definition*)
+ Section MORPHISM.
+ Variable C:Type.
+ Variable (cO cI : C) (cadd cmul csub : C->C->C) (copp : C->C).
+ Variable ceqb : C->C->bool.
+ (* [phi] est un morphisme de [C] dans [R] *)
+ Variable phi : C -> R.
+ Notation "x +! y" := (cadd x y). Notation "x -! y " := (csub x y).
+ Notation "x *! y " := (cmul x y). Notation "-! x" := (copp x).
+ Notation "x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x).
+
+(*for semi rings*)
+ Record semi_morph : Prop := mkRmorph {
+ Smorph0 : [cO] == 0;
+ Smorph1 : [cI] == 1;
+ Smorph_add : forall x y, [x +! y] == [x]+[y];
+ Smorph_mul : forall x y, [x *! y] == [x]*[y];
+ Smorph_eq : forall x y, x?=!y = true -> [x] == [y]
+ }.
+
+(* for rings*)
+ Record ring_morph : Prop := mkmorph {
+ morph0 : [cO] == 0;
+ morph1 : [cI] == 1;
+ morph_add : forall x y, [x +! y] == [x]+[y];
+ morph_sub : forall x y, [x -! y] == [x]-[y];
+ morph_mul : forall x y, [x *! y] == [x]*[y];
+ morph_opp : forall x, [-!x] == -[x];
+ morph_eq : forall x y, x?=!y = true -> [x] == [y]
+ }.
+
+ Section SIGN.
+ Variable get_sign : C -> option C.
+ Record sign_theory : Prop := mksign_th {
+ sign_spec : forall c c', get_sign c = Some c' -> c ?=! -! c' = true
+ }.
+ End SIGN.
+
+ Definition get_sign_None (c:C) := @None C.
+
+ Lemma get_sign_None_th : sign_theory get_sign_None.
+ Proof. constructor;intros;discriminate. Qed.
+
+ Section DIV.
+ Variable cdiv: C -> C -> C*C.
+ Record div_theory : Prop := mkdiv_th {
+ div_eucl_th : forall a b, let (q,r) := cdiv a b in [a] == [b *! q +! r]
+ }.
+ End DIV.
+
+ End MORPHISM.
+
+ (** Identity is a morphism *)
+ Variable Rsth : Setoid_Theory R req.
+ Add Setoid R req Rsth as R_setoid1.
+ Variable reqb : R->R->bool.
+ Hypothesis morph_req : forall x y, (reqb x y) = true -> x == y.
+ Definition IDphi (x:R) := x.
+ Lemma IDmorph : ring_morph rO rI radd rmul rsub ropp reqb IDphi.
+ Proof.
+ apply (mkmorph rO rI radd rmul rsub ropp reqb IDphi);intros;unfold IDphi;
+ try apply (Seq_refl _ _ Rsth);auto.
+ Qed.
+
+ (** Specification of the power function *)
+ Section POWER.
+ Variable Cpow : Set.
+ Variable Cp_phi : N -> Cpow.
+ Variable rpow : R -> Cpow -> R.
+
+ Record power_theory : Prop := mkpow_th {
+ rpow_pow_N : forall r n, req (rpow r (Cp_phi n)) (pow_N rI rmul r n)
+ }.
+
+ End POWER.
+
+ Definition pow_N_th := mkpow_th id_phi_N (pow_N rI rmul) (pow_N_pow_N rI rmul Rsth).
+
+
+End DEFINITIONS.
+
+
+
+Section ALMOST_RING.
+ Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable req : R -> R -> Prop.
+ 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).
+
+ (** Leibniz equality leads to a setoid theory and is extensional*)
+ Lemma Eqsth : Setoid_Theory R (@eq R).
+ Proof. constructor;red;intros;subst;trivial. Qed.
+
+ Lemma Eq_s_ext : sring_eq_ext radd rmul (@eq R).
+ Proof. constructor;intros;subst;trivial. Qed.
+
+ Lemma Eq_ext : ring_eq_ext radd rmul ropp (@eq R).
+ Proof. constructor;intros;subst;trivial. Qed.
+
+ Variable Rsth : Setoid_Theory R req.
+ Add Setoid R req Rsth as R_setoid2.
+ Ltac sreflexivity := apply (Seq_refl _ _ Rsth).
+
+ Section SEMI_RING.
+ Variable SReqe : sring_eq_ext radd rmul req.
+ Add Morphism radd : radd_ext1. exact (SRadd_ext SReqe). Qed.
+ Add Morphism rmul : rmul_ext1. exact (SRmul_ext SReqe). Qed.
+ Variable SRth : semi_ring_theory 0 1 radd rmul req.
+
+ (** Every semi ring can be seen as an almost ring, by taking :
+ -x = x and x - y = x + y *)
+ Definition SRopp (x:R) := x. Notation "- x" := (SRopp x).
+
+ Definition SRsub x y := x + -y. Notation "x - y " := (SRsub x y).
+
+ Lemma SRopp_ext : forall x y, x == y -> -x == -y.
+ Proof. intros x y H;exact H. Qed.
+
+ Lemma SReqe_Reqe : ring_eq_ext radd rmul SRopp req.
+ Proof.
+ constructor. exact (SRadd_ext SReqe). exact (SRmul_ext SReqe).
+ exact SRopp_ext.
+ Qed.
+
+ Lemma SRopp_mul_l : forall x y, -(x * y) == -x * y.
+ Proof. intros;sreflexivity. Qed.
+
+ Lemma SRopp_add : forall x y, -(x + y) == -x + -y.
+ Proof. intros;sreflexivity. Qed.
+
+
+ Lemma SRsub_def : forall x y, x - y == x + -y.
+ Proof. intros;sreflexivity. Qed.
+
+ Lemma SRth_ARth : almost_ring_theory 0 1 radd rmul SRsub SRopp req.
+ Proof (mk_art 0 1 radd rmul SRsub SRopp req
+ (SRadd_0_l SRth) (SRadd_comm SRth) (SRadd_assoc SRth)
+ (SRmul_1_l SRth) (SRmul_0_l SRth)
+ (SRmul_comm SRth) (SRmul_assoc SRth) (SRdistr_l SRth)
+ SRopp_mul_l SRopp_add SRsub_def).
+
+ (** Identity morphism for semi-ring equipped with their almost-ring structure*)
+ Variable reqb : R->R->bool.
+
+ Hypothesis morph_req : forall x y, (reqb x y) = true -> x == y.
+
+ Definition SRIDmorph : ring_morph 0 1 radd rmul SRsub SRopp req
+ 0 1 radd rmul SRsub SRopp reqb (@IDphi R).
+ Proof.
+ apply mkmorph;intros;try sreflexivity. unfold IDphi;auto.
+ Qed.
+
+ (* a semi_morph can be extended to a ring_morph for the almost_ring derived
+ from a semi_ring, provided the ring is a setoid (we only need
+ reflexivity) *)
+ Variable C : Type.
+ Variable (cO cI : C) (cadd cmul: C->C->C).
+ Variable (ceqb : C -> C -> bool).
+ Variable phi : C -> R.
+ Variable Smorph : semi_morph rO rI radd rmul req cO cI cadd cmul ceqb phi.
+
+ Lemma SRmorph_Rmorph :
+ ring_morph rO rI radd rmul SRsub SRopp req
+ cO cI cadd cmul cadd (fun x => x) ceqb phi.
+ Proof.
+ case Smorph; intros; constructor; auto.
+ unfold SRopp in |- *; intros.
+ setoid_reflexivity.
+ Qed.
+
+ End SEMI_RING.
+
+ Variable Reqe : ring_eq_ext radd rmul ropp req.
+ Add Morphism radd : radd_ext2. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul : rmul_ext2. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp : ropp_ext2. exact (Ropp_ext Reqe). Qed.
+
+ Section RING.
+ Variable Rth : ring_theory 0 1 radd rmul rsub ropp req.
+
+ (** Rings are almost rings*)
+ Lemma Rmul_0_l : forall x, 0 * x == 0.
+ Proof.
+ intro x; setoid_replace (0*x) with ((0+1)*x + -x).
+ rewrite (Radd_0_l Rth); rewrite (Rmul_1_l Rth).
+ rewrite (Ropp_def Rth);sreflexivity.
+
+ rewrite (Rdistr_l Rth);rewrite (Rmul_1_l Rth).
+ rewrite <- (Radd_assoc Rth); rewrite (Ropp_def Rth).
+ rewrite (Radd_comm Rth); rewrite (Radd_0_l Rth);sreflexivity.
+ Qed.
+
+ Lemma Ropp_mul_l : forall x y, -(x * y) == -x * y.
+ Proof.
+ intros x y;rewrite <-(Radd_0_l Rth (- x * y)).
+ rewrite (Radd_comm Rth).
+ rewrite <-(Ropp_def Rth (x*y)).
+ rewrite (Radd_assoc Rth).
+ rewrite <- (Rdistr_l Rth).
+ rewrite (Rth.(Radd_comm) (-x));rewrite (Ropp_def Rth).
+ rewrite Rmul_0_l;rewrite (Radd_0_l Rth);sreflexivity.
+ Qed.
+
+ Lemma Ropp_add : forall x y, -(x + y) == -x + -y.
+ Proof.
+ intros x y;rewrite <- ((Radd_0_l Rth) (-(x+y))).
+ rewrite <- ((Ropp_def Rth) x).
+ rewrite <- ((Radd_0_l Rth) (x + - x + - (x + y))).
+ rewrite <- ((Ropp_def Rth) y).
+ rewrite ((Radd_comm Rth) x).
+ rewrite ((Radd_comm Rth) y).
+ rewrite <- ((Radd_assoc Rth) (-y)).
+ rewrite <- ((Radd_assoc Rth) (- x)).
+ rewrite ((Radd_assoc Rth) y).
+ rewrite ((Radd_comm Rth) y).
+ rewrite <- ((Radd_assoc Rth) (- x)).
+ rewrite ((Radd_assoc Rth) y).
+ rewrite ((Radd_comm Rth) y);rewrite (Ropp_def Rth).
+ rewrite ((Radd_comm Rth) (-x) 0);rewrite (Radd_0_l Rth).
+ apply (Radd_comm Rth).
+ Qed.
+
+ Lemma Ropp_opp : forall x, - -x == x.
+ Proof.
+ intros x; rewrite <- (Radd_0_l Rth (- -x)).
+ rewrite <- (Ropp_def Rth x).
+ rewrite <- (Radd_assoc Rth); rewrite (Ropp_def Rth).
+ rewrite ((Radd_comm Rth) x);apply (Radd_0_l Rth).
+ Qed.
+
+ Lemma Rth_ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
+ Proof
+ (mk_art 0 1 radd rmul rsub ropp req (Radd_0_l Rth) (Radd_comm Rth) (Radd_assoc Rth)
+ (Rmul_1_l Rth) Rmul_0_l (Rmul_comm Rth) (Rmul_assoc Rth) (Rdistr_l Rth)
+ Ropp_mul_l Ropp_add (Rsub_def Rth)).
+
+ (** Every semi morphism between two rings is a morphism*)
+ Variable C : Type.
+ Variable (cO cI : C) (cadd cmul csub: C->C->C) (copp : C -> C).
+ Variable (ceq : C -> C -> Prop) (ceqb : C -> C -> bool).
+ Variable phi : C -> R.
+ 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).
+ Variable Csth : Setoid_Theory C ceq.
+ Variable Ceqe : ring_eq_ext cadd cmul copp ceq.
+ Add Setoid C ceq Csth as C_setoid.
+ Add Morphism cadd : cadd_ext. exact (Radd_ext Ceqe). Qed.
+ Add Morphism cmul : cmul_ext. exact (Rmul_ext Ceqe). Qed.
+ Add Morphism copp : copp_ext. exact (Ropp_ext Ceqe). Qed.
+ Variable Cth : ring_theory cO cI cadd cmul csub copp ceq.
+ Variable Smorph : semi_morph 0 1 radd rmul req cO cI cadd cmul ceqb phi.
+ Variable phi_ext : forall x y, ceq x y -> [x] == [y].
+ Add Morphism phi : phi_ext1. exact phi_ext. Qed.
+ Lemma Smorph_opp : forall x, [-!x] == -[x].
+ Proof.
+ intros x;rewrite <- (Rth.(Radd_0_l) [-!x]).
+ rewrite <- ((Ropp_def Rth) [x]).
+ rewrite ((Radd_comm Rth) [x]).
+ rewrite <- (Radd_assoc Rth).
+ rewrite <- (Smorph_add Smorph).
+ rewrite (Ropp_def Cth).
+ rewrite (Smorph0 Smorph).
+ rewrite (Radd_comm Rth (-[x])).
+ apply (Radd_0_l Rth);sreflexivity.
+ Qed.
+
+ Lemma Smorph_sub : forall x y, [x -! y] == [x] - [y].
+ Proof.
+ intros x y; rewrite (Rsub_def Cth);rewrite (Rsub_def Rth).
+ rewrite (Smorph_add Smorph);rewrite Smorph_opp;sreflexivity.
+ Qed.
+
+ Lemma Smorph_morph : ring_morph 0 1 radd rmul rsub ropp req
+ cO cI cadd cmul csub copp ceqb phi.
+ Proof
+ (mkmorph 0 1 radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi
+ (Smorph0 Smorph) (Smorph1 Smorph)
+ (Smorph_add Smorph) Smorph_sub (Smorph_mul Smorph) Smorph_opp
+ (Smorph_eq Smorph)).
+
+ End RING.
+
+ (** Useful lemmas on almost ring *)
+ Variable ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
+
+ Lemma ARth_SRth : semi_ring_theory 0 1 radd rmul req.
+Proof.
+elim ARth; intros.
+constructor; trivial.
+Qed.
+
+ Lemma ARsub_ext :
+ forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 - y1 == x2 - y2.
+ Proof.
+ intros.
+ setoid_replace (x1 - y1) with (x1 + -y1).
+ setoid_replace (x2 - y2) with (x2 + -y2).
+ rewrite H;rewrite H0;sreflexivity.
+ apply (ARsub_def ARth).
+ apply (ARsub_def ARth).
+ Qed.
+ Add Morphism rsub : rsub_ext. exact ARsub_ext. Qed.
+
+ Ltac mrewrite :=
+ repeat first
+ [ rewrite (ARadd_0_l ARth)
+ | rewrite <- ((ARadd_comm ARth) 0)
+ | rewrite (ARmul_1_l ARth)
+ | rewrite <- ((ARmul_comm ARth) 1)
+ | rewrite (ARmul_0_l ARth)
+ | rewrite <- ((ARmul_comm ARth) 0)
+ | rewrite (ARdistr_l ARth)
+ | sreflexivity
+ | match goal with
+ | |- context [?z * (?x + ?y)] => rewrite ((ARmul_comm ARth) z (x+y))
+ end].
+
+ Lemma ARadd_0_r : forall x, (x + 0) == x.
+ Proof. intros; mrewrite. Qed.
+
+ Lemma ARmul_1_r : forall x, x * 1 == x.
+ Proof. intros;mrewrite. Qed.
+
+ Lemma ARmul_0_r : forall x, x * 0 == 0.
+ Proof. intros;mrewrite. Qed.
+
+ Lemma ARdistr_r : forall x y z, z * (x + y) == z*x + z*y.
+ Proof.
+ intros;mrewrite.
+ repeat rewrite (ARth.(ARmul_comm) z);sreflexivity.
+ Qed.
+
+ Lemma ARadd_assoc1 : forall x y z, (x + y) + z == (y + z) + x.
+ Proof.
+ intros;rewrite <-(ARth.(ARadd_assoc) x).
+ rewrite (ARth.(ARadd_comm) x);sreflexivity.
+ Qed.
+
+ Lemma ARadd_assoc2 : forall x y z, (y + x) + z == (y + z) + x.
+ Proof.
+ intros; repeat rewrite <- (ARadd_assoc ARth);
+ rewrite ((ARadd_comm ARth) x); sreflexivity.
+ Qed.
+
+ Lemma ARmul_assoc1 : forall x y z, (x * y) * z == (y * z) * x.
+ Proof.
+ intros;rewrite <-((ARmul_assoc ARth) x).
+ rewrite ((ARmul_comm ARth) x);sreflexivity.
+ Qed.
+
+ Lemma ARmul_assoc2 : forall x y z, (y * x) * z == (y * z) * x.
+ Proof.
+ intros; repeat rewrite <- (ARmul_assoc ARth);
+ rewrite ((ARmul_comm ARth) x); sreflexivity.
+ Qed.
+
+ Lemma ARopp_mul_r : forall x y, - (x * y) == x * -y.
+ Proof.
+ intros;rewrite ((ARmul_comm ARth) x y);
+ rewrite (ARopp_mul_l ARth); apply (ARmul_comm ARth).
+ Qed.
+
+ Lemma ARopp_zero : -0 == 0.
+ Proof.
+ rewrite <- (ARmul_0_r 0); rewrite (ARopp_mul_l ARth).
+ repeat rewrite ARmul_0_r; sreflexivity.
+ Qed.
+
+
+
+End ALMOST_RING.
+
+
+Section AddRing.
+
+(* Variable R : Type.
+ Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
+ Variable req : R -> R -> Prop. *)
+
+Inductive ring_kind : Type :=
+| Abstract
+| Computational
+ (R:Type)
+ (req : R -> R -> Prop)
+ (reqb : R -> R -> bool)
+ (_ : forall x y, (reqb x y) = true -> req x y)
+| Morphism
+ (R : Type)
+ (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R)
+ (req : R -> R -> Prop)
+ (C : Type)
+ (cO cI : C) (cadd cmul csub : C->C->C) (copp : C->C)
+ (ceqb : C->C->bool)
+ phi
+ (_ : ring_morph rO rI radd rmul rsub ropp req
+ cO cI cadd cmul csub copp ceqb phi).
+
+
+End AddRing.
+
+
+(** Some simplification tactics*)
+Ltac gen_reflexivity Rsth := apply (Seq_refl _ _ Rsth).
+
+Ltac gen_srewrite Rsth Reqe ARth :=
+ repeat first
+ [ gen_reflexivity Rsth
+ | progress rewrite (ARopp_zero Rsth Reqe ARth)
+ | rewrite (ARadd_0_l ARth)
+ | rewrite (ARadd_0_r Rsth ARth)
+ | rewrite (ARmul_1_l ARth)
+ | rewrite (ARmul_1_r Rsth ARth)
+ | rewrite (ARmul_0_l ARth)
+ | rewrite (ARmul_0_r Rsth ARth)
+ | rewrite (ARdistr_l ARth)
+ | rewrite (ARdistr_r Rsth Reqe ARth)
+ | rewrite (ARadd_assoc ARth)
+ | rewrite (ARmul_assoc ARth)
+ | progress rewrite (ARopp_add ARth)
+ | progress rewrite (ARsub_def ARth)
+ | progress rewrite <- (ARopp_mul_l ARth)
+ | progress rewrite <- (ARopp_mul_r Rsth Reqe ARth) ].
+
+Ltac gen_add_push add Rsth Reqe ARth x :=
+ repeat (match goal with
+ | |- context [add (add ?y x) ?z] =>
+ progress rewrite (ARadd_assoc2 Rsth Reqe ARth x y z)
+ | |- context [add (add x ?y) ?z] =>
+ progress rewrite (ARadd_assoc1 Rsth ARth x y z)
+ end).
+
+Ltac gen_mul_push mul Rsth Reqe ARth x :=
+ repeat (match goal with
+ | |- context [mul (mul ?y x) ?z] =>
+ progress rewrite (ARmul_assoc2 Rsth Reqe ARth x y z)
+ | |- context [mul (mul x ?y) ?z] =>
+ progress rewrite (ARmul_assoc1 Rsth ARth x y z)
+ end).
+
diff --git a/plugins/setoid_ring/ZArithRing.v b/plugins/setoid_ring/ZArithRing.v
new file mode 100644
index 00000000..4cb5a05a
--- /dev/null
+++ b/plugins/setoid_ring/ZArithRing.v
@@ -0,0 +1,60 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Export Ring.
+Require Import ZArith_base.
+Require Import Zpow_def.
+
+Import InitialRing.
+
+Set Implicit Arguments.
+
+Ltac Zcst t :=
+ match isZcst t with
+ true => t
+ | _ => constr:NotConstant
+ end.
+
+Ltac isZpow_coef t :=
+ match t with
+ | Zpos ?p => isPcst p
+ | Z0 => constr:true
+ | _ => constr:false
+ end.
+
+Definition N_of_Z x :=
+ match x with
+ | Zpos p => Npos p
+ | _ => N0
+ end.
+
+Ltac Zpow_tac t :=
+ match isZpow_coef t with
+ | true => constr:(N_of_Z t)
+ | _ => constr:NotConstant
+ end.
+
+Ltac Zpower_neg :=
+ repeat match goal with
+ | [|- ?G] =>
+ match G with
+ | context c [Zpower _ (Zneg _)] =>
+ let t := context c [Z0] in
+ change t
+ end
+ end.
+
+Add Ring Zr : Zth
+ (decidable Zeq_bool_eq, constants [Zcst], preprocess [Zpower_neg;unfold Zsucc],
+ power_tac Zpower_theory [Zpow_tac],
+ (* The two following option are not needed, it is the default chose when the set of
+ coefficiant is usual ring Z *)
+ div (InitialRing.Ztriv_div_th (@Eqsth Z) (@IDphi Z)),
+ sign get_signZ_th).
+
+
diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml4
new file mode 100644
index 00000000..535dbdbd
--- /dev/null
+++ b/plugins/setoid_ring/newring.ml4
@@ -0,0 +1,1164 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i camlp4deps: "parsing/grammar.cma" i*)
+
+(*i $Id$ i*)
+
+open Pp
+open Util
+open Names
+open Term
+open Closure
+open Environ
+open Libnames
+open Tactics
+open Rawterm
+open Termops
+open Tacticals
+open Tacexpr
+open Pcoq
+open Tactic
+open Constr
+open Proof_type
+open Coqlib
+open Tacmach
+open Mod_subst
+open Tacinterp
+open Libobject
+open Printer
+open Declare
+open Decl_kinds
+open Entries
+
+(****************************************************************************)
+(* controlled reduction *)
+
+let mark_arg i c = mkEvar(i,[|c|])
+let unmark_arg f c =
+ match destEvar c with
+ | (i,[|c|]) -> f i c
+ | _ -> assert false
+
+type protect_flag = Eval|Prot|Rec
+
+let tag_arg tag_rec map subs i c =
+ match map i with
+ Eval -> mk_clos subs c
+ | Prot -> mk_atom c
+ | Rec -> if i = -1 then mk_clos subs c else tag_rec c
+
+let rec mk_clos_but f_map subs t =
+ match f_map t with
+ | Some map -> tag_arg (mk_clos_but f_map subs) map subs (-1) t
+ | None ->
+ (match kind_of_term t with
+ App(f,args) -> mk_clos_app_but f_map subs f args 0
+ | Prod _ -> mk_clos_deep (mk_clos_but f_map) subs t
+ | _ -> mk_atom t)
+
+and mk_clos_app_but f_map subs f args n =
+ if n >= Array.length args then mk_atom(mkApp(f, args))
+ else
+ let fargs, args' = array_chop n args in
+ let f' = mkApp(f,fargs) in
+ match f_map f' with
+ Some map ->
+ mk_clos_deep
+ (fun s' -> unmark_arg (tag_arg (mk_clos_but f_map s') map s'))
+ subs
+ (mkApp (mark_arg (-1) f', Array.mapi mark_arg args'))
+ | None -> mk_clos_app_but f_map subs f args (n+1)
+
+
+let interp_map l c =
+ try
+ let (im,am) = List.assoc c l in
+ Some(fun i ->
+ if List.mem i im then Eval
+ else if List.mem i am then Prot
+ else if i = -1 then Eval
+ else Rec)
+ with Not_found -> None
+
+let interp_map l t =
+ try Some(List.assoc t l) with Not_found -> None
+
+let protect_maps = ref Stringmap.empty
+let add_map s m = protect_maps := Stringmap.add s m !protect_maps
+let lookup_map map =
+ try Stringmap.find map !protect_maps
+ with Not_found ->
+ errorlabstrm"lookup_map"(str"map "++qs map++str"not found")
+
+let protect_red map env sigma c =
+ kl (create_clos_infos betadeltaiota env)
+ (mk_clos_but (lookup_map map c) (Esubst.ESID 0) c);;
+
+let protect_tac map =
+ Tactics.reduct_option (protect_red map,DEFAULTcast) None ;;
+
+let protect_tac_in map id =
+ Tactics.reduct_option (protect_red map,DEFAULTcast) (Some(id,InHyp));;
+
+
+TACTIC EXTEND protect_fv
+ [ "protect_fv" string(map) "in" ident(id) ] ->
+ [ protect_tac_in map id ]
+| [ "protect_fv" string(map) ] ->
+ [ protect_tac map ]
+END;;
+
+(****************************************************************************)
+
+let closed_term t l =
+ let l = List.map constr_of_global l in
+ let cs = List.fold_right Quote.ConstrSet.add l Quote.ConstrSet.empty in
+ if Quote.closed_under cs t then tclIDTAC else tclFAIL 0 (mt())
+;;
+
+TACTIC EXTEND closed_term
+ [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] ->
+ [ closed_term t l ]
+END
+;;
+
+TACTIC EXTEND echo
+| [ "echo" constr(t) ] ->
+ [ Pp.msg (Termops.print_constr t); Tacinterp.eval_tactic (TacId []) ]
+END;;
+
+(*
+let closed_term_ast l =
+ TacFun([Some(id_of_string"t")],
+ TacAtom(dummy_loc,TacExtend(dummy_loc,"closed_term",
+ [Genarg.in_gen Genarg.wit_constr (mkVar(id_of_string"t"));
+ Genarg.in_gen (Genarg.wit_list1 Genarg.wit_ref) l])))
+*)
+let closed_term_ast l =
+ let l = List.map (fun gr -> ArgArg(dummy_loc,gr)) l in
+ TacFun([Some(id_of_string"t")],
+ TacAtom(dummy_loc,TacExtend(dummy_loc,"closed_term",
+ [Genarg.in_gen Genarg.globwit_constr (RVar(dummy_loc,id_of_string"t"),None);
+ Genarg.in_gen (Genarg.wit_list1 Genarg.globwit_ref) l])))
+(*
+let _ = add_tacdef false ((dummy_loc,id_of_string"ring_closed_term"
+*)
+
+(****************************************************************************)
+
+let ic c =
+ let env = Global.env() and sigma = Evd.empty in
+ Constrintern.interp_constr sigma env c
+
+let ty c = Typing.type_of (Global.env()) Evd.empty c
+
+let decl_constant na c =
+ mkConst(declare_constant (id_of_string na) (DefinitionEntry
+ { const_entry_body = c;
+ const_entry_type = None;
+ const_entry_opaque = true;
+ const_entry_boxed = true},
+ IsProof Lemma))
+
+(* Calling a global tactic *)
+let ltac_call tac (args:glob_tactic_arg list) =
+ TacArg(TacCall(dummy_loc, ArgArg(dummy_loc, Lazy.force tac),args))
+
+(* Calling a locally bound tactic *)
+let ltac_lcall tac args =
+ TacArg(TacCall(dummy_loc, ArgVar(dummy_loc, id_of_string tac),args))
+
+let ltac_letin (x, e1) e2 =
+ TacLetIn(false,[(dummy_loc,id_of_string x),e1],e2)
+
+let ltac_apply (f:glob_tactic_expr) (args:glob_tactic_arg list) =
+ Tacinterp.eval_tactic
+ (ltac_letin ("F", Tacexp f) (ltac_lcall "F" args))
+
+let ltac_record flds =
+ TacFun([Some(id_of_string"proj")], ltac_lcall "proj" flds)
+
+
+let carg c = TacDynamic(dummy_loc,Pretyping.constr_in c)
+
+let dummy_goal env =
+ {Evd.it = Evd.make_evar (named_context_val env) mkProp;
+ Evd.sigma = Evd.empty}
+
+let exec_tactic env n f args =
+ let lid = list_tabulate(fun i -> id_of_string("x"^string_of_int i)) n in
+ let res = ref [||] in
+ let get_res ist =
+ let l = List.map (fun id -> List.assoc id ist.lfun) lid in
+ res := Array.of_list l;
+ TacId[] in
+ let getter =
+ Tacexp(TacFun(List.map(fun id -> Some id) lid,
+ glob_tactic(tacticIn get_res))) in
+ let _ =
+ Tacinterp.eval_tactic(ltac_call f (args@[getter])) (dummy_goal env) in
+ !res
+
+let constr_of = function
+ | VConstr ([],c) -> c
+ | _ -> failwith "Ring.exec_tactic: anomaly"
+
+let stdlib_modules =
+ [["Coq";"Setoids";"Setoid"];
+ ["Coq";"Lists";"List"];
+ ["Coq";"Init";"Datatypes"];
+ ["Coq";"Init";"Logic"];
+ ]
+
+let coq_constant c =
+ lazy (Coqlib.gen_constant_in_modules "Ring" stdlib_modules c)
+
+let coq_mk_Setoid = coq_constant "Build_Setoid_Theory"
+let coq_cons = coq_constant "cons"
+let coq_nil = coq_constant "nil"
+let coq_None = coq_constant "None"
+let coq_Some = coq_constant "Some"
+let coq_eq = coq_constant "eq"
+
+let lapp f args = mkApp(Lazy.force f,args)
+
+let dest_rel0 t =
+ match kind_of_term t with
+ | App(f,args) when Array.length args >= 2 ->
+ let rel = mkApp(f,Array.sub args 0 (Array.length args - 2)) in
+ if closed0 rel then
+ (rel,args.(Array.length args - 2),args.(Array.length args - 1))
+ else error "ring: cannot find relation (not closed)"
+ | _ -> error "ring: cannot find relation"
+
+let rec dest_rel t =
+ match kind_of_term t with
+ | Prod(_,_,c) -> dest_rel c
+ | _ -> dest_rel0 t
+
+(****************************************************************************)
+(* Library linking *)
+
+let plugin_dir = "setoid_ring"
+
+let cdir = ["Coq";plugin_dir]
+let plugin_modules =
+ List.map (fun d -> cdir@d)
+ [["Ring_theory"];["Ring_polynom"]; ["Ring_tac"];["InitialRing"];
+ ["Field_tac"]; ["Field_theory"]
+ ]
+
+let my_constant c =
+ lazy (Coqlib.gen_constant_in_modules "Ring" plugin_modules c)
+
+let new_ring_path =
+ make_dirpath (List.map id_of_string ["Ring_tac";plugin_dir;"Coq"])
+let ltac s =
+ lazy(make_kn (MPfile new_ring_path) (make_dirpath []) (mk_label s))
+let znew_ring_path =
+ make_dirpath (List.map id_of_string ["InitialRing";plugin_dir;"Coq"])
+let zltac s =
+ lazy(make_kn (MPfile znew_ring_path) (make_dirpath []) (mk_label s))
+
+let mk_cst l s = lazy (Coqlib.gen_constant "newring" l s);;
+let pol_cst s = mk_cst [plugin_dir;"Ring_polynom"] s ;;
+
+(* Ring theory *)
+
+(* almost_ring defs *)
+let coq_almost_ring_theory = my_constant "almost_ring_theory"
+
+(* setoid and morphism utilities *)
+let coq_eq_setoid = my_constant "Eqsth"
+let coq_eq_morph = my_constant "Eq_ext"
+let coq_eq_smorph = my_constant "Eq_s_ext"
+
+(* ring -> almost_ring utilities *)
+let coq_ring_theory = my_constant "ring_theory"
+let coq_mk_reqe = my_constant "mk_reqe"
+
+(* semi_ring -> almost_ring utilities *)
+let coq_semi_ring_theory = my_constant "semi_ring_theory"
+let coq_mk_seqe = my_constant "mk_seqe"
+
+let ltac_inv_morph_gen = zltac"inv_gen_phi"
+let ltac_inv_morphZ = zltac"inv_gen_phiZ"
+let ltac_inv_morphN = zltac"inv_gen_phiN"
+let ltac_inv_morphNword = zltac"inv_gen_phiNword"
+let coq_abstract = my_constant"Abstract"
+let coq_comp = my_constant"Computational"
+let coq_morph = my_constant"Morphism"
+
+(* morphism *)
+let coq_ring_morph = my_constant "ring_morph"
+let coq_semi_morph = my_constant "semi_morph"
+
+(* power function *)
+let ltac_inv_morph_nothing = zltac"inv_morph_nothing"
+let coq_pow_N_pow_N = my_constant "pow_N_pow_N"
+
+(* hypothesis *)
+let coq_mkhypo = my_constant "mkhypo"
+let coq_hypo = my_constant "hypo"
+
+(* Equality: do not evaluate but make recursive call on both sides *)
+let map_with_eq arg_map c =
+ let (req,_,_) = dest_rel c in
+ interp_map
+ ((req,(function -1->Prot|_->Rec))::
+ List.map (fun (c,map) -> (Lazy.force c,map)) arg_map)
+
+let _ = add_map "ring"
+ (map_with_eq
+ [coq_cons,(function -1->Eval|2->Rec|_->Prot);
+ coq_nil, (function -1->Eval|_ -> Prot);
+ (* Pphi_dev: evaluate polynomial and coef operations, protect
+ ring operations and make recursive call on the var map *)
+ pol_cst "Pphi_dev", (function -1|8|9|10|11|12|14->Eval|13->Rec|_->Prot);
+ pol_cst "Pphi_pow",
+ (function -1|8|9|10|11|13|15|17->Eval|16->Rec|_->Prot);
+ (* PEeval: evaluate morphism and polynomial, protect ring
+ operations and make recursive call on the var map *)
+ pol_cst "PEeval", (function -1|7|9|12->Eval|11->Rec|_->Prot)])
+
+(****************************************************************************)
+(* Ring database *)
+
+type ring_info =
+ { ring_carrier : types;
+ ring_req : constr;
+ ring_setoid : constr;
+ ring_ext : constr;
+ ring_morph : constr;
+ ring_th : constr;
+ ring_cst_tac : glob_tactic_expr;
+ ring_pow_tac : glob_tactic_expr;
+ ring_lemma1 : constr;
+ ring_lemma2 : constr;
+ ring_pre_tac : glob_tactic_expr;
+ ring_post_tac : glob_tactic_expr }
+
+module Cmap = Map.Make(struct type t = constr let compare = compare end)
+
+let from_carrier = ref Cmap.empty
+let from_relation = ref Cmap.empty
+let from_name = ref Spmap.empty
+
+let ring_for_carrier r = Cmap.find r !from_carrier
+let ring_for_relation rel = Cmap.find rel !from_relation
+
+
+let find_ring_structure env sigma l =
+ match l with
+ | t::cl' ->
+ let ty = Retyping.get_type_of env sigma t in
+ let check c =
+ let ty' = Retyping.get_type_of env sigma c in
+ if not (Reductionops.is_conv env sigma ty ty') then
+ errorlabstrm "ring"
+ (str"arguments of ring_simplify do not have all the same type")
+ in
+ List.iter check cl';
+ (try ring_for_carrier ty
+ with Not_found ->
+ errorlabstrm "ring"
+ (str"cannot find a declared ring structure over"++
+ spc()++str"\""++pr_constr ty++str"\""))
+ | [] -> assert false
+(*
+ let (req,_,_) = dest_rel cl in
+ (try ring_for_relation req
+ with Not_found ->
+ errorlabstrm "ring"
+ (str"cannot find a declared ring structure for equality"++
+ spc()++str"\""++pr_constr req++str"\"")) *)
+
+let _ =
+ Summary.declare_summary "tactic-new-ring-table"
+ { Summary.freeze_function =
+ (fun () -> !from_carrier,!from_relation,!from_name);
+ Summary.unfreeze_function =
+ (fun (ct,rt,nt) ->
+ from_carrier := ct; from_relation := rt; from_name := nt);
+ Summary.init_function =
+ (fun () ->
+ from_carrier := Cmap.empty; from_relation := Cmap.empty;
+ from_name := Spmap.empty) }
+
+let add_entry (sp,_kn) e =
+(* let _ = ty e.ring_lemma1 in
+ let _ = ty e.ring_lemma2 in
+*)
+ from_carrier := Cmap.add e.ring_carrier e !from_carrier;
+ from_relation := Cmap.add e.ring_req e !from_relation;
+ from_name := Spmap.add sp e !from_name
+
+
+let subst_th (subst,th) =
+ let c' = subst_mps subst th.ring_carrier in
+ let eq' = subst_mps subst th.ring_req in
+ let set' = subst_mps subst th.ring_setoid in
+ let ext' = subst_mps subst th.ring_ext in
+ let morph' = subst_mps subst th.ring_morph in
+ let th' = subst_mps subst th.ring_th in
+ let thm1' = subst_mps subst th.ring_lemma1 in
+ let thm2' = subst_mps subst th.ring_lemma2 in
+ let tac'= subst_tactic subst th.ring_cst_tac in
+ let pow_tac'= subst_tactic subst th.ring_pow_tac in
+ let pretac'= subst_tactic subst th.ring_pre_tac in
+ let posttac'= subst_tactic subst th.ring_post_tac in
+ if c' == th.ring_carrier &&
+ eq' == th.ring_req &&
+ set' = th.ring_setoid &&
+ ext' == th.ring_ext &&
+ morph' == th.ring_morph &&
+ th' == th.ring_th &&
+ thm1' == th.ring_lemma1 &&
+ thm2' == th.ring_lemma2 &&
+ tac' == th.ring_cst_tac &&
+ pow_tac' == th.ring_pow_tac &&
+ pretac' == th.ring_pre_tac &&
+ posttac' == th.ring_post_tac then th
+ else
+ { ring_carrier = c';
+ ring_req = eq';
+ ring_setoid = set';
+ ring_ext = ext';
+ ring_morph = morph';
+ ring_th = th';
+ ring_cst_tac = tac';
+ ring_pow_tac = pow_tac';
+ ring_lemma1 = thm1';
+ ring_lemma2 = thm2';
+ ring_pre_tac = pretac';
+ ring_post_tac = posttac' }
+
+
+let (theory_to_obj, obj_to_theory) =
+ let cache_th (name,th) = add_entry name th in
+ declare_object
+ {(default_object "tactic-new-ring-theory") with
+ open_function = (fun i o -> if i=1 then cache_th o);
+ cache_function = cache_th;
+ subst_function = subst_th;
+ classify_function = (fun x -> Substitute x)}
+
+
+let setoid_of_relation env a r =
+ let evm = Evd.empty in
+ try
+ lapp coq_mk_Setoid
+ [|a ; r ;
+ Rewrite.get_reflexive_proof env evm a r ;
+ Rewrite.get_symmetric_proof env evm a r ;
+ Rewrite.get_transitive_proof env evm a r |]
+ with Not_found ->
+ error "cannot find setoid relation"
+
+let op_morph r add mul opp req m1 m2 m3 =
+ lapp coq_mk_reqe [| r; add; mul; opp; req; m1; m2; m3 |]
+
+let op_smorph r add mul req m1 m2 =
+ lapp coq_mk_seqe [| r; add; mul; req; m1; m2 |]
+
+(* let default_ring_equality (r,add,mul,opp,req) = *)
+(* let is_setoid = function *)
+(* {rel_refl=Some _; rel_sym=Some _;rel_trans=Some _;rel_aeq=rel} -> *)
+(* eq_constr req rel (\* Qu: use conversion ? *\) *)
+(* | _ -> false in *)
+(* match default_relation_for_carrier ~filter:is_setoid r with *)
+(* Leibniz _ -> *)
+(* let setoid = lapp coq_eq_setoid [|r|] in *)
+(* let op_morph = *)
+(* match opp with *)
+(* Some opp -> lapp coq_eq_morph [|r;add;mul;opp|] *)
+(* | None -> lapp coq_eq_smorph [|r;add;mul|] in *)
+(* (setoid,op_morph) *)
+(* | Relation rel -> *)
+(* let setoid = setoid_of_relation rel in *)
+(* let is_endomorphism = function *)
+(* { args=args } -> List.for_all *)
+(* (function (var,Relation rel) -> *)
+(* var=None && eq_constr req rel *)
+(* | _ -> false) args in *)
+(* let add_m = *)
+(* try default_morphism ~filter:is_endomorphism add *)
+(* with Not_found -> *)
+(* error "ring addition should be declared as a morphism" in *)
+(* let mul_m = *)
+(* try default_morphism ~filter:is_endomorphism mul *)
+(* with Not_found -> *)
+(* error "ring multiplication should be declared as a morphism" in *)
+(* let op_morph = *)
+(* match opp with *)
+(* | Some opp -> *)
+(* (let opp_m = *)
+(* try default_morphism ~filter:is_endomorphism opp *)
+(* with Not_found -> *)
+(* error "ring opposite should be declared as a morphism" in *)
+(* let op_morph = *)
+(* op_morph r add mul opp req add_m.lem mul_m.lem opp_m.lem in *)
+(* msgnl *)
+(* (str"Using setoid \""++pr_constr rel.rel_aeq++str"\""++spc()++ *)
+(* str"and morphisms \""++pr_constr add_m.morphism_theory++ *)
+(* str"\","++spc()++ str"\""++pr_constr mul_m.morphism_theory++ *)
+(* str"\""++spc()++str"and \""++pr_constr opp_m.morphism_theory++ *)
+(* str"\""); *)
+(* op_morph) *)
+(* | None -> *)
+(* (msgnl *)
+(* (str"Using setoid \""++pr_constr rel.rel_aeq++str"\"" ++ spc() ++ *)
+(* str"and morphisms \""++pr_constr add_m.morphism_theory++ *)
+(* str"\""++spc()++str"and \""++ *)
+(* pr_constr mul_m.morphism_theory++str"\""); *)
+(* op_smorph r add mul req add_m.lem mul_m.lem) in *)
+(* (setoid,op_morph) *)
+
+let ring_equality (r,add,mul,opp,req) =
+ match kind_of_term req with
+ | App (f, [| _ |]) when eq_constr f (Lazy.force coq_eq) ->
+ let setoid = lapp coq_eq_setoid [|r|] in
+ let op_morph =
+ match opp with
+ Some opp -> lapp coq_eq_morph [|r;add;mul;opp|]
+ | None -> lapp coq_eq_smorph [|r;add;mul|] in
+ (setoid,op_morph)
+ | _ ->
+ let setoid = setoid_of_relation (Global.env ()) r req in
+ let signature = [Some (r,req);Some (r,req)],Some(r,req) in
+ let add_m, add_m_lem =
+ try Rewrite.default_morphism signature add
+ with Not_found ->
+ error "ring addition should be declared as a morphism" in
+ let mul_m, mul_m_lem =
+ try Rewrite.default_morphism signature mul
+ with Not_found ->
+ error "ring multiplication should be declared as a morphism" in
+ let op_morph =
+ match opp with
+ | Some opp ->
+ (let opp_m,opp_m_lem =
+ try Rewrite.default_morphism ([Some(r,req)],Some(r,req)) opp
+ with Not_found ->
+ error "ring opposite should be declared as a morphism" in
+ let op_morph =
+ op_morph r add mul opp req add_m_lem mul_m_lem opp_m_lem in
+ Flags.if_verbose
+ msgnl
+ (str"Using setoid \""++pr_constr req++str"\""++spc()++
+ str"and morphisms \""++pr_constr add_m_lem ++
+ str"\","++spc()++ str"\""++pr_constr mul_m_lem++
+ str"\""++spc()++str"and \""++pr_constr opp_m_lem++
+ str"\"");
+ op_morph)
+ | None ->
+ (Flags.if_verbose
+ msgnl
+ (str"Using setoid \""++pr_constr req ++str"\"" ++ spc() ++
+ str"and morphisms \""++pr_constr add_m_lem ++
+ str"\""++spc()++str"and \""++
+ pr_constr mul_m_lem++str"\"");
+ op_smorph r add mul req add_m_lem mul_m_lem) in
+ (setoid,op_morph)
+
+let build_setoid_params r add mul opp req eqth =
+ match eqth with
+ Some th -> th
+ | None -> ring_equality (r,add,mul,opp,req)
+
+let dest_ring env sigma th_spec =
+ let th_typ = Retyping.get_type_of env sigma th_spec in
+ match kind_of_term th_typ with
+ App(f,[|r;zero;one;add;mul;sub;opp;req|])
+ when f = Lazy.force coq_almost_ring_theory ->
+ (None,r,zero,one,add,mul,Some sub,Some opp,req)
+ | App(f,[|r;zero;one;add;mul;req|])
+ when f = Lazy.force coq_semi_ring_theory ->
+ (Some true,r,zero,one,add,mul,None,None,req)
+ | App(f,[|r;zero;one;add;mul;sub;opp;req|])
+ when f = Lazy.force coq_ring_theory ->
+ (Some false,r,zero,one,add,mul,Some sub,Some opp,req)
+ | _ -> error "bad ring structure"
+
+
+let dest_morph env sigma m_spec =
+ let m_typ = Retyping.get_type_of env sigma m_spec in
+ match kind_of_term m_typ with
+ App(f,[|r;zero;one;add;mul;sub;opp;req;
+ c;czero;cone;cadd;cmul;csub;copp;ceqb;phi|])
+ when f = Lazy.force coq_ring_morph ->
+ (c,czero,cone,cadd,cmul,Some csub,Some copp,ceqb,phi)
+ | App(f,[|r;zero;one;add;mul;req;c;czero;cone;cadd;cmul;ceqb;phi|])
+ when f = Lazy.force coq_semi_morph ->
+ (c,czero,cone,cadd,cmul,None,None,ceqb,phi)
+ | _ -> error "bad morphism structure"
+
+
+type coeff_spec =
+ Computational of constr (* equality test *)
+ | Abstract (* coeffs = Z *)
+ | Morphism of constr (* general morphism *)
+
+
+let reflect_coeff rkind =
+ (* We build an ill-typed terms on purpose... *)
+ match rkind with
+ Abstract -> Lazy.force coq_abstract
+ | Computational c -> lapp coq_comp [|c|]
+ | Morphism m -> lapp coq_morph [|m|]
+
+type cst_tac_spec =
+ CstTac of raw_tactic_expr
+ | Closed of reference list
+
+let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac =
+ match cst_tac with
+ Some (CstTac t) -> Tacinterp.glob_tactic t
+ | Some (Closed lc) ->
+ closed_term_ast (List.map Smartlocate.global_with_alias lc)
+ | None ->
+ (match rk, opp, kind with
+ Abstract, None, _ ->
+ let t = ArgArg(dummy_loc,Lazy.force ltac_inv_morphN) in
+ TacArg(TacCall(dummy_loc,t,List.map carg [zero;one;add;mul]))
+ | Abstract, Some opp, Some _ ->
+ let t = ArgArg(dummy_loc, Lazy.force ltac_inv_morphZ) in
+ TacArg(TacCall(dummy_loc,t,List.map carg [zero;one;add;mul;opp]))
+ | Abstract, Some opp, None ->
+ let t = ArgArg(dummy_loc, Lazy.force ltac_inv_morphNword) in
+ TacArg
+ (TacCall(dummy_loc,t,List.map carg [zero;one;add;mul;opp]))
+ | Computational _,_,_ ->
+ let t = ArgArg(dummy_loc, Lazy.force ltac_inv_morph_gen) in
+ TacArg
+ (TacCall(dummy_loc,t,List.map carg [zero;one;zero;one]))
+ | Morphism mth,_,_ ->
+ let (_,czero,cone,_,_,_,_,_,_) = dest_morph env sigma mth in
+ let t = ArgArg(dummy_loc, Lazy.force ltac_inv_morph_gen) in
+ TacArg
+ (TacCall(dummy_loc,t,List.map carg [zero;one;czero;cone])))
+
+let make_hyp env c =
+ let t = Retyping.get_type_of env Evd.empty c in
+ lapp coq_mkhypo [|t;c|]
+
+let make_hyp_list env lH =
+ let carrier = Lazy.force coq_hypo in
+ List.fold_right
+ (fun c l -> lapp coq_cons [|carrier; (make_hyp env c); l|]) lH
+ (lapp coq_nil [|carrier|])
+
+let interp_power env pow =
+ let carrier = Lazy.force coq_hypo in
+ match pow with
+ | None ->
+ let t = ArgArg(dummy_loc, Lazy.force ltac_inv_morph_nothing) in
+ (TacArg(TacCall(dummy_loc,t,[])), lapp coq_None [|carrier|])
+ | Some (tac, spec) ->
+ let tac =
+ match tac with
+ | CstTac t -> Tacinterp.glob_tactic t
+ | Closed lc ->
+ closed_term_ast (List.map Smartlocate.global_with_alias lc) in
+ let spec = make_hyp env (ic spec) in
+ (tac, lapp coq_Some [|carrier; spec|])
+
+let interp_sign env sign =
+ let carrier = Lazy.force coq_hypo in
+ match sign with
+ | None -> lapp coq_None [|carrier|]
+ | Some spec ->
+ let spec = make_hyp env (ic spec) in
+ lapp coq_Some [|carrier;spec|]
+ (* Same remark on ill-typed terms ... *)
+
+let interp_div env div =
+ let carrier = Lazy.force coq_hypo in
+ match div with
+ | None -> lapp coq_None [|carrier|]
+ | Some spec ->
+ let spec = make_hyp env (ic spec) in
+ lapp coq_Some [|carrier;spec|]
+ (* Same remark on ill-typed terms ... *)
+
+let add_theory name rth eqth morphth cst_tac (pre,post) power sign div =
+ check_required_library (cdir@["Ring_base"]);
+ let env = Global.env() in
+ let sigma = Evd.empty in
+ let (kind,r,zero,one,add,mul,sub,opp,req) = dest_ring env sigma rth in
+ let (sth,ext) = build_setoid_params r add mul opp req eqth in
+ let (pow_tac, pspec) = interp_power env power in
+ let sspec = interp_sign env sign in
+ let dspec = interp_div env div in
+ let rk = reflect_coeff morphth in
+ let params =
+ exec_tactic env 5 (zltac "ring_lemmas")
+ (List.map carg[sth;ext;rth;pspec;sspec;dspec;rk]) in
+ let lemma1 = constr_of params.(3) in
+ let lemma2 = constr_of params.(4) in
+
+ let lemma1 = decl_constant (string_of_id name^"_ring_lemma1") lemma1 in
+ let lemma2 = decl_constant (string_of_id name^"_ring_lemma2") lemma2 in
+ let cst_tac =
+ interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in
+ let pretac =
+ match pre with
+ Some t -> Tacinterp.glob_tactic t
+ | _ -> TacId [] in
+ let posttac =
+ match post with
+ Some t -> Tacinterp.glob_tactic t
+ | _ -> TacId [] in
+ let _ =
+ Lib.add_leaf name
+ (theory_to_obj
+ { ring_carrier = r;
+ ring_req = req;
+ ring_setoid = sth;
+ ring_ext = constr_of params.(1);
+ ring_morph = constr_of params.(2);
+ ring_th = constr_of params.(0);
+ ring_cst_tac = cst_tac;
+ ring_pow_tac = pow_tac;
+ ring_lemma1 = lemma1;
+ ring_lemma2 = lemma2;
+ ring_pre_tac = pretac;
+ ring_post_tac = posttac }) in
+ ()
+
+type ring_mod =
+ Ring_kind of coeff_spec
+ | Const_tac of cst_tac_spec
+ | Pre_tac of raw_tactic_expr
+ | Post_tac of raw_tactic_expr
+ | Setoid of Topconstr.constr_expr * Topconstr.constr_expr
+ | Pow_spec of cst_tac_spec * Topconstr.constr_expr
+ (* Syntaxification tactic , correctness lemma *)
+ | Sign_spec of Topconstr.constr_expr
+ | Div_spec of Topconstr.constr_expr
+
+
+VERNAC ARGUMENT EXTEND ring_mod
+ | [ "decidable" constr(eq_test) ] -> [ Ring_kind(Computational (ic eq_test)) ]
+ | [ "abstract" ] -> [ Ring_kind Abstract ]
+ | [ "morphism" constr(morph) ] -> [ Ring_kind(Morphism (ic morph)) ]
+ | [ "constants" "[" tactic(cst_tac) "]" ] -> [ Const_tac(CstTac cst_tac) ]
+ | [ "closed" "[" ne_global_list(l) "]" ] -> [ Const_tac(Closed l) ]
+ | [ "preprocess" "[" tactic(pre) "]" ] -> [ Pre_tac pre ]
+ | [ "postprocess" "[" tactic(post) "]" ] -> [ Post_tac post ]
+ | [ "setoid" constr(sth) constr(ext) ] -> [ Setoid(sth,ext) ]
+ | [ "sign" constr(sign_spec) ] -> [ Sign_spec sign_spec ]
+ | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] ->
+ [ Pow_spec (Closed l, pow_spec) ]
+ | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] ->
+ [ Pow_spec (CstTac cst_tac, pow_spec) ]
+ | [ "div" constr(div_spec) ] -> [ Div_spec div_spec ]
+END
+
+let set_once s r v =
+ if !r = None then r := Some v else error (s^" cannot be set twice")
+
+let process_ring_mods l =
+ let kind = ref None in
+ let set = ref None in
+ let cst_tac = ref None in
+ let pre = ref None in
+ let post = ref None in
+ let sign = ref None in
+ let power = ref None in
+ let div = ref None in
+ List.iter(function
+ Ring_kind k -> set_once "ring kind" kind k
+ | Const_tac t -> set_once "tactic recognizing constants" cst_tac t
+ | Pre_tac t -> set_once "preprocess tactic" pre t
+ | Post_tac t -> set_once "postprocess tactic" post t
+ | Setoid(sth,ext) -> set_once "setoid" set (ic sth,ic ext)
+ | Pow_spec(t,spec) -> set_once "power" power (t,spec)
+ | Sign_spec t -> set_once "sign" sign t
+ | Div_spec t -> set_once "div" div t) l;
+ let k = match !kind with Some k -> k | None -> Abstract in
+ (k, !set, !cst_tac, !pre, !post, !power, !sign, !div)
+
+VERNAC COMMAND EXTEND AddSetoidRing
+ | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods(l) ] ->
+ [ let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in
+ add_theory id (ic t) set k cst (pre,post) power sign div]
+END
+
+(*****************************************************************************)
+(* The tactics consist then only in a lookup in the ring database and
+ call the appropriate ltac. *)
+
+let make_args_list rl t =
+ match rl with
+ | [] -> let (_,t1,t2) = dest_rel0 t in [t1;t2]
+ | _ -> rl
+
+let make_term_list carrier rl =
+ List.fold_right
+ (fun x l -> lapp coq_cons [|carrier;x;l|]) rl
+ (lapp coq_nil [|carrier|])
+
+let ltac_ring_structure e =
+ let req = carg e.ring_req in
+ let sth = carg e.ring_setoid in
+ let ext = carg e.ring_ext in
+ let morph = carg e.ring_morph in
+ let th = carg e.ring_th in
+ let cst_tac = Tacexp e.ring_cst_tac in
+ let pow_tac = Tacexp e.ring_pow_tac in
+ let lemma1 = carg e.ring_lemma1 in
+ let lemma2 = carg e.ring_lemma2 in
+ let pretac = Tacexp(TacFun([None],e.ring_pre_tac)) in
+ let posttac = Tacexp(TacFun([None],e.ring_post_tac)) in
+ [req;sth;ext;morph;th;cst_tac;pow_tac;
+ lemma1;lemma2;pretac;posttac]
+
+let ring_lookup (f:glob_tactic_expr) lH rl t gl =
+ let env = pf_env gl in
+ let sigma = project gl in
+ let rl = make_args_list rl t in
+ let e = find_ring_structure env sigma rl in
+ let rl = carg (make_term_list e.ring_carrier rl) in
+ let lH = carg (make_hyp_list env lH) in
+ let ring = ltac_ring_structure e in
+ ltac_apply f (ring@[lH;rl]) gl
+
+TACTIC EXTEND ring_lookup
+| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] ->
+ [ let (t,lr) = list_sep_last lrt in ring_lookup (fst f) lH lr t]
+END
+
+
+
+(***********************************************************************)
+
+let new_field_path =
+ make_dirpath (List.map id_of_string ["Field_tac";plugin_dir;"Coq"])
+
+let field_ltac s =
+ lazy(make_kn (MPfile new_field_path) (make_dirpath []) (mk_label s))
+
+
+let _ = add_map "field"
+ (map_with_eq
+ [coq_cons,(function -1->Eval|2->Rec|_->Prot);
+ coq_nil, (function -1->Eval|_ -> Prot);
+ (* display_linear: evaluate polynomials and coef operations, protect
+ field operations and make recursive call on the var map *)
+ my_constant "display_linear",
+ (function -1|9|10|11|12|13|15|16->Eval|14->Rec|_->Prot);
+ my_constant "display_pow_linear",
+ (function -1|9|10|11|12|13|14|16|18|19->Eval|17->Rec|_->Prot);
+ (* Pphi_dev: evaluate polynomial and coef operations, protect
+ ring operations and make recursive call on the var map *)
+ pol_cst "Pphi_dev", (function -1|8|9|10|11|12|14->Eval|13->Rec|_->Prot);
+ pol_cst "Pphi_pow",
+ (function -1|8|9|10|11|13|15|17->Eval|16->Rec|_->Prot);
+ (* PEeval: evaluate morphism and polynomial, protect ring
+ operations and make recursive call on the var map *)
+ pol_cst "PEeval", (function -1|7|9|12->Eval|11->Rec|_->Prot);
+ (* FEeval: evaluate morphism, protect field
+ operations and make recursive call on the var map *)
+ my_constant "FEeval", (function -1|8|9|10|11|14->Eval|13->Rec|_->Prot)]);;
+
+let _ = add_map "field_cond"
+ (map_with_eq
+ [coq_cons,(function -1->Eval|2->Rec|_->Prot);
+ coq_nil, (function -1->Eval|_ -> Prot);
+ (* PCond: evaluate morphism and denum list, protect ring
+ operations and make recursive call on the var map *)
+ my_constant "PCond", (function -1|8|10|13->Eval|12->Rec|_->Prot)]);;
+(* (function -1|8|10->Eval|9->Rec|_->Prot)]);;*)
+
+
+let _ = Redexpr.declare_reduction "simpl_field_expr"
+ (protect_red "field")
+
+
+
+let afield_theory = my_constant "almost_field_theory"
+let field_theory = my_constant "field_theory"
+let sfield_theory = my_constant "semi_field_theory"
+let af_ar = my_constant"AF_AR"
+let f_r = my_constant"F_R"
+let sf_sr = my_constant"SF_SR"
+let dest_field env sigma th_spec =
+ let th_typ = Retyping.get_type_of env sigma th_spec in
+ match kind_of_term th_typ with
+ | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|])
+ when f = Lazy.force afield_theory ->
+ let rth = lapp af_ar
+ [|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in
+ (None,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth)
+ | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|])
+ when f = Lazy.force field_theory ->
+ let rth =
+ lapp f_r
+ [|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in
+ (Some false,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth)
+ | App(f,[|r;zero;one;add;mul;div;inv;req|])
+ when f = Lazy.force sfield_theory ->
+ let rth = lapp sf_sr
+ [|r;zero;one;add;mul;div;inv;req;th_spec|] in
+ (Some true,r,zero,one,add,mul,None,None,div,inv,req,rth)
+ | _ -> error "bad field structure"
+
+type field_info =
+ { field_carrier : types;
+ field_req : constr;
+ field_cst_tac : glob_tactic_expr;
+ field_pow_tac : glob_tactic_expr;
+ field_ok : constr;
+ field_simpl_eq_ok : constr;
+ field_simpl_ok : constr;
+ field_simpl_eq_in_ok : constr;
+ field_cond : constr;
+ field_pre_tac : glob_tactic_expr;
+ field_post_tac : glob_tactic_expr }
+
+let field_from_carrier = ref Cmap.empty
+let field_from_relation = ref Cmap.empty
+let field_from_name = ref Spmap.empty
+
+
+let field_for_carrier r = Cmap.find r !field_from_carrier
+let field_for_relation rel = Cmap.find rel !field_from_relation
+
+let find_field_structure env sigma l =
+ check_required_library (cdir@["Field_tac"]);
+ match l with
+ | t::cl' ->
+ let ty = Retyping.get_type_of env sigma t in
+ let check c =
+ let ty' = Retyping.get_type_of env sigma c in
+ if not (Reductionops.is_conv env sigma ty ty') then
+ errorlabstrm "field"
+ (str"arguments of field_simplify do not have all the same type")
+ in
+ List.iter check cl';
+ (try field_for_carrier ty
+ with Not_found ->
+ errorlabstrm "field"
+ (str"cannot find a declared field structure over"++
+ spc()++str"\""++pr_constr ty++str"\""))
+ | [] -> assert false
+(* let (req,_,_) = dest_rel cl in
+ (try field_for_relation req
+ with Not_found ->
+ errorlabstrm "field"
+ (str"cannot find a declared field structure for equality"++
+ spc()++str"\""++pr_constr req++str"\"")) *)
+
+let _ =
+ Summary.declare_summary "tactic-new-field-table"
+ { Summary.freeze_function =
+ (fun () -> !field_from_carrier,!field_from_relation,!field_from_name);
+ Summary.unfreeze_function =
+ (fun (ct,rt,nt) ->
+ field_from_carrier := ct; field_from_relation := rt;
+ field_from_name := nt);
+ Summary.init_function =
+ (fun () ->
+ field_from_carrier := Cmap.empty; field_from_relation := Cmap.empty;
+ field_from_name := Spmap.empty) }
+
+let add_field_entry (sp,_kn) e =
+(*
+ let _ = ty e.field_ok in
+ let _ = ty e.field_simpl_eq_ok in
+ let _ = ty e.field_simpl_ok in
+ let _ = ty e.field_cond in
+*)
+ field_from_carrier := Cmap.add e.field_carrier e !field_from_carrier;
+ field_from_relation := Cmap.add e.field_req e !field_from_relation;
+ field_from_name := Spmap.add sp e !field_from_name
+
+let subst_th (subst,th) =
+ let c' = subst_mps subst th.field_carrier in
+ let eq' = subst_mps subst th.field_req in
+ let thm1' = subst_mps subst th.field_ok in
+ let thm2' = subst_mps subst th.field_simpl_eq_ok in
+ let thm3' = subst_mps subst th.field_simpl_ok in
+ let thm4' = subst_mps subst th.field_simpl_eq_in_ok in
+ let thm5' = subst_mps subst th.field_cond in
+ let tac'= subst_tactic subst th.field_cst_tac in
+ let pow_tac' = subst_tactic subst th.field_pow_tac in
+ let pretac'= subst_tactic subst th.field_pre_tac in
+ let posttac'= subst_tactic subst th.field_post_tac in
+ if c' == th.field_carrier &&
+ eq' == th.field_req &&
+ thm1' == th.field_ok &&
+ thm2' == th.field_simpl_eq_ok &&
+ thm3' == th.field_simpl_ok &&
+ thm4' == th.field_simpl_eq_in_ok &&
+ thm5' == th.field_cond &&
+ tac' == th.field_cst_tac &&
+ pow_tac' == th.field_pow_tac &&
+ pretac' == th.field_pre_tac &&
+ posttac' == th.field_post_tac then th
+ else
+ { field_carrier = c';
+ field_req = eq';
+ field_cst_tac = tac';
+ field_pow_tac = pow_tac';
+ field_ok = thm1';
+ field_simpl_eq_ok = thm2';
+ field_simpl_ok = thm3';
+ field_simpl_eq_in_ok = thm4';
+ field_cond = thm5';
+ field_pre_tac = pretac';
+ field_post_tac = posttac' }
+
+let (ftheory_to_obj, obj_to_ftheory) =
+ let cache_th (name,th) = add_field_entry name th in
+ declare_object
+ {(default_object "tactic-new-field-theory") with
+ open_function = (fun i o -> if i=1 then cache_th o);
+ cache_function = cache_th;
+ subst_function = subst_th;
+ classify_function = (fun x -> Substitute x) }
+
+let field_equality r inv req =
+ match kind_of_term req with
+ | App (f, [| _ |]) when eq_constr f (Lazy.force coq_eq) ->
+ mkApp((Coqlib.build_coq_eq_data()).congr,[|r;r;inv|])
+ | _ ->
+ let _setoid = setoid_of_relation (Global.env ()) r req in
+ let signature = [Some (r,req)],Some(r,req) in
+ let inv_m, inv_m_lem =
+ try Rewrite.default_morphism signature inv
+ with Not_found ->
+ error "field inverse should be declared as a morphism" in
+ inv_m_lem
+
+let add_field_theory name fth eqth morphth cst_tac inj (pre,post) power sign odiv =
+ check_required_library (cdir@["Field_tac"]);
+ let env = Global.env() in
+ let sigma = Evd.empty in
+ let (kind,r,zero,one,add,mul,sub,opp,div,inv,req,rth) =
+ dest_field env sigma fth in
+ let (sth,ext) = build_setoid_params r add mul opp req eqth in
+ let eqth = Some(sth,ext) in
+ let _ = add_theory name rth eqth morphth cst_tac (None,None) power sign odiv in
+ let (pow_tac, pspec) = interp_power env power in
+ let sspec = interp_sign env sign in
+ let dspec = interp_div env odiv in
+ let inv_m = field_equality r inv req in
+ let rk = reflect_coeff morphth in
+ let params =
+ exec_tactic env 9 (field_ltac"field_lemmas")
+ (List.map carg[sth;ext;inv_m;fth;pspec;sspec;dspec;rk]) in
+ let lemma1 = constr_of params.(3) in
+ let lemma2 = constr_of params.(4) in
+ let lemma3 = constr_of params.(5) in
+ let lemma4 = constr_of params.(6) in
+ let cond_lemma =
+ match inj with
+ | Some thm -> mkApp(constr_of params.(8),[|thm|])
+ | None -> constr_of params.(7) in
+ let lemma1 = decl_constant (string_of_id name^"_field_lemma1") lemma1 in
+ let lemma2 = decl_constant (string_of_id name^"_field_lemma2") lemma2 in
+ let lemma3 = decl_constant (string_of_id name^"_field_lemma3") lemma3 in
+ let lemma4 = decl_constant (string_of_id name^"_field_lemma4") lemma4 in
+ let cond_lemma = decl_constant (string_of_id name^"_lemma5") cond_lemma in
+ let cst_tac =
+ interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in
+ let pretac =
+ match pre with
+ Some t -> Tacinterp.glob_tactic t
+ | _ -> TacId [] in
+ let posttac =
+ match post with
+ Some t -> Tacinterp.glob_tactic t
+ | _ -> TacId [] in
+ let _ =
+ Lib.add_leaf name
+ (ftheory_to_obj
+ { field_carrier = r;
+ field_req = req;
+ field_cst_tac = cst_tac;
+ field_pow_tac = pow_tac;
+ field_ok = lemma1;
+ field_simpl_eq_ok = lemma2;
+ field_simpl_ok = lemma3;
+ field_simpl_eq_in_ok = lemma4;
+ field_cond = cond_lemma;
+ field_pre_tac = pretac;
+ field_post_tac = posttac }) in ()
+
+type field_mod =
+ Ring_mod of ring_mod
+ | Inject of Topconstr.constr_expr
+
+VERNAC ARGUMENT EXTEND field_mod
+ | [ ring_mod(m) ] -> [ Ring_mod m ]
+ | [ "completeness" constr(inj) ] -> [ Inject inj ]
+END
+
+let process_field_mods l =
+ let kind = ref None in
+ let set = ref None in
+ let cst_tac = ref None in
+ let pre = ref None in
+ let post = ref None in
+ let inj = ref None in
+ let sign = ref None in
+ let power = ref None in
+ let div = ref None in
+ List.iter(function
+ Ring_mod(Ring_kind k) -> set_once "field kind" kind k
+ | Ring_mod(Const_tac t) ->
+ set_once "tactic recognizing constants" cst_tac t
+ | Ring_mod(Pre_tac t) -> set_once "preprocess tactic" pre t
+ | Ring_mod(Post_tac t) -> set_once "postprocess tactic" post t
+ | Ring_mod(Setoid(sth,ext)) -> set_once "setoid" set (ic sth,ic ext)
+ | Ring_mod(Pow_spec(t,spec)) -> set_once "power" power (t,spec)
+ | Ring_mod(Sign_spec t) -> set_once "sign" sign t
+ | Ring_mod(Div_spec t) -> set_once "div" div t
+ | Inject i -> set_once "infinite property" inj (ic i)) l;
+ let k = match !kind with Some k -> k | None -> Abstract in
+ (k, !set, !inj, !cst_tac, !pre, !post, !power, !sign, !div)
+
+VERNAC COMMAND EXTEND AddSetoidField
+| [ "Add" "Field" ident(id) ":" constr(t) field_mods(l) ] ->
+ [ let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in
+ add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div]
+END
+
+
+let ltac_field_structure e =
+ let req = carg e.field_req in
+ let cst_tac = Tacexp e.field_cst_tac in
+ let pow_tac = Tacexp e.field_pow_tac in
+ let field_ok = carg e.field_ok in
+ let field_simpl_ok = carg e.field_simpl_ok in
+ let field_simpl_eq_ok = carg e.field_simpl_eq_ok in
+ let field_simpl_eq_in_ok = carg e.field_simpl_eq_in_ok in
+ let cond_ok = carg e.field_cond in
+ let pretac = Tacexp(TacFun([None],e.field_pre_tac)) in
+ let posttac = Tacexp(TacFun([None],e.field_post_tac)) in
+ [req;cst_tac;pow_tac;field_ok;field_simpl_ok;field_simpl_eq_ok;
+ field_simpl_eq_in_ok;cond_ok;pretac;posttac]
+
+let field_lookup (f:glob_tactic_expr) lH rl t gl =
+ let env = pf_env gl in
+ let sigma = project gl in
+ let rl = make_args_list rl t in
+ let e = find_field_structure env sigma rl in
+ let rl = carg (make_term_list e.field_carrier rl) in
+ let lH = carg (make_hyp_list env lH) in
+ let field = ltac_field_structure e in
+ ltac_apply f (field@[lH;rl]) gl
+
+
+TACTIC EXTEND field_lookup
+| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] ->
+ [ let (t,l) = list_sep_last lt in field_lookup (fst f) lH l t ]
+END
diff --git a/plugins/setoid_ring/newring_plugin.mllib b/plugins/setoid_ring/newring_plugin.mllib
new file mode 100644
index 00000000..a98392f1
--- /dev/null
+++ b/plugins/setoid_ring/newring_plugin.mllib
@@ -0,0 +1,2 @@
+Newring
+Newring_plugin_mod
diff --git a/plugins/setoid_ring/vo.itarget b/plugins/setoid_ring/vo.itarget
new file mode 100644
index 00000000..6934375b
--- /dev/null
+++ b/plugins/setoid_ring/vo.itarget
@@ -0,0 +1,15 @@
+ArithRing.vo
+BinList.vo
+Field_tac.vo
+Field_theory.vo
+Field.vo
+InitialRing.vo
+NArithRing.vo
+RealField.vo
+Ring_base.vo
+Ring_equiv.vo
+Ring_polynom.vo
+Ring_tac.vo
+Ring_theory.vo
+Ring.vo
+ZArithRing.vo