summaryrefslogtreecommitdiff
path: root/contrib7
diff options
context:
space:
mode:
Diffstat (limited to 'contrib7')
-rw-r--r--contrib7/cc/CCSolve.v20
-rw-r--r--contrib7/correctness/ArrayPermut.v183
-rw-r--r--contrib7/correctness/Arrays.v75
-rw-r--r--contrib7/correctness/Arrays_stuff.v16
-rw-r--r--contrib7/correctness/Correctness.v25
-rw-r--r--contrib7/correctness/Exchange.v94
-rw-r--r--contrib7/correctness/ProgBool.v66
-rw-r--r--contrib7/correctness/ProgInt.v19
-rw-r--r--contrib7/correctness/ProgramsExtraction.v30
-rw-r--r--contrib7/correctness/Programs_stuff.v13
-rw-r--r--contrib7/correctness/Sorted.v198
-rw-r--r--contrib7/correctness/Tuples.v106
-rw-r--r--contrib7/correctness/preuves.v128
-rw-r--r--contrib7/extraction/test_extraction.v533
-rw-r--r--contrib7/field/Field.v15
-rw-r--r--contrib7/field/Field_Compl.v62
-rw-r--r--contrib7/field/Field_Tactic.v397
-rw-r--r--contrib7/field/Field_Theory.v612
-rw-r--r--contrib7/fourier/Fourier.v28
-rw-r--r--contrib7/fourier/Fourier_util.v236
-rw-r--r--contrib7/interface/AddDad.v19
-rw-r--r--contrib7/interface/Centaur.v88
-rw-r--r--contrib7/interface/vernacrc17
-rw-r--r--contrib7/omega/Omega.v57
-rw-r--r--contrib7/omega/OmegaLemmas.v399
-rw-r--r--contrib7/ring/ArithRing.v81
-rw-r--r--contrib7/ring/NArithRing.v44
-rw-r--r--contrib7/ring/Quote.v85
-rw-r--r--contrib7/ring/Ring.v34
-rw-r--r--contrib7/ring/Ring_abstract.v699
-rw-r--r--contrib7/ring/Ring_normalize.v893
-rw-r--r--contrib7/ring/Ring_theory.v384
-rw-r--r--contrib7/ring/Setoid_ring.v13
-rw-r--r--contrib7/ring/Setoid_ring_normalize.v1141
-rw-r--r--contrib7/ring/Setoid_ring_theory.v429
-rw-r--r--contrib7/ring/ZArithRing.v35
-rw-r--r--contrib7/romega/ROmega.v12
-rw-r--r--contrib7/romega/ReflOmegaCore.v2602
38 files changed, 9888 insertions, 0 deletions
diff --git a/contrib7/cc/CCSolve.v b/contrib7/cc/CCSolve.v
new file mode 100644
index 00000000..388763ed
--- /dev/null
+++ b/contrib7/cc/CCSolve.v
@@ -0,0 +1,20 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: CCSolve.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Tactic Definition CCsolve :=
+ Repeat (Match Context With
+ [ H: ?1 |- ?2] ->
+ Let Heq = FreshId "Heq" In
+ (Assert Heq:(?2==?1);[Congruence|(Rewrite Heq;Exact H)])
+ |[ H: ?1; G: ?2 -> ?3 |- ?] ->
+ Let Heq = FreshId "Heq" In
+ (Assert Heq:(?2==?1) ;[Congruence|
+ (Rewrite Heq in G;Generalize (G H);Clear G;Intro G)])).
+
diff --git a/contrib7/correctness/ArrayPermut.v b/contrib7/correctness/ArrayPermut.v
new file mode 100644
index 00000000..4a0025ca
--- /dev/null
+++ b/contrib7/correctness/ArrayPermut.v
@@ -0,0 +1,183 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: ArrayPermut.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+(****************************************************************************)
+(* Permutations of elements in arrays *)
+(* Definition and properties *)
+(****************************************************************************)
+
+Require ProgInt.
+Require Arrays.
+Require Export Exchange.
+
+Require Omega.
+
+Set Implicit Arguments.
+
+(* We define "permut" as the smallest equivalence relation which contains
+ * transpositions i.e. exchange of two elements.
+ *)
+
+Inductive permut [n:Z; A:Set] : (array n A)->(array n A)->Prop :=
+ exchange_is_permut :
+ (t,t':(array n A))(i,j:Z)(exchange t t' i j) -> (permut t t')
+ | permut_refl :
+ (t:(array n A))(permut t t)
+ | permut_sym :
+ (t,t':(array n A))(permut t t') -> (permut t' t)
+ | permut_trans :
+ (t,t',t'':(array n A))
+ (permut t t') -> (permut t' t'') -> (permut t t'').
+
+Hints Resolve exchange_is_permut permut_refl permut_sym permut_trans : v62 datatypes.
+
+(* We also define the permutation on a segment of an array, "sub_permut",
+ * the other parts of the array being unchanged
+ *
+ * One again we define it as the smallest equivalence relation containing
+ * transpositions on the given segment.
+ *)
+
+Inductive sub_permut [n:Z; A:Set; g,d:Z] : (array n A)->(array n A)->Prop :=
+ exchange_is_sub_permut :
+ (t,t':(array n A))(i,j:Z)`g <= i <= d` -> `g <= j <= d`
+ -> (exchange t t' i j) -> (sub_permut g d t t')
+ | sub_permut_refl :
+ (t:(array n A))(sub_permut g d t t)
+ | sub_permut_sym :
+ (t,t':(array n A))(sub_permut g d t t') -> (sub_permut g d t' t)
+ | sub_permut_trans :
+ (t,t',t'':(array n A))
+ (sub_permut g d t t') -> (sub_permut g d t' t'')
+ -> (sub_permut g d t t'').
+
+Hints Resolve exchange_is_sub_permut sub_permut_refl sub_permut_sym sub_permut_trans
+ : v62 datatypes.
+
+(* To express that some parts of arrays are equal we introduce the
+ * property "array_id" which says that a segment is the same on two
+ * arrays.
+ *)
+
+Definition array_id := [n:Z][A:Set][t,t':(array n A)][g,d:Z]
+ (i:Z) `g <= i <= d` -> #t[i] = #t'[i].
+
+(* array_id is an equivalence relation *)
+
+Lemma array_id_refl :
+ (n:Z)(A:Set)(t:(array n A))(g,d:Z)
+ (array_id t t g d).
+Proof.
+Unfold array_id.
+Auto with datatypes.
+Save.
+
+Hints Resolve array_id_refl : v62 datatypes.
+
+Lemma array_id_sym :
+ (n:Z)(A:Set)(t,t':(array n A))(g,d:Z)
+ (array_id t t' g d)
+ -> (array_id t' t g d).
+Proof.
+Unfold array_id. Intros.
+Symmetry; Auto with datatypes.
+Save.
+
+Hints Resolve array_id_sym : v62 datatypes.
+
+Lemma array_id_trans :
+ (n:Z)(A:Set)(t,t',t'':(array n A))(g,d:Z)
+ (array_id t t' g d)
+ -> (array_id t' t'' g d)
+ -> (array_id t t'' g d).
+Proof.
+Unfold array_id. Intros.
+Apply trans_eq with y:=#t'[i]; Auto with datatypes.
+Save.
+
+Hints Resolve array_id_trans: v62 datatypes.
+
+(* Outside the segment [g,d] the elements are equal *)
+
+Lemma sub_permut_id :
+ (n:Z)(A:Set)(t,t':(array n A))(g,d:Z)
+ (sub_permut g d t t') ->
+ (array_id t t' `0` `g-1`) /\ (array_id t t' `d+1` `n-1`).
+Proof.
+Intros n A t t' g d. Induction 1; Intros.
+Elim H2; Intros.
+Unfold array_id; Split; Intros.
+Apply H7; Omega.
+Apply H7; Omega.
+Auto with datatypes.
+Decompose [and] H1; Auto with datatypes.
+Decompose [and] H1; Decompose [and] H3; EAuto with datatypes.
+Save.
+
+Hints Resolve sub_permut_id.
+
+Lemma sub_permut_eq :
+ (n:Z)(A:Set)(t,t':(array n A))(g,d:Z)
+ (sub_permut g d t t') ->
+ (i:Z) (`0<=i<g` \/ `d<i<n`) -> #t[i]=#t'[i].
+Proof.
+Intros n A t t' g d Htt' i Hi.
+Elim (sub_permut_id Htt'). Unfold array_id.
+Intros.
+Elim Hi; [ Intro; Apply H; Omega | Intro; Apply H0; Omega ].
+Save.
+
+(* sub_permut is a particular case of permutation *)
+
+Lemma sub_permut_is_permut :
+ (n:Z)(A:Set)(t,t':(array n A))(g,d:Z)
+ (sub_permut g d t t') ->
+ (permut t t').
+Proof.
+Intros n A t t' g d. Induction 1; Intros; EAuto with datatypes.
+Save.
+
+Hints Resolve sub_permut_is_permut.
+
+(* If we have a sub-permutation on an empty segment, then we have a
+ * sub-permutation on any segment.
+ *)
+
+Lemma sub_permut_void :
+ (N:Z)(A:Set)(t,t':(array N A))
+ (g,g',d,d':Z) `d < g`
+ -> (sub_permut g d t t') -> (sub_permut g' d' t t').
+Proof.
+Intros N A t t' g g' d d' Hdg.
+(Induction 1; Intros).
+(Absurd `g <= d`; Omega).
+Auto with datatypes.
+Auto with datatypes.
+EAuto with datatypes.
+Save.
+
+(* A sub-permutation on a segment may be extended to any segment that
+ * contains the first one.
+ *)
+
+Lemma sub_permut_extension :
+ (N:Z)(A:Set)(t,t':(array N A))
+ (g,g',d,d':Z) `g' <= g` -> `d <= d'`
+ -> (sub_permut g d t t') -> (sub_permut g' d' t t').
+Proof.
+Intros N A t t' g g' d d' Hgg' Hdd'.
+(Induction 1; Intros).
+Apply exchange_is_sub_permut with i:=i j:=j; [ Omega | Omega | Assumption ].
+Auto with datatypes.
+Auto with datatypes.
+EAuto with datatypes.
+Save.
diff --git a/contrib7/correctness/Arrays.v b/contrib7/correctness/Arrays.v
new file mode 100644
index 00000000..3fdc78c1
--- /dev/null
+++ b/contrib7/correctness/Arrays.v
@@ -0,0 +1,75 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: Arrays.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+(**********************************************)
+(* Functional arrays, for use in Correctness. *)
+(**********************************************)
+
+(* This is an axiomatization of arrays.
+ *
+ * The type (array N T) is the type of arrays ranging from 0 to N-1
+ * which elements are of type T.
+ *
+ * Arrays are created with new, accessed with access and modified with store.
+ *
+ * Operations of accessing and storing are not guarded, but axioms are.
+ * So these arrays can be viewed as arrays where accessing and storing
+ * out of the bounds has no effect.
+ *)
+
+
+Require Export ProgInt.
+
+Set Implicit Arguments.
+
+
+(* The type of arrays *)
+
+Parameter array : Z -> Set -> Set.
+
+
+(* Functions to create, access and modify arrays *)
+
+Parameter new : (n:Z)(T:Set) T -> (array n T).
+
+Parameter access : (n:Z)(T:Set) (array n T) -> Z -> T.
+
+Parameter store : (n:Z)(T:Set) (array n T) -> Z -> T -> (array n T).
+
+
+(* Axioms *)
+
+Axiom new_def : (n:Z)(T:Set)(v0:T)
+ (i:Z) `0<=i<n` -> (access (new n v0) i) = v0.
+
+Axiom store_def_1 : (n:Z)(T:Set)(t:(array n T))(v:T)
+ (i:Z) `0<=i<n` ->
+ (access (store t i v) i) = v.
+
+Axiom store_def_2 : (n:Z)(T:Set)(t:(array n T))(v:T)
+ (i:Z)(j:Z) `0<=i<n` -> `0<=j<n` ->
+ `i <> j` ->
+ (access (store t i v) j) = (access t j).
+
+Hints Resolve new_def store_def_1 store_def_2 : datatypes v62.
+
+(* A tactic to simplify access in arrays *)
+
+Tactic Definition ArrayAccess i j H :=
+ Elim (Z_eq_dec i j); [
+ Intro H; Rewrite H; Rewrite store_def_1
+ | Intro H; Rewrite store_def_2; [ Idtac | Idtac | Idtac | Exact H ] ].
+
+(* Symbolic notation for access *)
+
+Notation "# t [ c ]" := (access t c) (at level 0, t ident)
+ V8only (at level 0, t at level 0).
diff --git a/contrib7/correctness/Arrays_stuff.v b/contrib7/correctness/Arrays_stuff.v
new file mode 100644
index 00000000..448b0ab6
--- /dev/null
+++ b/contrib7/correctness/Arrays_stuff.v
@@ -0,0 +1,16 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: Arrays_stuff.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Require Export Exchange.
+Require Export ArrayPermut.
+Require Export Sorted.
+
diff --git a/contrib7/correctness/Correctness.v b/contrib7/correctness/Correctness.v
new file mode 100644
index 00000000..b0fde165
--- /dev/null
+++ b/contrib7/correctness/Correctness.v
@@ -0,0 +1,25 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: Correctness.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+(* Correctness is base on the tactic Refine (developped on purpose) *)
+
+Require Export Tuples.
+
+Require Export ProgInt.
+Require Export ProgBool.
+Require Export Zwf.
+
+Require Export Arrays.
+
+(*
+Token "'".
+*)
diff --git a/contrib7/correctness/Exchange.v b/contrib7/correctness/Exchange.v
new file mode 100644
index 00000000..12c8c9de
--- /dev/null
+++ b/contrib7/correctness/Exchange.v
@@ -0,0 +1,94 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: Exchange.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+(****************************************************************************)
+(* Exchange of two elements in an array *)
+(* Definition and properties *)
+(****************************************************************************)
+
+Require ProgInt.
+Require Arrays.
+
+Set Implicit Arguments.
+
+(* Definition *)
+
+Inductive exchange [n:Z; A:Set; t,t':(array n A); i,j:Z] : Prop :=
+ exchange_c :
+ `0<=i<n` -> `0<=j<n` ->
+ (#t[i] = #t'[j]) ->
+ (#t[j] = #t'[i]) ->
+ ((k:Z)`0<=k<n` -> `k<>i` -> `k<>j` -> #t[k] = #t'[k]) ->
+ (exchange t t' i j).
+
+(* Properties about exchanges *)
+
+Lemma exchange_1 : (n:Z)(A:Set)(t:(array n A))
+ (i,j:Z) `0<=i<n` -> `0<=j<n` ->
+ (access (store (store t i #t[j]) j #t[i]) i) = #t[j].
+Proof.
+Intros n A t i j H_i H_j.
+Case (dec_eq j i).
+Intro eq_i_j. Rewrite eq_i_j.
+Auto with datatypes.
+Intro not_j_i.
+Rewrite (store_def_2 (store t i #t[j]) #t[i] H_j H_i not_j_i).
+Auto with datatypes.
+Save.
+
+Hints Resolve exchange_1 : v62 datatypes.
+
+
+Lemma exchange_proof :
+ (n:Z)(A:Set)(t:(array n A))
+ (i,j:Z) `0<=i<n` -> `0<=j<n` ->
+ (exchange (store (store t i (access t j)) j (access t i)) t i j).
+Proof.
+Intros n A t i j H_i H_j.
+Apply exchange_c; Auto with datatypes.
+Intros k H_k not_k_i not_k_j.
+Cut ~j=k; Auto with datatypes. Intro not_j_k.
+Rewrite (store_def_2 (store t i (access t j)) (access t i) H_j H_k not_j_k).
+Auto with datatypes.
+Save.
+
+Hints Resolve exchange_proof : v62 datatypes.
+
+
+Lemma exchange_sym :
+ (n:Z)(A:Set)(t,t':(array n A))(i,j:Z)
+ (exchange t t' i j) -> (exchange t' t i j).
+Proof.
+Intros n A t t' i j H1.
+Elim H1. Clear H1. Intros.
+Constructor 1; Auto with datatypes.
+Intros. Rewrite (H3 k); Auto with datatypes.
+Save.
+
+Hints Resolve exchange_sym : v62 datatypes.
+
+
+Lemma exchange_id :
+ (n:Z)(A:Set)(t,t':(array n A))(i,j:Z)
+ (exchange t t' i j) ->
+ i=j ->
+ (k:Z) `0 <= k < n` -> (access t k)=(access t' k).
+Proof.
+Intros n A t t' i j Hex Heq k Hk.
+Elim Hex. Clear Hex. Intros.
+Rewrite Heq in H1. Rewrite Heq in H2.
+Case (Z_eq_dec k j).
+ Intro Heq'. Rewrite Heq'. Assumption.
+ Intro Hnoteq. Apply (H3 k); Auto with datatypes. Rewrite Heq. Assumption.
+Save.
+
+Hints Resolve exchange_id : v62 datatypes.
diff --git a/contrib7/correctness/ProgBool.v b/contrib7/correctness/ProgBool.v
new file mode 100644
index 00000000..c7a7687d
--- /dev/null
+++ b/contrib7/correctness/ProgBool.v
@@ -0,0 +1,66 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: ProgBool.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Require ZArith.
+Require Export Bool_nat.
+Require Export Sumbool.
+
+Definition annot_bool :
+ (b:bool) { b':bool | if b' then b=true else b=false }.
+Proof.
+Intro b.
+Exists b. Case b; Trivial.
+Save.
+
+
+(* Logical connectives *)
+
+Definition spec_and := [A,B,C,D:Prop][b:bool]if b then A /\ C else B \/ D.
+
+Definition prog_bool_and :
+ (Q1,Q2:bool->Prop) (sig bool Q1) -> (sig bool Q2)
+ -> { b:bool | if b then (Q1 true) /\ (Q2 true)
+ else (Q1 false) \/ (Q2 false) }.
+Proof.
+Intros Q1 Q2 H1 H2.
+Elim H1. Intro b1. Elim H2. Intro b2.
+Case b1; Case b2; Intros.
+Exists true; Auto.
+Exists false; Auto. Exists false; Auto. Exists false; Auto.
+Save.
+
+Definition spec_or := [A,B,C,D:Prop][b:bool]if b then A \/ C else B /\ D.
+
+Definition prog_bool_or :
+ (Q1,Q2:bool->Prop) (sig bool Q1) -> (sig bool Q2)
+ -> { b:bool | if b then (Q1 true) \/ (Q2 true)
+ else (Q1 false) /\ (Q2 false) }.
+Proof.
+Intros Q1 Q2 H1 H2.
+Elim H1. Intro b1. Elim H2. Intro b2.
+Case b1; Case b2; Intros.
+Exists true; Auto. Exists true; Auto. Exists true; Auto.
+Exists false; Auto.
+Save.
+
+Definition spec_not:= [A,B:Prop][b:bool]if b then B else A.
+
+Definition prog_bool_not :
+ (Q:bool->Prop) (sig bool Q)
+ -> { b:bool | if b then (Q false) else (Q true) }.
+Proof.
+Intros Q H.
+Elim H. Intro b.
+Case b; Intro.
+Exists false; Auto. Exists true; Auto.
+Save.
+
diff --git a/contrib7/correctness/ProgInt.v b/contrib7/correctness/ProgInt.v
new file mode 100644
index 00000000..0ca830c2
--- /dev/null
+++ b/contrib7/correctness/ProgInt.v
@@ -0,0 +1,19 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: ProgInt.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Require Export ZArith.
+Require Export ZArith_dec.
+
+Theorem Znotzero : (x:Z){`x<>0`}+{`x=0`}.
+Proof.
+Intro x. Elim (Z_eq_dec x `0`) ; Auto.
+Save.
diff --git a/contrib7/correctness/ProgramsExtraction.v b/contrib7/correctness/ProgramsExtraction.v
new file mode 100644
index 00000000..20f82ce4
--- /dev/null
+++ b/contrib7/correctness/ProgramsExtraction.v
@@ -0,0 +1,30 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: ProgramsExtraction.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Require Export Extraction.
+
+Extract Inductive unit => unit [ "()" ].
+Extract Inductive bool => bool [ true false ].
+Extract Inductive sumbool => bool [ true false ].
+
+Require Export Correctness.
+
+Declare ML Module "pextract".
+
+Grammar vernac vernac : ast :=
+ imperative_ocaml [ "Write" "Caml" "File" stringarg($file)
+ "[" ne_identarg_list($idl) "]" "." ]
+ -> [ (IMPERATIVEEXTRACTION $file (VERNACARGLIST ($LIST $idl))) ]
+
+| initialize [ "Initialize" identarg($id) "with" comarg($c) "." ]
+ -> [ (INITIALIZE $id $c) ]
+.
diff --git a/contrib7/correctness/Programs_stuff.v b/contrib7/correctness/Programs_stuff.v
new file mode 100644
index 00000000..00beeaeb
--- /dev/null
+++ b/contrib7/correctness/Programs_stuff.v
@@ -0,0 +1,13 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: Programs_stuff.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Require Export Arrays_stuff.
diff --git a/contrib7/correctness/Sorted.v b/contrib7/correctness/Sorted.v
new file mode 100644
index 00000000..f476142e
--- /dev/null
+++ b/contrib7/correctness/Sorted.v
@@ -0,0 +1,198 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Library about sorted (sub-)arrays / Nicolas Magaud, July 1998 *)
+
+(* $Id: Sorted.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+Require Export Arrays.
+Require ArrayPermut.
+
+Require ZArithRing.
+Require Omega.
+V7only [Import Z_scope.].
+Open Local Scope Z_scope.
+
+Set Implicit Arguments.
+
+(* Definition *)
+
+Definition sorted_array :=
+ [N:Z][A:(array N Z)][deb:Z][fin:Z]
+ `deb<=fin` -> (x:Z) `x>=deb` -> `x<fin` -> (Zle #A[x] #A[`x+1`]).
+
+(* Elements of a sorted sub-array are in increasing order *)
+
+(* one element and the next one *)
+
+Lemma sorted_elements_1 :
+ (N:Z)(A:(array N Z))(n:Z)(m:Z)
+ (sorted_array A n m)
+ -> (k:Z)`k>=n`
+ -> (i:Z) `0<=i` -> `k+i<=m`
+ -> (Zle (access A k) (access A `k+i`)).
+Proof.
+Intros N A n m H_sorted k H_k i H_i.
+Pattern i. Apply natlike_ind.
+Intro.
+Replace `k+0` with k; Omega. (*** Ring `k+0` => BUG ***)
+
+Intros.
+Apply Zle_trans with m:=(access A `k+x`).
+Apply H0 ; Omega.
+
+Unfold Zs.
+Replace `k+(x+1)` with `(k+x)+1`.
+Unfold sorted_array in H_sorted.
+Apply H_sorted ; Omega.
+
+Omega.
+
+Assumption.
+Save.
+
+(* one element and any of the following *)
+
+Lemma sorted_elements :
+ (N:Z)(A:(array N Z))(n:Z)(m:Z)(k:Z)(l:Z)
+ (sorted_array A n m)
+ -> `k>=n` -> `l<N` -> `k<=l` -> `l<=m`
+ -> (Zle (access A k) (access A l)).
+Proof.
+Intros.
+Replace l with `k+(l-k)`.
+Apply sorted_elements_1 with n:=n m:=m; [Assumption | Omega | Omega | Omega].
+Omega.
+Save.
+
+Hints Resolve sorted_elements : datatypes v62.
+
+(* A sub-array of a sorted array is sorted *)
+
+Lemma sub_sorted_array : (N:Z)(A:(array N Z))(deb:Z)(fin:Z)(i:Z)(j:Z)
+ (sorted_array A deb fin) ->
+ (`i>=deb` -> `j<=fin` -> `i<=j` -> (sorted_array A i j)).
+Proof.
+Unfold sorted_array.
+Intros.
+Apply H ; Omega.
+Save.
+
+Hints Resolve sub_sorted_array : datatypes v62.
+
+(* Extension on the left of the property of being sorted *)
+
+Lemma left_extension : (N:Z)(A:(array N Z))(i:Z)(j:Z)
+ `i>0` -> `j<N` -> (sorted_array A i j)
+ -> (Zle #A[`i-1`] #A[i]) -> (sorted_array A `i-1` j).
+Proof.
+(Intros; Unfold sorted_array ; Intros).
+Elim (Z_ge_lt_dec x i). (* (`x >= i`) + (`x < i`) *)
+Intro Hcut.
+Apply H1 ; Omega.
+
+Intro Hcut.
+Replace x with `i-1`.
+Replace `i-1+1` with i ; [Assumption | Omega].
+
+Omega.
+Save.
+
+(* Extension on the right *)
+
+Lemma right_extension : (N:Z)(A:(array N Z))(i:Z)(j:Z)
+ `i>=0` -> `j<N-1` -> (sorted_array A i j)
+ -> (Zle #A[j] #A[`j+1`]) -> (sorted_array A i `j+1`).
+Proof.
+(Intros; Unfold sorted_array ; Intros).
+Elim (Z_lt_ge_dec x j).
+Intro Hcut.
+Apply H1 ; Omega.
+
+Intro HCut.
+Replace x with j ; [Assumption | Omega].
+Save.
+
+(* Substitution of the leftmost value by a smaller value *)
+
+Lemma left_substitution :
+ (N:Z)(A:(array N Z))(i:Z)(j:Z)(v:Z)
+ `i>=0` -> `j<N` -> (sorted_array A i j)
+ -> (Zle v #A[i])
+ -> (sorted_array (store A i v) i j).
+Proof.
+Intros N A i j v H_i H_j H_sorted H_v.
+Unfold sorted_array ; Intros.
+
+Cut `x = i`\/`x > i`.
+(Intro Hcut ; Elim Hcut ; Clear Hcut ; Intro).
+Rewrite H2.
+Rewrite store_def_1 ; Try Omega.
+Rewrite store_def_2 ; Try Omega.
+Apply Zle_trans with m:=(access A i) ; [Assumption | Apply H_sorted ; Omega].
+
+(Rewrite store_def_2; Try Omega).
+(Rewrite store_def_2; Try Omega).
+Apply H_sorted ; Omega.
+Omega.
+Save.
+
+(* Substitution of the rightmost value by a larger value *)
+
+Lemma right_substitution :
+ (N:Z)(A:(array N Z))(i:Z)(j:Z)(v:Z)
+ `i>=0` -> `j<N` -> (sorted_array A i j)
+ -> (Zle #A[j] v)
+ -> (sorted_array (store A j v) i j).
+Proof.
+Intros N A i j v H_i H_j H_sorted H_v.
+Unfold sorted_array ; Intros.
+
+Cut `x = j-1`\/`x < j-1`.
+(Intro Hcut ; Elim Hcut ; Clear Hcut ; Intro).
+Rewrite H2.
+Replace `j-1+1` with j; [ Idtac | Omega ]. (*** Ring `j-1+1`. => BUG ***)
+Rewrite store_def_2 ; Try Omega.
+Rewrite store_def_1 ; Try Omega.
+Apply Zle_trans with m:=(access A j).
+Apply sorted_elements with n:=i m:=j ; Try Omega ; Assumption.
+Assumption.
+
+(Rewrite store_def_2; Try Omega).
+(Rewrite store_def_2; Try Omega).
+Apply H_sorted ; Omega.
+
+Omega.
+Save.
+
+(* Affectation outside of the sorted region *)
+
+Lemma no_effect :
+ (N:Z)(A:(array N Z))(i:Z)(j:Z)(k:Z)(v:Z)
+ `i>=0` -> `j<N` -> (sorted_array A i j)
+ -> `0<=k<i`\/`j<k<N`
+ -> (sorted_array (store A k v) i j).
+Proof.
+Intros.
+Unfold sorted_array ; Intros.
+Rewrite store_def_2 ; Try Omega.
+Rewrite store_def_2 ; Try Omega.
+Apply H1 ; Assumption.
+Save.
+
+Lemma sorted_array_id : (N:Z)(t1,t2:(array N Z))(g,d:Z)
+ (sorted_array t1 g d) -> (array_id t1 t2 g d) -> (sorted_array t2 g d).
+Proof.
+Intros N t1 t2 g d Hsorted Hid.
+Unfold array_id in Hid.
+Unfold sorted_array in Hsorted. Unfold sorted_array.
+Intros Hgd x H1x H2x.
+Rewrite <- (Hid x); [ Idtac | Omega ].
+Rewrite <- (Hid `x+1`); [ Idtac | Omega ].
+Apply Hsorted; Assumption.
+Save.
diff --git a/contrib7/correctness/Tuples.v b/contrib7/correctness/Tuples.v
new file mode 100644
index 00000000..6e1eb03a
--- /dev/null
+++ b/contrib7/correctness/Tuples.v
@@ -0,0 +1,106 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* Certification of Imperative Programs / Jean-Christophe Filliâtre *)
+
+(* $Id: Tuples.v,v 1.1.2.1 2004/07/16 19:30:16 herbelin Exp $ *)
+
+(* Tuples *)
+
+Definition tuple_1 := [X:Set]X.
+Definition tuple_2 := prod.
+Definition Build_tuple_2 := pair.
+Definition proj_2_1 := fst.
+Definition proj_2_2 := snd.
+
+Record tuple_3 [ T1,T2,T3 : Set ] : Set :=
+ { proj_3_1 : T1 ;
+ proj_3_2 : T2 ;
+ proj_3_3 : T3 }.
+
+Record tuple_4 [ T1,T2,T3,T4 : Set ] : Set :=
+ { proj_4_1 : T1 ;
+ proj_4_2 : T2 ;
+ proj_4_3 : T3 ;
+ proj_4_4 : T4 }.
+
+Record tuple_5 [ T1,T2,T3,T4,T5 : Set ] : Set :=
+ { proj_5_1 : T1 ;
+ proj_5_2 : T2 ;
+ proj_5_3 : T3 ;
+ proj_5_4 : T4 ;
+ proj_5_5 : T5 }.
+
+Record tuple_6 [ T1,T2,T3,T4,T5,T6 : Set ] : Set :=
+ { proj_6_1 : T1 ;
+ proj_6_2 : T2 ;
+ proj_6_3 : T3 ;
+ proj_6_4 : T4 ;
+ proj_6_5 : T5 ;
+ proj_6_6 : T6 }.
+
+Record tuple_7 [ T1,T2,T3,T4,T5,T6,T7 : Set ] : Set :=
+ { proj_7_1 : T1 ;
+ proj_7_2 : T2 ;
+ proj_7_3 : T3 ;
+ proj_7_4 : T4 ;
+ proj_7_5 : T5 ;
+ proj_7_6 : T6 ;
+ proj_7_7 : T7 }.
+
+
+(* Existentials *)
+
+Definition sig_1 := sig.
+Definition exist_1 := exist.
+
+Inductive sig_2 [ T1,T2 : Set; P:T1->T2->Prop ] : Set :=
+ exist_2 : (x1:T1)(x2:T2)(P x1 x2) -> (sig_2 T1 T2 P).
+
+Inductive sig_3 [ T1,T2,T3 : Set; P:T1->T2->T3->Prop ] : Set :=
+ exist_3 : (x1:T1)(x2:T2)(x3:T3)(P x1 x2 x3) -> (sig_3 T1 T2 T3 P).
+
+
+Inductive sig_4 [ T1,T2,T3,T4 : Set;
+ P:T1->T2->T3->T4->Prop ] : Set :=
+ exist_4 : (x1:T1)(x2:T2)(x3:T3)(x4:T4)
+ (P x1 x2 x3 x4)
+ -> (sig_4 T1 T2 T3 T4 P).
+
+Inductive sig_5 [ T1,T2,T3,T4,T5 : Set;
+ P:T1->T2->T3->T4->T5->Prop ] : Set :=
+ exist_5 : (x1:T1)(x2:T2)(x3:T3)(x4:T4)(x5:T5)
+ (P x1 x2 x3 x4 x5)
+ -> (sig_5 T1 T2 T3 T4 T5 P).
+
+Inductive sig_6 [ T1,T2,T3,T4,T5,T6 : Set;
+ P:T1->T2->T3->T4->T5->T6->Prop ] : Set :=
+ exist_6 : (x1:T1)(x2:T2)(x3:T3)(x4:T4)(x5:T5)(x6:T6)
+ (P x1 x2 x3 x4 x5 x6)
+ -> (sig_6 T1 T2 T3 T4 T5 T6 P).
+
+Inductive sig_7 [ T1,T2,T3,T4,T5,T6,T7 : Set;
+ P:T1->T2->T3->T4->T5->T6->T7->Prop ] : Set :=
+ exist_7 : (x1:T1)(x2:T2)(x3:T3)(x4:T4)(x5:T5)(x6:T6)(x7:T7)
+ (P x1 x2 x3 x4 x5 x6 x7)
+ -> (sig_7 T1 T2 T3 T4 T5 T6 T7 P).
+
+Inductive sig_8 [ T1,T2,T3,T4,T5,T6,T7,T8 : Set;
+ P:T1->T2->T3->T4->T5->T6->T7->T8->Prop ] : Set :=
+ exist_8 : (x1:T1)(x2:T2)(x3:T3)(x4:T4)(x5:T5)(x6:T6)(x7:T7)(x8:T8)
+ (P x1 x2 x3 x4 x5 x6 x7 x8)
+ -> (sig_8 T1 T2 T3 T4 T5 T6 T7 T8 P).
+
+Inductive dep_tuple_2 [ T1,T2 : Set; P:T1->T2->Set ] : Set :=
+ Build_dep_tuple_2 : (x1:T1)(x2:T2)(P x1 x2) -> (dep_tuple_2 T1 T2 P).
+
+Inductive dep_tuple_3 [ T1,T2,T3 : Set; P:T1->T2->T3->Set ] : Set :=
+ Build_dep_tuple_3 : (x1:T1)(x2:T2)(x3:T3)(P x1 x2 x3)
+ -> (dep_tuple_3 T1 T2 T3 P).
+
+
diff --git a/contrib7/correctness/preuves.v b/contrib7/correctness/preuves.v
new file mode 100644
index 00000000..33659b43
--- /dev/null
+++ b/contrib7/correctness/preuves.v
@@ -0,0 +1,128 @@
+
+(* Quelques preuves sur des programmes simples,
+ * juste histoire d'avoir un petit bench.
+ *)
+
+Require Correctness.
+Require Omega.
+
+Global Variable x : Z ref.
+Global Variable y : Z ref.
+Global Variable z : Z ref.
+Global Variable i : Z ref.
+Global Variable j : Z ref.
+Global Variable n : Z ref.
+Global Variable m : Z ref.
+Variable r : Z.
+Variable N : Z.
+Global Variable t : array N of Z.
+
+(**********************************************************************)
+
+Require Exchange.
+Require ArrayPermut.
+
+Correctness swap
+ fun (N:Z)(t:array N of Z)(i,j:Z) ->
+ { `0 <= i < N` /\ `0 <= j < N` }
+ (let v = t[i] in
+ begin
+ t[i] := t[j];
+ t[j] := v
+ end)
+ { (exchange t t@ i j) }.
+Proof.
+Auto with datatypes.
+Save.
+
+Correctness downheap
+ let rec downheap (N:Z)(t:array N of Z) : unit { variant `0` } =
+ (swap N t 0 0) { True }
+.
+
+(**********************************************************************)
+
+Global Variable x : Z ref.
+Debug on.
+Correctness assign0 (x := 0) { `x=0` }.
+Save.
+
+(**********************************************************************)
+
+Global Variable i : Z ref.
+Debug on.
+Correctness assign1 { `0 <= i` } (i := !i + 1) { `0 < i` }.
+Omega.
+Save.
+
+(**********************************************************************)
+
+Global Variable i : Z ref.
+Debug on.
+Correctness if0 { `0 <= i` } (if !i>0 then i:=!i-1 else tt) { `0 <= i` }.
+Omega.
+Save.
+
+(**********************************************************************)
+
+Global Variable i : Z ref.
+Debug on.
+Correctness assert0 { `0 <= i` } begin assert { `i=2` }; i:=!i-1 end { `i=1` }.
+
+(**********************************************************************)
+
+Correctness echange
+ { `0 <= i < N` /\ `0 <= j < N` }
+ begin
+ label B;
+ x := t[!i]; t[!i] := t[!j]; t[!j] := !x;
+ assert { #t[i] = #t@B[j] /\ #t[j] = #t@B[i] }
+ end.
+Proof.
+Auto with datatypes.
+Save.
+
+
+(**********************************************************************)
+
+(*
+ * while x <= y do x := x+1 done { y < x }
+ *)
+
+Correctness incrementation
+ while !x < !y do
+ { invariant True variant `(Zs y)-x` }
+ x := !x + 1
+ done
+ { `y < x` }.
+Proof.
+Exact (Zwf_well_founded `0`).
+Unfold Zwf. Omega.
+Exact I.
+Save.
+
+
+(************************************************************************)
+
+Correctness pivot1
+ begin
+ while (Z_lt_ge_dec !i r) do
+ { invariant True variant (Zminus (Zs r) i) } i := (Zs !i)
+ done;
+ while (Z_lt_ge_dec r !j) do
+ { invariant True variant (Zminus (Zs j) r) } j := (Zpred !j)
+ done
+ end
+ { `j <= r` /\ `r <= i` }.
+Proof.
+Exact (Zwf_well_founded `0`).
+Unfold Zwf. Omega.
+Exact I.
+Exact (Zwf_well_founded `0`).
+Unfold Zwf. Unfold Zpred. Omega.
+Exact I.
+Omega.
+Save.
+
+
+
diff --git a/contrib7/extraction/test_extraction.v b/contrib7/extraction/test_extraction.v
new file mode 100644
index 00000000..e76b1c69
--- /dev/null
+++ b/contrib7/extraction/test_extraction.v
@@ -0,0 +1,533 @@
+(************************************************************************)
+(* 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 Arith.
+Require PolyList.
+
+(*** STANDARD EXAMPLES *)
+
+(** Functions. *)
+
+Definition idnat := [x:nat]x.
+Extraction idnat.
+(* let idnat x = x *)
+
+Definition id := [X:Type][x:X]x.
+Extraction id. (* let id x = x *)
+Definition id' := (id Set nat).
+Extraction id'. (* type id' = nat *)
+
+Definition test2 := [f:nat->nat][x:nat](f x).
+Extraction test2.
+(* let test2 f x = f x *)
+
+Definition test3 := [f:nat->Set->nat][x:nat](f x nat).
+Extraction test3.
+(* let test3 f x = f x __ *)
+
+Definition test4 := [f:(nat->nat)->nat][x:nat][g:nat->nat](f g).
+Extraction test4.
+(* let test4 f x g = f g *)
+
+Definition test5 := ((1),(0)).
+Extraction test5.
+(* let test5 = Pair ((S O), O) *)
+
+Definition cf := [x:nat][_:(le x O)](S x).
+Extraction NoInline cf.
+Definition test6 := (cf O (le_n O)).
+Extraction test6.
+(* let test6 = cf O *)
+
+Definition test7 := ([X:Set][x:X]x nat).
+Extraction test7.
+(* let test7 x = x *)
+
+Definition d := [X:Type]X.
+Extraction d. (* type 'x d = 'x *)
+Definition d2 := (d Set).
+Extraction d2. (* type d2 = __ d *)
+Definition d3 := [x:(d Set)]O.
+Extraction d3. (* let d3 _ = O *)
+Definition d4 := (d nat).
+Extraction d4. (* type d4 = nat d *)
+Definition d5 := ([x:(d Type)]O Type).
+Extraction d5. (* let d5 = O *)
+Definition d6 := ([x:(d Type)]x).
+Extraction d6. (* type 'x d6 = 'x *)
+
+Definition test8 := ([X:Type][x:X]x Set nat).
+Extraction test8. (* type test8 = nat *)
+
+Definition test9 := let t = nat in (id Set t).
+Extraction test9. (* type test9 = nat *)
+
+Definition test10 := ([X:Type][x:X]O Type Type).
+Extraction test10. (* let test10 = O *)
+
+Definition test11 := let n=O in let p=(S n) in (S p).
+Extraction test11. (* let test11 = S (S O) *)
+
+Definition test12 := (x:(X:Type)X->X)(x Type Type).
+Extraction test12.
+(* type test12 = (__ -> __ -> __) -> __ *)
+
+
+Definition test13 := Cases (left True True I) of (left x)=>(S O) | (right x)=>O end.
+Extraction test13. (* let test13 = S O *)
+
+
+(** example with more arguments that given by the type *)
+
+Definition test19 := (nat_rec [n:nat]nat->nat [n:nat]O [n:nat][f:nat->nat]f O O).
+Extraction test19.
+(* let test19 =
+ let rec f = function
+ | O -> (fun n0 -> O)
+ | S n0 -> f n0
+ in f O O
+*)
+
+
+(** casts *)
+
+Definition test20 := (True :: Type).
+Extraction test20.
+(* type test20 = __ *)
+
+
+(** Simple inductive type and recursor. *)
+
+Extraction nat.
+(*
+type nat =
+ | O
+ | S of nat
+*)
+
+Extraction sumbool_rect.
+(*
+let sumbool_rect f f0 = function
+ | Left -> f __
+ | Right -> f0 __
+*)
+
+(** Less simple inductive type. *)
+
+Inductive c [x:nat] : nat -> Set :=
+ refl : (c x x)
+ | trans : (y,z:nat)(c x y)->(le y z)->(c x z).
+Extraction c.
+(*
+type c =
+ | Refl
+ | Trans of nat * nat * c
+*)
+
+Definition Ensemble := [U:Type]U->Prop.
+Definition Empty_set := [U:Type][x:U]False.
+Definition Add := [U:Type][A:(Ensemble U)][x:U][y:U](A y) \/ x==y.
+
+Inductive Finite [U:Type] : (Ensemble U) -> Set :=
+ Empty_is_finite: (Finite U (Empty_set U))
+ | Union_is_finite:
+ (A: (Ensemble U)) (Finite U A) ->
+ (x: U) ~ (A x) -> (Finite U (Add U A x)).
+Extraction Finite.
+(*
+type 'u finite =
+ | Empty_is_finite
+ | Union_is_finite of 'u finite * 'u
+*)
+
+
+(** Mutual Inductive *)
+
+Inductive tree : Set :=
+ Node : nat -> forest -> tree
+with forest : Set :=
+ | Leaf : nat -> forest
+ | Cons : tree -> forest -> forest .
+
+Extraction tree.
+(*
+type tree =
+ | Node of nat * forest
+and forest =
+ | Leaf of nat
+ | Cons of tree * forest
+*)
+
+Fixpoint tree_size [t:tree] : nat :=
+ Cases t of (Node a f) => (S (forest_size f)) end
+with forest_size [f:forest] : nat :=
+ Cases f of
+ | (Leaf b) => (S O)
+ | (Cons t f') => (plus (tree_size t) (forest_size f'))
+ end.
+
+Extraction tree_size.
+(*
+let rec tree_size = function
+ | Node (a, f) -> S (forest_size f)
+and forest_size = function
+ | Leaf b -> S O
+ | Cons (t, f') -> plus (tree_size t) (forest_size f')
+*)
+
+
+(** Eta-expansions of inductive constructor *)
+
+Inductive titi : Set := tata : nat->nat->nat->nat->titi.
+Definition test14 := (tata O).
+Extraction test14.
+(* let test14 x x0 x1 = Tata (O, x, x0, x1) *)
+Definition test15 := (tata O (S O)).
+Extraction test15.
+(* let test15 x x0 = Tata (O, (S O), x, x0) *)
+
+Inductive eta : Set := eta_c : nat->Prop->nat->Prop->eta.
+Extraction eta_c.
+(*
+type eta =
+ | Eta_c of nat * nat
+*)
+Definition test16 := (eta_c O).
+Extraction test16.
+(* let test16 x = Eta_c (O, x) *)
+Definition test17 := (eta_c O True).
+Extraction test17.
+(* let test17 x = Eta_c (O, x) *)
+Definition test18 := (eta_c O True O).
+Extraction test18.
+(* let test18 _ = Eta_c (O, O) *)
+
+
+(** Example of singleton inductive type *)
+
+Inductive bidon [A:Prop;B:Type] : Set := tb : (x:A)(y:B)(bidon A B).
+Definition fbidon := [A,B:Type][f:A->B->(bidon True nat)][x:A][y:B](f x y).
+Extraction bidon.
+(* type 'b bidon = 'b *)
+Extraction tb.
+(* tb : singleton inductive constructor *)
+Extraction fbidon.
+(* let fbidon f x y =
+ f x y
+*)
+
+Definition fbidon2 := (fbidon True nat (tb True nat)).
+Extraction fbidon2. (* let fbidon2 y = y *)
+Extraction NoInline fbidon.
+Extraction fbidon2.
+(* let fbidon2 y = fbidon (fun _ x -> x) __ y *)
+
+(* NB: first argument of fbidon2 has type [True], so it disappears. *)
+
+(** mutual inductive on many sorts *)
+
+Inductive
+ test_0 : Prop := ctest0 : test_0
+with
+ test_1 : Set := ctest1 : test_0-> test_1.
+Extraction test_0.
+(* test0 : logical inductive *)
+Extraction test_1.
+(*
+type test1 =
+ | Ctest1
+*)
+
+(** logical singleton *)
+
+Extraction eq.
+(* eq : logical inductive *)
+Extraction eq_rect.
+(* let eq_rect x f y =
+ f
+*)
+
+(** No more propagation of type parameters. Obj.t instead. *)
+
+Inductive tp1 : Set :=
+ T : (C:Set)(c:C)tp2 -> tp1 with tp2 : Set := T' : tp1->tp2.
+Extraction tp1.
+(*
+type tp1 =
+ | T of __ * tp2
+and tp2 =
+ | T' of tp1
+*)
+
+Inductive tp1bis : Set :=
+ Tbis : tp2bis -> tp1bis
+with tp2bis : Set := T'bis : (C:Set)(c:C)tp1bis->tp2bis.
+Extraction tp1bis.
+(*
+type tp1bis =
+ | Tbis of tp2bis
+and tp2bis =
+ | T'bis of __ * tp1bis
+*)
+
+
+(** Strange inductive type. *)
+
+Inductive Truc : Set->Set :=
+ chose : (A:Set)(Truc A)
+ | machin : (A:Set)A->(Truc bool)->(Truc A).
+Extraction Truc.
+(*
+type 'x truc =
+ | Chose
+ | Machin of 'x * bool truc
+*)
+
+
+(** Dependant type over Type *)
+
+Definition test24:= (sigT Set [a:Set](option a)).
+Extraction test24.
+(* type test24 = (__, __ option) sigT *)
+
+
+(** Coq term non strongly-normalizable after extraction *)
+
+Require Gt.
+Definition loop :=
+ [Ax:(Acc nat gt O)]
+ (Fix F {F [a:nat;b:(Acc nat gt a)] : nat :=
+ (F (S a) (Acc_inv nat gt a b (S a) (gt_Sn_n a)))}
+ O Ax).
+Extraction loop.
+(* let loop _ =
+ let rec f a =
+ f (S a)
+ in f O
+*)
+
+(*** EXAMPLES NEEDING OBJ.MAGIC *)
+
+(** False conversion of type: *)
+
+Lemma oups : (H:(nat==(list nat)))nat -> nat.
+Intros.
+Generalize H0;Intros.
+Rewrite H in H1.
+Case H1.
+Exact H0.
+Intros.
+Exact n.
+Qed.
+Extraction oups.
+(*
+let oups h0 =
+ match Obj.magic h0 with
+ | Nil -> h0
+ | Cons0 (n, l) -> n
+*)
+
+
+(** hybrids *)
+
+Definition horibilis := [b:bool]<[b:bool]if b then Type else nat>if b then Set else O.
+Extraction horibilis.
+(*
+let horibilis = function
+ | True -> Obj.magic __
+ | False -> Obj.magic O
+*)
+
+Definition PropSet := [b:bool]if b then Prop else Set.
+Extraction PropSet. (* type propSet = __ *)
+
+Definition natbool := [b:bool]if b then nat else bool.
+Extraction natbool. (* type natbool = __ *)
+
+Definition zerotrue := [b:bool]<natbool>if b then O else true.
+Extraction zerotrue.
+(*
+let zerotrue = function
+ | True -> Obj.magic O
+ | False -> Obj.magic True
+*)
+
+Definition natProp := [b:bool]<[_:bool]Type>if b then nat else Prop.
+
+Definition natTrue := [b:bool]<[_:bool]Type>if b then nat else True.
+
+Definition zeroTrue := [b:bool]<natProp>if b then O else True.
+Extraction zeroTrue.
+(*
+let zeroTrue = function
+ | True -> Obj.magic O
+ | False -> Obj.magic __
+*)
+
+Definition natTrue2 := [b:bool]<[_:bool]Type>if b then nat else True.
+
+Definition zeroprop := [b:bool]<natTrue>if b then O else I.
+Extraction zeroprop.
+(*
+let zeroprop = function
+ | True -> Obj.magic O
+ | False -> Obj.magic __
+*)
+
+(** polymorphic f applied several times *)
+
+Definition test21 := (id nat O, id bool true).
+Extraction test21.
+(* let test21 = Pair ((id O), (id True)) *)
+
+(** ok *)
+
+Definition test22 := ([f:(X:Type)X->X](f nat O, f bool true) [X:Type][x:X]x).
+Extraction test22.
+(* let test22 =
+ let f = fun x -> x in Pair ((f O), (f True)) *)
+
+(* still ok via optim beta -> let *)
+
+Definition test23 := [f:(X:Type)X->X](f nat O, f bool true).
+Extraction test23.
+(* let test23 f = Pair ((Obj.magic f __ O), (Obj.magic f __ True)) *)
+
+(* problem: fun f -> (f 0, f true) not legal in ocaml *)
+(* solution: magic ... *)
+
+
+(** Dummy constant __ can be applied.... *)
+
+Definition f : (X:Type)(nat->X)->(X->bool)->bool :=
+ [X:Type;x:nat->X;y:X->bool](y (x O)).
+Extraction f.
+(* let f x y =
+ y (x O)
+*)
+
+Definition f_prop := (f (O=O) [_](refl_equal ? O) [_]true).
+Extraction NoInline f.
+Extraction f_prop.
+(* let f_prop =
+ f (Obj.magic __) (fun _ -> True)
+*)
+
+Definition f_arity := (f Set [_:nat]nat [_:Set]true).
+Extraction f_arity.
+(* let f_arity =
+ f (Obj.magic __) (fun _ -> True)
+*)
+
+Definition f_normal := (f nat [x]x [x](Cases x of O => true | _ => false end)).
+Extraction f_normal.
+(* let f_normal =
+ f (fun x -> x) (fun x -> match x with
+ | O -> True
+ | S n -> False)
+*)
+
+
+(* inductive with magic needed *)
+
+Inductive Boite : Set :=
+ boite : (b:bool)(if b then nat else nat*nat)->Boite.
+Extraction Boite.
+(*
+type boite =
+ | Boite of bool * __
+*)
+
+
+Definition boite1 := (boite true O).
+Extraction boite1.
+(* let boite1 = Boite (True, (Obj.magic O)) *)
+
+Definition boite2 := (boite false (O,O)).
+Extraction boite2.
+(* let boite2 = Boite (False, (Obj.magic (Pair (O, O)))) *)
+
+Definition test_boite := [B:Boite]<nat>Cases B of
+ (boite true n) => n
+| (boite false n) => (plus (fst ? ? n) (snd ? ? n))
+end.
+Extraction test_boite.
+(*
+let test_boite = function
+ | Boite (b0, n) ->
+ (match b0 with
+ | True -> Obj.magic n
+ | False -> plus (fst (Obj.magic n)) (snd (Obj.magic n)))
+*)
+
+(* singleton inductive with magic needed *)
+
+Inductive Box : Set :=
+ box : (A:Set)A -> Box.
+Extraction Box.
+(* type box = __ *)
+
+Definition box1 := (box nat O).
+Extraction box1. (* let box1 = Obj.magic O *)
+
+(* applied constant, magic needed *)
+
+Definition idzarb := [b:bool][x:(if b then nat else bool)]x.
+Definition zarb := (idzarb true O).
+Extraction NoInline idzarb.
+Extraction zarb.
+(* let zarb = Obj.magic idzarb True (Obj.magic O) *)
+
+(** function of variable arity. *)
+(** Fun n = nat -> nat -> ... -> nat *)
+
+Fixpoint Fun [n:nat] : Set :=
+ Cases n of
+ O => nat
+ | (S n) => nat -> (Fun n)
+ end.
+
+Fixpoint Const [k,n:nat] : (Fun n) :=
+ <Fun>Cases n of
+ O => k
+ | (S n) => [p:nat](Const k n)
+ end.
+
+Fixpoint proj [k,n:nat] : (Fun n) :=
+ <Fun>Cases n of
+ O => O (* ou assert false ....*)
+ | (S n) => Cases k of
+ O => [x](Const x n)
+ | (S k) => [x](proj k n)
+ end
+ end.
+
+Definition test_proj := (proj (2) (4) (0) (1) (2) (3)).
+
+Eval Compute in test_proj.
+
+Recursive Extraction test_proj.
+
+
+
+(*** TO SUM UP: ***)
+
+
+Extraction "test_extraction.ml"
+ idnat id id' test2 test3 test4 test5 test6 test7
+ d d2 d3 d4 d5 d6 test8 id id' test9 test10 test11
+ test12 test13 test19 test20
+ nat sumbool_rect c Finite tree tree_size
+ test14 test15 eta_c test16 test17 test18 bidon tb fbidon fbidon2
+ fbidon2 test_0 test_1 eq eq_rect tp1 tp1bis Truc oups test24 loop
+ horibilis PropSet natbool zerotrue zeroTrue zeroprop test21 test22
+ test23 f f_prop f_arity f_normal
+ Boite boite1 boite2 test_boite
+ Box box1 zarb test_proj.
+
+
diff --git a/contrib7/field/Field.v b/contrib7/field/Field.v
new file mode 100644
index 00000000..f282e246
--- /dev/null
+++ b/contrib7/field/Field.v
@@ -0,0 +1,15 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Field.v,v 1.1.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+Require Export Field_Compl.
+Require Export Field_Theory.
+Require Export Field_Tactic.
+
+(* Command declarations are moved to the ML side *)
diff --git a/contrib7/field/Field_Compl.v b/contrib7/field/Field_Compl.v
new file mode 100644
index 00000000..2cc01038
--- /dev/null
+++ b/contrib7/field/Field_Compl.v
@@ -0,0 +1,62 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Field_Compl.v,v 1.2.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+Inductive listT [A:Type] : Type :=
+ nilT : (listT A) | consT : A->(listT A)->(listT A).
+
+Fixpoint appT [A:Type][l:(listT A)] : (listT A) -> (listT A) :=
+ [m:(listT A)]
+ Cases l of
+ | nilT => m
+ | (consT a l1) => (consT A a (appT A l1 m))
+ end.
+
+Inductive prodT [A,B:Type] : Type :=
+ pairT: A->B->(prodT A B).
+
+Definition assoc_2nd :=
+Fix assoc_2nd_rec
+ {assoc_2nd_rec
+ [A:Type;B:Set;eq_dec:(e1,e2:B){e1=e2}+{~e1=e2};lst:(listT (prodT A B))]
+ : B->A->A:=
+ [key:B;default:A]
+ Cases lst of
+ | nilT => default
+ | (consT (pairT v e) l) =>
+ (Cases (eq_dec e key) of
+ | (left _) => v
+ | (right _) => (assoc_2nd_rec A B eq_dec l key default)
+ end)
+ end}.
+
+Definition fstT [A,B:Type;c:(prodT A B)] :=
+ Cases c of
+ | (pairT a _) => a
+ end.
+
+Definition sndT [A,B:Type;c:(prodT A B)] :=
+ Cases c of
+ | (pairT _ a) => a
+ end.
+
+Definition mem :=
+Fix mem {mem [A:Set;eq_dec:(e1,e2:A){e1=e2}+{~e1=e2};a:A;l:(listT A)] : bool :=
+ Cases l of
+ | nilT => false
+ | (consT a1 l1) =>
+ Cases (eq_dec a a1) of
+ | (left _) => true
+ | (right _) => (mem A eq_dec a l1)
+ end
+ end}.
+
+Inductive field_rel_option [A:Type] : Type :=
+ | Field_None : (field_rel_option A)
+ | Field_Some : (A -> A -> A) -> (field_rel_option A).
diff --git a/contrib7/field/Field_Tactic.v b/contrib7/field/Field_Tactic.v
new file mode 100644
index 00000000..ffd2aad4
--- /dev/null
+++ b/contrib7/field/Field_Tactic.v
@@ -0,0 +1,397 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Field_Tactic.v,v 1.2.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+Require Ring.
+Require Export Field_Compl.
+Require Export Field_Theory.
+
+(**** Interpretation A --> ExprA ****)
+
+Recursive Tactic Definition MemAssoc var lvar :=
+ Match lvar With
+ | [(nilT ?)] -> false
+ | [(consT ? ?1 ?2)] ->
+ (Match ?1=var With
+ | [?1=?1] -> true
+ | _ -> (MemAssoc var ?2)).
+
+Recursive Tactic Definition SeekVarAux FT lvar trm :=
+ Let AT = Eval Cbv Beta Delta [A] Iota in (A FT)
+ And AzeroT = Eval Cbv Beta Delta [Azero] Iota in (Azero FT)
+ And AoneT = Eval Cbv Beta Delta [Aone] Iota in (Aone FT)
+ And AplusT = Eval Cbv Beta Delta [Aplus] Iota in (Aplus FT)
+ And AmultT = Eval Cbv Beta Delta [Amult] Iota in (Amult FT)
+ And AoppT = Eval Cbv Beta Delta [Aopp] Iota in (Aopp FT)
+ And AinvT = Eval Cbv Beta Delta [Ainv] Iota in (Ainv FT) In
+ Match trm With
+ | [(AzeroT)] -> lvar
+ | [(AoneT)] -> lvar
+ | [(AplusT ?1 ?2)] ->
+ Let l1 = (SeekVarAux FT lvar ?1) In
+ (SeekVarAux FT l1 ?2)
+ | [(AmultT ?1 ?2)] ->
+ Let l1 = (SeekVarAux FT lvar ?1) In
+ (SeekVarAux FT l1 ?2)
+ | [(AoppT ?1)] -> (SeekVarAux FT lvar ?1)
+ | [(AinvT ?1)] -> (SeekVarAux FT lvar ?1)
+ | [?1] ->
+ Let res = (MemAssoc ?1 lvar) In
+ Match res With
+ | [(true)] -> lvar
+ | [(false)] -> '(consT AT ?1 lvar).
+
+Tactic Definition SeekVar FT trm :=
+ Let AT = Eval Cbv Beta Delta [A] Iota in (A FT) In
+ (SeekVarAux FT '(nilT AT) trm).
+
+Recursive Tactic Definition NumberAux lvar cpt :=
+ Match lvar With
+ | [(nilT ?1)] -> '(nilT (prodT ?1 nat))
+ | [(consT ?1 ?2 ?3)] ->
+ Let l2 = (NumberAux ?3 '(S cpt)) In
+ '(consT (prodT ?1 nat) (pairT ?1 nat ?2 cpt) l2).
+
+Tactic Definition Number lvar := (NumberAux lvar O).
+
+Tactic Definition BuildVarList FT trm :=
+ Let lvar = (SeekVar FT trm) In
+ (Number lvar).
+V7only [
+(*Used by contrib Maple *)
+Tactic Definition build_var_list := BuildVarList.
+].
+
+Recursive Tactic Definition Assoc elt lst :=
+ Match lst With
+ | [(nilT ?)] -> Fail
+ | [(consT (prodT ? nat) (pairT ? nat ?1 ?2) ?3)] ->
+ Match elt= ?1 With
+ | [?1= ?1] -> ?2
+ | _ -> (Assoc elt ?3).
+
+Recursive Meta Definition interp_A FT lvar trm :=
+ Let AT = Eval Cbv Beta Delta [A] Iota in (A FT)
+ And AzeroT = Eval Cbv Beta Delta [Azero] Iota in (Azero FT)
+ And AoneT = Eval Cbv Beta Delta [Aone] Iota in (Aone FT)
+ And AplusT = Eval Cbv Beta Delta [Aplus] Iota in (Aplus FT)
+ And AmultT = Eval Cbv Beta Delta [Amult] Iota in (Amult FT)
+ And AoppT = Eval Cbv Beta Delta [Aopp] Iota in (Aopp FT)
+ And AinvT = Eval Cbv Beta Delta [Ainv] Iota in (Ainv FT) In
+ Match trm With
+ | [(AzeroT)] -> EAzero
+ | [(AoneT)] -> EAone
+ | [(AplusT ?1 ?2)] ->
+ Let e1 = (interp_A FT lvar ?1)
+ And e2 = (interp_A FT lvar ?2) In
+ '(EAplus e1 e2)
+ | [(AmultT ?1 ?2)] ->
+ Let e1 = (interp_A FT lvar ?1)
+ And e2 = (interp_A FT lvar ?2) In
+ '(EAmult e1 e2)
+ | [(AoppT ?1)] ->
+ Let e = (interp_A FT lvar ?1) In
+ '(EAopp e)
+ | [(AinvT ?1)] ->
+ Let e = (interp_A FT lvar ?1) In
+ '(EAinv e)
+ | [?1] ->
+ Let idx = (Assoc ?1 lvar) In
+ '(EAvar idx).
+
+(************************)
+(* Simplification *)
+(************************)
+
+(**** Generation of the multiplier ****)
+
+Recursive Tactic Definition Remove e l :=
+ Match l With
+ | [(nilT ?)] -> l
+ | [(consT ?1 e ?2)] -> ?2
+ | [(consT ?1 ?2 ?3)] ->
+ Let nl = (Remove e ?3) In
+ '(consT ?1 ?2 nl).
+
+Recursive Tactic Definition Union l1 l2 :=
+ Match l1 With
+ | [(nilT ?)] -> l2
+ | [(consT ?1 ?2 ?3)] ->
+ Let nl2 = (Remove ?2 l2) In
+ Let nl = (Union ?3 nl2) In
+ '(consT ?1 ?2 nl).
+
+Recursive Tactic Definition RawGiveMult trm :=
+ Match trm With
+ | [(EAinv ?1)] -> '(consT ExprA ?1 (nilT ExprA))
+ | [(EAopp ?1)] -> (RawGiveMult ?1)
+ | [(EAplus ?1 ?2)] ->
+ Let l1 = (RawGiveMult ?1)
+ And l2 = (RawGiveMult ?2) In
+ (Union l1 l2)
+ | [(EAmult ?1 ?2)] ->
+ Let l1 = (RawGiveMult ?1)
+ And l2 = (RawGiveMult ?2) In
+ Eval Compute in (appT ExprA l1 l2)
+ | _ -> '(nilT ExprA).
+
+Tactic Definition GiveMult trm :=
+ Let ltrm = (RawGiveMult trm) In
+ '(mult_of_list ltrm).
+
+(**** Associativity ****)
+
+Tactic Definition ApplyAssoc FT lvar trm :=
+ Let t=Eval Compute in (assoc trm) In
+ Match t=trm With
+ | [ ?1=?1 ] -> Idtac
+ | _ -> Rewrite <- (assoc_correct FT trm); Change (assoc trm) with t.
+
+(**** Distribution *****)
+
+Tactic Definition ApplyDistrib FT lvar trm :=
+ Let t=Eval Compute in (distrib trm) In
+ Match t=trm With
+ | [ ?1=?1 ] -> Idtac
+ | _ -> Rewrite <- (distrib_correct FT trm); Change (distrib trm) with t.
+
+(**** Multiplication by the inverse product ****)
+
+Tactic Definition GrepMult :=
+ Match Context With
+ | [ id: ~(interp_ExprA ? ? ?)= ? |- ?] -> id.
+
+Tactic Definition WeakReduce :=
+ Match Context With
+ | [|-[(interp_ExprA ?1 ?2 ?)]] ->
+ Cbv Beta Delta [interp_ExprA assoc_2nd eq_nat_dec mult_of_list ?1 ?2 A
+ Azero Aone Aplus Amult Aopp Ainv] Zeta Iota.
+
+Tactic Definition Multiply mul :=
+ Match Context With
+ | [|-(interp_ExprA ?1 ?2 ?3)=(interp_ExprA ?1 ?2 ?4)] ->
+ Let AzeroT = Eval Cbv Beta Delta [Azero ?1] Iota in (Azero ?1) In
+ Cut ~(interp_ExprA ?1 ?2 mul)=AzeroT;
+ [Intro;
+ Let id = GrepMult In
+ Apply (mult_eq ?1 ?3 ?4 mul ?2 id)
+ |WeakReduce;
+ Let AoneT = Eval Cbv Beta Delta [Aone ?1] Iota in (Aone ?1)
+ And AmultT = Eval Cbv Beta Delta [Amult ?1] Iota in (Amult ?1) In
+ Try (Match Context With
+ | [|-[(AmultT ? AoneT)]] -> Rewrite (AmultT_1r ?1));Clear ?1 ?2].
+
+Tactic Definition ApplyMultiply FT lvar trm :=
+ Let t=Eval Compute in (multiply trm) In
+ Match t=trm With
+ | [ ?1=?1 ] -> Idtac
+ | _ -> Rewrite <- (multiply_correct FT trm); Change (multiply trm) with t.
+
+(**** Permutations and simplification ****)
+
+Tactic Definition ApplyInverse mul FT lvar trm :=
+ Let t=Eval Compute in (inverse_simplif mul trm) In
+ Match t=trm With
+ | [ ?1=?1 ] -> Idtac
+ | _ -> Rewrite <- (inverse_correct FT trm mul);
+ [Change (inverse_simplif mul trm) with t|Assumption].
+(**** Inverse test ****)
+
+Tactic Definition StrongFail tac := First [tac|Fail 2].
+
+Recursive Tactic Definition InverseTestAux FT trm :=
+ Let AplusT = Eval Cbv Beta Delta [Aplus] Iota in (Aplus FT)
+ And AmultT = Eval Cbv Beta Delta [Amult] Iota in (Amult FT)
+ And AoppT = Eval Cbv Beta Delta [Aopp] Iota in (Aopp FT)
+ And AinvT = Eval Cbv Beta Delta [Ainv] Iota in (Ainv FT) In
+ Match trm With
+ | [(AinvT ?)] -> Fail 1
+ | [(AoppT ?1)] -> StrongFail ((InverseTestAux FT ?1);Idtac)
+ | [(AplusT ?1 ?2)] ->
+ StrongFail ((InverseTestAux FT ?1);(InverseTestAux FT ?2))
+ | [(AmultT ?1 ?2)] ->
+ StrongFail ((InverseTestAux FT ?1);(InverseTestAux FT ?2))
+ | _ -> Idtac.
+
+Tactic Definition InverseTest FT :=
+ Let AplusT = Eval Cbv Beta Delta [Aplus] Iota in (Aplus FT) In
+ Match Context With
+ | [|- ?1=?2] -> (InverseTestAux FT '(AplusT ?1 ?2)).
+
+(**** Field itself ****)
+
+Tactic Definition ApplySimplif sfun :=
+ (Match Context With
+ | [|- (interp_ExprA ?1 ?2 ?3)=(interp_ExprA ? ? ?)] ->
+ (sfun ?1 ?2 ?3));
+ (Match Context With
+ | [|- (interp_ExprA ? ? ?)=(interp_ExprA ?1 ?2 ?3)] ->
+ (sfun ?1 ?2 ?3)).
+
+Tactic Definition Unfolds FT :=
+ (Match Eval Cbv Beta Delta [Aminus] Iota in (Aminus FT) With
+ | [(Field_Some ? ?1)] -> Unfold ?1
+ | _ -> Idtac);
+ (Match Eval Cbv Beta Delta [Adiv] Iota in (Adiv FT) With
+ | [(Field_Some ? ?1)] -> Unfold ?1
+ | _ -> Idtac).
+
+Tactic Definition Reduce FT :=
+ Let AzeroT = Eval Cbv Beta Delta [Azero] Iota in (Azero FT)
+ And AoneT = Eval Cbv Beta Delta [Aone] Iota in (Aone FT)
+ And AplusT = Eval Cbv Beta Delta [Aplus] Iota in (Aplus FT)
+ And AmultT = Eval Cbv Beta Delta [Amult] Iota in (Amult FT)
+ And AoppT = Eval Cbv Beta Delta [Aopp] Iota in (Aopp FT)
+ And AinvT = Eval Cbv Beta Delta [Ainv] Iota in (Ainv FT) In
+ Cbv Beta Delta -[AzeroT AoneT AplusT AmultT AoppT AinvT] Zeta Iota
+ Orelse Compute.
+
+Recursive Tactic Definition Field_Gen_Aux FT :=
+ Let AplusT = Eval Cbv Beta Delta [Aplus] Iota in (Aplus FT) In
+ Match Context With
+ | [|- ?1=?2] ->
+ Let lvar = (BuildVarList FT '(AplusT ?1 ?2)) In
+ Let trm1 = (interp_A FT lvar ?1)
+ And trm2 = (interp_A FT lvar ?2) In
+ Let mul = (GiveMult '(EAplus trm1 trm2)) In
+ Cut [ft:=FT][vm:=lvar](interp_ExprA ft vm trm1)=(interp_ExprA ft vm trm2);
+ [Compute;Auto
+ |Intros ft vm;(ApplySimplif ApplyDistrib);(ApplySimplif ApplyAssoc);
+ (Multiply mul);[(ApplySimplif ApplyMultiply);
+ (ApplySimplif (ApplyInverse mul));
+ (Let id = GrepMult In Clear id);WeakReduce;Clear ft vm;
+ First [(InverseTest FT);Ring|(Field_Gen_Aux FT)]|Idtac]].
+
+Tactic Definition Field_Gen FT :=
+ Unfolds FT;((InverseTest FT);Ring) Orelse (Field_Gen_Aux FT).
+V7only [Tactic Definition field_gen := Field_Gen.].
+
+(*****************************)
+(* Term Simplification *)
+(*****************************)
+
+(**** Minus and division expansions ****)
+
+Meta Definition InitExp FT trm :=
+ Let e =
+ (Match Eval Cbv Beta Delta [Aminus] Iota in (Aminus FT) With
+ | [(Field_Some ? ?1)] -> Eval Cbv Beta Delta [?1] in trm
+ | _ -> trm) In
+ Match Eval Cbv Beta Delta [Adiv] Iota in (Adiv FT) With
+ | [(Field_Some ? ?1)] -> Eval Cbv Beta Delta [?1] in e
+ | _ -> e.
+V7only [
+(*Used by contrib Maple *)
+Tactic Definition init_exp := InitExp.
+].
+
+(**** Inverses simplification ****)
+
+Recursive Meta Definition SimplInv trm:=
+ Match trm With
+ | [(EAplus ?1 ?2)] ->
+ Let e1 = (SimplInv ?1)
+ And e2 = (SimplInv ?2) In
+ '(EAplus e1 e2)
+ | [(EAmult ?1 ?2)] ->
+ Let e1 = (SimplInv ?1)
+ And e2 = (SimplInv ?2) In
+ '(EAmult e1 e2)
+ | [(EAopp ?1)] -> Let e = (SimplInv ?1) In '(EAopp e)
+ | [(EAinv ?1)] -> (SimplInvAux ?1)
+ | [?1] -> ?1
+And SimplInvAux trm :=
+ Match trm With
+ | [(EAinv ?1)] -> (SimplInv ?1)
+ | [(EAmult ?1 ?2)] ->
+ Let e1 = (SimplInv '(EAinv ?1))
+ And e2 = (SimplInv '(EAinv ?2)) In
+ '(EAmult e1 e2)
+ | [?1] -> Let e = (SimplInv ?1) In '(EAinv e).
+
+(**** Monom simplification ****)
+
+Recursive Meta Definition Map fcn lst :=
+ Match lst With
+ | [(nilT ?)] -> lst
+ | [(consT ?1 ?2 ?3)] ->
+ Let r = (fcn ?2)
+ And t = (Map fcn ?3) In
+ '(consT ?1 r t).
+
+Recursive Meta Definition BuildMonomAux lst trm :=
+ Match lst With
+ | [(nilT ?)] -> Eval Compute in (assoc trm)
+ | [(consT ? ?1 ?2)] -> BuildMonomAux ?2 '(EAmult trm ?1).
+
+Recursive Meta Definition BuildMonom lnum lden :=
+ Let ildn = (Map (Fun e -> '(EAinv e)) lden) In
+ Let ltot = Eval Compute in (appT ExprA lnum ildn) In
+ Let trm = (BuildMonomAux ltot EAone) In
+ Match trm With
+ | [(EAmult ? ?1)] -> ?1
+ | [?1] -> ?1.
+
+Recursive Meta Definition SimplMonomAux lnum lden trm :=
+ Match trm With
+ | [(EAmult (EAinv ?1) ?2)] ->
+ Let mma = (MemAssoc ?1 lnum) In
+ (Match mma With
+ | [true] ->
+ Let newlnum = (Remove ?1 lnum) In SimplMonomAux newlnum lden ?2
+ | [false] -> SimplMonomAux lnum '(consT ExprA ?1 lden) ?2)
+ | [(EAmult ?1 ?2)] ->
+ Let mma = (MemAssoc ?1 lden) In
+ (Match mma With
+ | [true] ->
+ Let newlden = (Remove ?1 lden) In SimplMonomAux lnum newlden ?2
+ | [false] -> SimplMonomAux '(consT ExprA ?1 lnum) lden ?2)
+ | [(EAinv ?1)] ->
+ Let mma = (MemAssoc ?1 lnum) In
+ (Match mma With
+ | [true] ->
+ Let newlnum = (Remove ?1 lnum) In BuildMonom newlnum lden
+ | [false] -> BuildMonom lnum '(consT ExprA ?1 lden))
+ | [?1] ->
+ Let mma = (MemAssoc ?1 lden) In
+ (Match mma With
+ | [true] ->
+ Let newlden = (Remove ?1 lden) In BuildMonom lnum newlden
+ | [false] -> BuildMonom '(consT ExprA ?1 lnum) lden).
+
+Meta Definition SimplMonom trm :=
+ SimplMonomAux '(nilT ExprA) '(nilT ExprA) trm.
+
+Recursive Meta Definition SimplAllMonoms trm :=
+ Match trm With
+ | [(EAplus ?1 ?2)] ->
+ Let e1 = (SimplMonom ?1)
+ And e2 = (SimplAllMonoms ?2) In
+ '(EAplus e1 e2)
+ | [?1] -> SimplMonom ?1.
+
+(**** Associativity and distribution ****)
+
+Meta Definition AssocDistrib trm := Eval Compute in (assoc (distrib trm)).
+
+(**** The tactic Field_Term ****)
+
+Tactic Definition EvalWeakReduce trm :=
+ Eval Cbv Beta Delta [interp_ExprA assoc_2nd eq_nat_dec mult_of_list A Azero
+ Aone Aplus Amult Aopp Ainv] Zeta Iota in trm.
+
+Tactic Definition Field_Term FT exp :=
+ Let newexp = (InitExp FT exp) In
+ Let lvar = (BuildVarList FT newexp) In
+ Let trm = (interp_A FT lvar newexp) In
+ Let tma = Eval Compute in (assoc trm) In
+ Let tsmp = (SimplAllMonoms (AssocDistrib (SimplAllMonoms
+ (SimplInv tma)))) In
+ Let trep = (EvalWeakReduce '(interp_ExprA FT lvar tsmp)) In
+ Replace exp with trep;[Ring trep|Field_Gen FT].
diff --git a/contrib7/field/Field_Theory.v b/contrib7/field/Field_Theory.v
new file mode 100644
index 00000000..3ba2fbc0
--- /dev/null
+++ b/contrib7/field/Field_Theory.v
@@ -0,0 +1,612 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Field_Theory.v,v 1.2.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+Require Peano_dec.
+Require Ring.
+Require Field_Compl.
+
+Record Field_Theory : Type :=
+{ A : Type;
+ Aplus : A -> A -> A;
+ Amult : A -> A -> A;
+ Aone : A;
+ Azero : A;
+ Aopp : A -> A;
+ Aeq : A -> A -> bool;
+ Ainv : A -> A;
+ Aminus : (field_rel_option A);
+ Adiv : (field_rel_option A);
+ RT : (Ring_Theory Aplus Amult Aone Azero Aopp Aeq);
+ Th_inv_def : (n:A)~(n=Azero)->(Amult (Ainv n) n)=Aone
+}.
+
+(* The reflexion structure *)
+Inductive ExprA : Set :=
+| EAzero : ExprA
+| EAone : ExprA
+| EAplus : ExprA -> ExprA -> ExprA
+| EAmult : ExprA -> ExprA -> ExprA
+| EAopp : ExprA -> ExprA
+| EAinv : ExprA -> ExprA
+| EAvar : nat -> ExprA.
+
+(**** Decidability of equality ****)
+
+Lemma eqExprA_O:(e1,e2:ExprA){e1=e2}+{~e1=e2}.
+Proof.
+ Double Induction e1 e2;Try Intros;
+ Try (Left;Reflexivity) Orelse Try (Right;Discriminate).
+ Elim (H1 e0);Intro y;Elim (H2 e);Intro y0;
+ Try (Left; Rewrite y; Rewrite y0;Auto)
+ Orelse (Right;Red;Intro;Inversion H3;Auto).
+ Elim (H1 e0);Intro y;Elim (H2 e);Intro y0;
+ Try (Left; Rewrite y; Rewrite y0;Auto)
+ Orelse (Right;Red;Intro;Inversion H3;Auto).
+ Elim (H0 e);Intro y.
+ Left; Rewrite y; Auto.
+ Right;Red; Intro;Inversion H1;Auto.
+ Elim (H0 e);Intro y.
+ Left; Rewrite y; Auto.
+ Right;Red; Intro;Inversion H1;Auto.
+ Elim (eq_nat_dec n n0);Intro y.
+ Left; Rewrite y; Auto.
+ Right;Red;Intro;Inversion H;Auto.
+Defined.
+
+Definition eq_nat_dec := Eval Compute in Peano_dec.eq_nat_dec.
+Definition eqExprA := Eval Compute in eqExprA_O.
+
+(**** Generation of the multiplier ****)
+
+Fixpoint mult_of_list [e:(listT ExprA)]: ExprA :=
+ Cases e of
+ | nilT => EAone
+ | (consT e1 l1) => (EAmult e1 (mult_of_list l1))
+ end.
+
+Section Theory_of_fields.
+
+Variable T : Field_Theory.
+
+Local AT := (A T).
+Local AplusT := (Aplus T).
+Local AmultT := (Amult T).
+Local AoneT := (Aone T).
+Local AzeroT := (Azero T).
+Local AoppT := (Aopp T).
+Local AeqT := (Aeq T).
+Local AinvT := (Ainv T).
+Local RTT := (RT T).
+Local Th_inv_defT := (Th_inv_def T).
+
+Add Abstract Ring (A T) (Aplus T) (Amult T) (Aone T) (Azero T) (Aopp T)
+ (Aeq T) (RT T).
+
+Add Abstract Ring AT AplusT AmultT AoneT AzeroT AoppT AeqT RTT.
+
+(***************************)
+(* Lemmas to be used *)
+(***************************)
+
+Lemma AplusT_sym:(r1,r2:AT)(AplusT r1 r2)=(AplusT r2 r1).
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AplusT_assoc:(r1,r2,r3:AT)(AplusT (AplusT r1 r2) r3)=
+ (AplusT r1 (AplusT r2 r3)).
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AmultT_sym:(r1,r2:AT)(AmultT r1 r2)=(AmultT r2 r1).
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AmultT_assoc:(r1,r2,r3:AT)(AmultT (AmultT r1 r2) r3)=
+ (AmultT r1 (AmultT r2 r3)).
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AplusT_Ol:(r:AT)(AplusT AzeroT r)=r.
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AmultT_1l:(r:AT)(AmultT AoneT r)=r.
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AplusT_AoppT_r:(r:AT)(AplusT r (AoppT r))=AzeroT.
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma AmultT_AplusT_distr:(r1,r2,r3:AT)(AmultT r1 (AplusT r2 r3))=
+ (AplusT (AmultT r1 r2) (AmultT r1 r3)).
+Proof.
+ Intros;Ring.
+Save.
+
+Lemma r_AplusT_plus:(r,r1,r2:AT)(AplusT r r1)=(AplusT r r2)->r1=r2.
+Proof.
+ Intros; Transitivity (AplusT (AplusT (AoppT r) r) r1).
+ Ring.
+ Transitivity (AplusT (AplusT (AoppT r) r) r2).
+ Repeat Rewrite -> AplusT_assoc; Rewrite <- H; Reflexivity.
+ Ring.
+Save.
+
+Lemma r_AmultT_mult:
+ (r,r1,r2:AT)(AmultT r r1)=(AmultT r r2)->~r=AzeroT->r1=r2.
+Proof.
+ Intros; Transitivity (AmultT (AmultT (AinvT r) r) r1).
+ Rewrite Th_inv_defT;[Symmetry; Apply AmultT_1l;Auto|Auto].
+ Transitivity (AmultT (AmultT (AinvT r) r) r2).
+ Repeat Rewrite AmultT_assoc; Rewrite H; Trivial.
+ Rewrite Th_inv_defT;[Apply AmultT_1l;Auto|Auto].
+Save.
+
+Lemma AmultT_Or:(r:AT) (AmultT r AzeroT)=AzeroT.
+Proof.
+ Intro; Ring.
+Save.
+
+Lemma AmultT_Ol:(r:AT)(AmultT AzeroT r)=AzeroT.
+Proof.
+ Intro; Ring.
+Save.
+
+Lemma AmultT_1r:(r:AT)(AmultT r AoneT)=r.
+Proof.
+ Intro; Ring.
+Save.
+
+Lemma AinvT_r:(r:AT)~r=AzeroT->(AmultT r (AinvT r))=AoneT.
+Proof.
+ Intros; Rewrite -> AmultT_sym; Apply Th_inv_defT; Auto.
+Save.
+
+Lemma without_div_O_contr:
+ (r1,r2:AT)~(AmultT r1 r2)=AzeroT ->~r1=AzeroT/\~r2=AzeroT.
+Proof.
+ Intros r1 r2 H; Split; Red; Intro; Apply H; Rewrite H0; Ring.
+Save.
+
+(************************)
+(* Interpretation *)
+(************************)
+
+(**** ExprA --> A ****)
+
+Fixpoint interp_ExprA [lvar:(listT (prodT AT nat));e:ExprA] : AT :=
+ Cases e of
+ | EAzero => AzeroT
+ | EAone => AoneT
+ | (EAplus e1 e2) => (AplusT (interp_ExprA lvar e1) (interp_ExprA lvar e2))
+ | (EAmult e1 e2) => (AmultT (interp_ExprA lvar e1) (interp_ExprA lvar e2))
+ | (EAopp e) => ((Aopp T) (interp_ExprA lvar e))
+ | (EAinv e) => ((Ainv T) (interp_ExprA lvar e))
+ | (EAvar n) => (assoc_2nd AT nat eq_nat_dec lvar n AzeroT)
+ end.
+
+(************************)
+(* Simplification *)
+(************************)
+
+(**** Associativity ****)
+
+Definition merge_mult :=
+ Fix merge_mult {merge_mult [e1:ExprA] : ExprA -> ExprA :=
+ [e2:ExprA]Cases e1 of
+ | (EAmult t1 t2) =>
+ Cases t2 of
+ | (EAmult t2 t3) => (EAmult t1 (EAmult t2 (merge_mult t3 e2)))
+ | _ => (EAmult t1 (EAmult t2 e2))
+ end
+ | _ => (EAmult e1 e2)
+ end}.
+
+Fixpoint assoc_mult [e:ExprA] : ExprA :=
+ Cases e of
+ | (EAmult e1 e3) =>
+ Cases e1 of
+ | (EAmult e1 e2) =>
+ (merge_mult (merge_mult (assoc_mult e1) (assoc_mult e2))
+ (assoc_mult e3))
+ | _ => (EAmult e1 (assoc_mult e3))
+ end
+ | _ => e
+ end.
+
+Definition merge_plus :=
+ Fix merge_plus {merge_plus [e1:ExprA]:ExprA->ExprA:=
+ [e2:ExprA]Cases e1 of
+ | (EAplus t1 t2) =>
+ Cases t2 of
+ | (EAplus t2 t3) => (EAplus t1 (EAplus t2 (merge_plus t3 e2)))
+ | _ => (EAplus t1 (EAplus t2 e2))
+ end
+ | _ => (EAplus e1 e2)
+ end}.
+
+Fixpoint assoc [e:ExprA] : ExprA :=
+ Cases e of
+ | (EAplus e1 e3) =>
+ Cases e1 of
+ | (EAplus e1 e2) =>
+ (merge_plus (merge_plus (assoc e1) (assoc e2)) (assoc e3))
+ | _ => (EAplus (assoc_mult e1) (assoc e3))
+ end
+ | _ => (assoc_mult e)
+ end.
+
+Lemma merge_mult_correct1:
+ (e1,e2,e3:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (merge_mult (EAmult e1 e2) e3))=
+ (interp_ExprA lvar (EAmult e1 (merge_mult e2 e3))).
+Proof.
+Intros e1 e2;Generalize e1;Generalize e2;Clear e1 e2.
+Induction e2;Auto;Intros.
+Unfold 1 merge_mult;Fold merge_mult;
+ Unfold 2 interp_ExprA;Fold interp_ExprA;
+ Rewrite (H0 e e3 lvar);
+ Unfold 1 interp_ExprA;Fold interp_ExprA;
+ Unfold 5 interp_ExprA;Fold interp_ExprA;Auto.
+Save.
+
+Lemma merge_mult_correct:
+ (e1,e2:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (merge_mult e1 e2))=
+ (interp_ExprA lvar (EAmult e1 e2)).
+Proof.
+Induction e1;Auto;Intros.
+Elim e0;Try (Intros;Simpl;Ring).
+Unfold interp_ExprA in H2;Fold interp_ExprA in H2;
+ Cut (AmultT (interp_ExprA lvar e2) (AmultT (interp_ExprA lvar e4)
+ (AmultT (interp_ExprA lvar e) (interp_ExprA lvar e3))))=
+ (AmultT (AmultT (AmultT (interp_ExprA lvar e) (interp_ExprA lvar e4))
+ (interp_ExprA lvar e2)) (interp_ExprA lvar e3)).
+Intro H3;Rewrite H3;Rewrite <-H2;
+ Rewrite merge_mult_correct1;Simpl;Ring.
+Ring.
+Save.
+
+Lemma assoc_mult_correct1:(e1,e2:ExprA)(lvar:(listT (prodT AT nat)))
+ (AmultT (interp_ExprA lvar (assoc_mult e1))
+ (interp_ExprA lvar (assoc_mult e2)))=
+ (interp_ExprA lvar (assoc_mult (EAmult e1 e2))).
+Proof.
+Induction e1;Auto;Intros.
+Rewrite <-(H e0 lvar);Simpl;Rewrite merge_mult_correct;Simpl;
+ Rewrite merge_mult_correct;Simpl;Auto.
+Save.
+
+Lemma assoc_mult_correct:
+ (e:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (assoc_mult e))=(interp_ExprA lvar e).
+Proof.
+Induction e;Auto;Intros.
+Elim e0;Intros.
+Intros;Simpl;Ring.
+Simpl;Rewrite (AmultT_1l (interp_ExprA lvar (assoc_mult e1)));
+ Rewrite (AmultT_1l (interp_ExprA lvar e1)); Apply H0.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite merge_mult_correct;Simpl;Rewrite merge_mult_correct;Simpl;
+ Rewrite AmultT_assoc;Rewrite assoc_mult_correct1;Rewrite H2;Simpl;
+ Rewrite <-assoc_mult_correct1 in H1;
+ Unfold 3 interp_ExprA in H1;Fold interp_ExprA in H1;
+ Rewrite (H0 lvar) in H1;
+ Rewrite (AmultT_sym (interp_ExprA lvar e3) (interp_ExprA lvar e1));
+ Rewrite <-AmultT_assoc;Rewrite H1;Rewrite AmultT_assoc;Ring.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite (H0 lvar);Auto.
+Save.
+
+Lemma merge_plus_correct1:
+ (e1,e2,e3:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (merge_plus (EAplus e1 e2) e3))=
+ (interp_ExprA lvar (EAplus e1 (merge_plus e2 e3))).
+Proof.
+Intros e1 e2;Generalize e1;Generalize e2;Clear e1 e2.
+Induction e2;Auto;Intros.
+Unfold 1 merge_plus;Fold merge_plus;
+ Unfold 2 interp_ExprA;Fold interp_ExprA;
+ Rewrite (H0 e e3 lvar);
+ Unfold 1 interp_ExprA;Fold interp_ExprA;
+ Unfold 5 interp_ExprA;Fold interp_ExprA;Auto.
+Save.
+
+Lemma merge_plus_correct:
+ (e1,e2:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (merge_plus e1 e2))=
+ (interp_ExprA lvar (EAplus e1 e2)).
+Proof.
+Induction e1;Auto;Intros.
+Elim e0;Try Intros;Try (Simpl;Ring).
+Unfold interp_ExprA in H2;Fold interp_ExprA in H2;
+ Cut (AplusT (interp_ExprA lvar e2) (AplusT (interp_ExprA lvar e4)
+ (AplusT (interp_ExprA lvar e) (interp_ExprA lvar e3))))=
+ (AplusT (AplusT (AplusT (interp_ExprA lvar e) (interp_ExprA lvar e4))
+ (interp_ExprA lvar e2)) (interp_ExprA lvar e3)).
+Intro H3;Rewrite H3;Rewrite <-H2;Rewrite merge_plus_correct1;Simpl;Ring.
+Ring.
+Save.
+
+Lemma assoc_plus_correct:(e1,e2:ExprA)(lvar:(listT (prodT AT nat)))
+ (AplusT (interp_ExprA lvar (assoc e1)) (interp_ExprA lvar (assoc e2)))=
+ (interp_ExprA lvar (assoc (EAplus e1 e2))).
+Proof.
+Induction e1;Auto;Intros.
+Rewrite <-(H e0 lvar);Simpl;Rewrite merge_plus_correct;Simpl;
+ Rewrite merge_plus_correct;Simpl;Auto.
+Save.
+
+Lemma assoc_correct:
+ (e:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (assoc e))=(interp_ExprA lvar e).
+Proof.
+Induction e;Auto;Intros.
+Elim e0;Intros.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite merge_plus_correct;Simpl;Rewrite merge_plus_correct;
+ Simpl;Rewrite AplusT_assoc;Rewrite assoc_plus_correct;Rewrite H2;
+ Simpl;Apply (r_AplusT_plus (interp_ExprA lvar (assoc e1))
+ (AplusT (interp_ExprA lvar (assoc e2))
+ (AplusT (interp_ExprA lvar e3) (interp_ExprA lvar e1)))
+ (AplusT (AplusT (interp_ExprA lvar e2) (interp_ExprA lvar e3))
+ (interp_ExprA lvar e1)));Rewrite <-AplusT_assoc;
+ Rewrite (AplusT_sym (interp_ExprA lvar (assoc e1))
+ (interp_ExprA lvar (assoc e2)));
+ Rewrite assoc_plus_correct;Rewrite H1;Simpl;Rewrite (H0 lvar);
+ Rewrite <-(AplusT_assoc (AplusT (interp_ExprA lvar e2)
+ (interp_ExprA lvar e1))
+ (interp_ExprA lvar e3) (interp_ExprA lvar e1));
+ Rewrite (AplusT_assoc (interp_ExprA lvar e2) (interp_ExprA lvar e1)
+ (interp_ExprA lvar e3));
+ Rewrite (AplusT_sym (interp_ExprA lvar e1) (interp_ExprA lvar e3));
+ Rewrite <-(AplusT_assoc (interp_ExprA lvar e2) (interp_ExprA lvar e3)
+ (interp_ExprA lvar e1));Apply AplusT_sym.
+Unfold assoc;Fold assoc;Unfold interp_ExprA;Fold interp_ExprA;
+ Rewrite assoc_mult_correct;Rewrite (H0 lvar);Simpl;Auto.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite (H0 lvar);Auto.
+Simpl;Rewrite (H0 lvar);Auto.
+Unfold assoc;Fold assoc;Unfold interp_ExprA;Fold interp_ExprA;
+ Rewrite assoc_mult_correct;Simpl;Auto.
+Save.
+
+(**** Distribution *****)
+
+Fixpoint distrib_EAopp [e:ExprA] : ExprA :=
+ Cases e of
+ | (EAplus e1 e2) => (EAplus (distrib_EAopp e1) (distrib_EAopp e2))
+ | (EAmult e1 e2) => (EAmult (distrib_EAopp e1) (distrib_EAopp e2))
+ | (EAopp e) => (EAmult (EAopp EAone) (distrib_EAopp e))
+ | e => e
+ end.
+
+Definition distrib_mult_right :=
+ Fix distrib_mult_right {distrib_mult_right [e1:ExprA]:ExprA->ExprA:=
+ [e2:ExprA]Cases e1 of
+ | (EAplus t1 t2) =>
+ (EAplus (distrib_mult_right t1 e2) (distrib_mult_right t2 e2))
+ | _ => (EAmult e1 e2)
+ end}.
+
+Fixpoint distrib_mult_left [e1:ExprA] : ExprA->ExprA :=
+ [e2:ExprA]
+ Cases e1 of
+ | (EAplus t1 t2) =>
+ (EAplus (distrib_mult_left t1 e2) (distrib_mult_left t2 e2))
+ | _ => (distrib_mult_right e2 e1)
+ end.
+
+Fixpoint distrib_main [e:ExprA] : ExprA :=
+ Cases e of
+ | (EAmult e1 e2) => (distrib_mult_left (distrib_main e1) (distrib_main e2))
+ | (EAplus e1 e2) => (EAplus (distrib_main e1) (distrib_main e2))
+ | (EAopp e) => (EAopp (distrib_main e))
+ | _ => e
+ end.
+
+Definition distrib [e:ExprA] : ExprA := (distrib_main (distrib_EAopp e)).
+
+Lemma distrib_mult_right_correct:
+ (e1,e2:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (distrib_mult_right e1 e2))=
+ (AmultT (interp_ExprA lvar e1) (interp_ExprA lvar e2)).
+Proof.
+Induction e1;Try Intros;Simpl;Auto.
+Rewrite AmultT_sym;Rewrite AmultT_AplusT_distr;
+ Rewrite (H e2 lvar);Rewrite (H0 e2 lvar);Ring.
+Save.
+
+Lemma distrib_mult_left_correct:
+ (e1,e2:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (distrib_mult_left e1 e2))=
+ (AmultT (interp_ExprA lvar e1) (interp_ExprA lvar e2)).
+Proof.
+Induction e1;Try Intros;Simpl.
+Rewrite AmultT_Ol;Rewrite distrib_mult_right_correct;Simpl;Apply AmultT_Or.
+Rewrite distrib_mult_right_correct;Simpl;
+ Apply AmultT_sym.
+Rewrite AmultT_sym;
+ Rewrite (AmultT_AplusT_distr (interp_ExprA lvar e2) (interp_ExprA lvar e)
+ (interp_ExprA lvar e0));
+ Rewrite (AmultT_sym (interp_ExprA lvar e2) (interp_ExprA lvar e));
+ Rewrite (AmultT_sym (interp_ExprA lvar e2) (interp_ExprA lvar e0));
+ Rewrite (H e2 lvar);Rewrite (H0 e2 lvar);Auto.
+Rewrite distrib_mult_right_correct;Simpl;Apply AmultT_sym.
+Rewrite distrib_mult_right_correct;Simpl;Apply AmultT_sym.
+Rewrite distrib_mult_right_correct;Simpl;Apply AmultT_sym.
+Rewrite distrib_mult_right_correct;Simpl;Apply AmultT_sym.
+Save.
+
+Lemma distrib_correct:
+ (e:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (distrib e))=(interp_ExprA lvar e).
+Proof.
+Induction e;Intros;Auto.
+Simpl;Rewrite <- (H lvar);Rewrite <- (H0 lvar); Unfold distrib;Simpl;Auto.
+Simpl;Rewrite <- (H lvar);Rewrite <- (H0 lvar); Unfold distrib;Simpl;
+ Apply distrib_mult_left_correct.
+Simpl;Fold AoppT;Rewrite <- (H lvar);Unfold distrib;Simpl;
+ Rewrite distrib_mult_right_correct;
+ Simpl;Fold AoppT;Ring.
+Save.
+
+(**** Multiplication by the inverse product ****)
+
+Lemma mult_eq:
+ (e1,e2,a:ExprA)(lvar:(listT (prodT AT nat)))
+ ~((interp_ExprA lvar a)=AzeroT)->
+ (interp_ExprA lvar (EAmult a e1))=(interp_ExprA lvar (EAmult a e2))->
+ (interp_ExprA lvar e1)=(interp_ExprA lvar e2).
+Proof.
+ Simpl;Intros;
+ Apply (r_AmultT_mult (interp_ExprA lvar a) (interp_ExprA lvar e1)
+ (interp_ExprA lvar e2));Assumption.
+Save.
+
+Fixpoint multiply_aux [a,e:ExprA] : ExprA :=
+ Cases e of
+ | (EAplus e1 e2) =>
+ (EAplus (EAmult a e1) (multiply_aux a e2))
+ | _ => (EAmult a e)
+ end.
+
+Definition multiply [e:ExprA] : ExprA :=
+ Cases e of
+ | (EAmult a e1) => (multiply_aux a e1)
+ | _ => e
+ end.
+
+Lemma multiply_aux_correct:
+ (a,e:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (multiply_aux a e))=
+ (AmultT (interp_ExprA lvar a) (interp_ExprA lvar e)).
+Proof.
+Induction e;Simpl;Intros;Try (Rewrite merge_mult_correct);Auto.
+ Simpl;Rewrite (H0 lvar);Ring.
+Save.
+
+Lemma multiply_correct:
+ (e:ExprA)(lvar:(listT (prodT AT nat)))
+ (interp_ExprA lvar (multiply e))=(interp_ExprA lvar e).
+Proof.
+ Induction e;Simpl;Auto.
+ Intros;Apply multiply_aux_correct.
+Save.
+
+(**** Permutations and simplification ****)
+
+Fixpoint monom_remove [a,m:ExprA] : ExprA :=
+ Cases m of
+ | (EAmult m0 m1) =>
+ (Cases (eqExprA m0 (EAinv a)) of
+ | (left _) => m1
+ | (right _) => (EAmult m0 (monom_remove a m1))
+ end)
+ | _ =>
+ (Cases (eqExprA m (EAinv a)) of
+ | (left _) => EAone
+ | (right _) => (EAmult a m)
+ end)
+ end.
+
+Definition monom_simplif_rem :=
+ Fix monom_simplif_rem {monom_simplif_rem/1:ExprA->ExprA->ExprA:=
+ [a,m:ExprA]
+ Cases a of
+ | (EAmult a0 a1) => (monom_simplif_rem a1 (monom_remove a0 m))
+ | _ => (monom_remove a m)
+ end}.
+
+Definition monom_simplif [a,m:ExprA] : ExprA :=
+ Cases m of
+ | (EAmult a' m') =>
+ (Cases (eqExprA a a') of
+ | (left _) => (monom_simplif_rem a m')
+ | (right _) => m
+ end)
+ | _ => m
+ end.
+
+Fixpoint inverse_simplif [a,e:ExprA] : ExprA :=
+ Cases e of
+ | (EAplus e1 e2) => (EAplus (monom_simplif a e1) (inverse_simplif a e2))
+ | _ => (monom_simplif a e)
+ end.
+
+Lemma monom_remove_correct:(e,a:ExprA)
+ (lvar:(listT (prodT AT nat)))~((interp_ExprA lvar a)=AzeroT)->
+ (interp_ExprA lvar (monom_remove a e))=
+ (AmultT (interp_ExprA lvar a) (interp_ExprA lvar e)).
+Proof.
+Induction e; Intros.
+Simpl;Case (eqExprA EAzero (EAinv a));Intros;[Inversion e0|Simpl;Trivial].
+Simpl;Case (eqExprA EAone (EAinv a));Intros;[Inversion e0|Simpl;Trivial].
+Simpl;Case (eqExprA (EAplus e0 e1) (EAinv a));Intros;[Inversion e2|
+ Simpl;Trivial].
+Simpl;Case (eqExprA e0 (EAinv a));Intros.
+Rewrite e2;Simpl;Fold AinvT.
+Rewrite <-(AmultT_assoc (interp_ExprA lvar a) (AinvT (interp_ExprA lvar a))
+ (interp_ExprA lvar e1));
+ Rewrite AinvT_r;[Ring|Assumption].
+Simpl;Rewrite H0;Auto; Ring.
+Simpl;Fold AoppT;Case (eqExprA (EAopp e0) (EAinv a));Intros;[Inversion e1|
+ Simpl;Trivial].
+Unfold monom_remove;Case (eqExprA (EAinv e0) (EAinv a));Intros.
+Case (eqExprA e0 a);Intros.
+Rewrite e2;Simpl;Fold AinvT;Rewrite AinvT_r;Auto.
+Inversion e1;Simpl;ElimType False;Auto.
+Simpl;Trivial.
+Unfold monom_remove;Case (eqExprA (EAvar n) (EAinv a));Intros;
+ [Inversion e0|Simpl;Trivial].
+Save.
+
+Lemma monom_simplif_rem_correct:(a,e:ExprA)
+ (lvar:(listT (prodT AT nat)))~((interp_ExprA lvar a)=AzeroT)->
+ (interp_ExprA lvar (monom_simplif_rem a e))=
+ (AmultT (interp_ExprA lvar a) (interp_ExprA lvar e)).
+Proof.
+Induction a;Simpl;Intros; Try Rewrite monom_remove_correct;Auto.
+Elim (without_div_O_contr (interp_ExprA lvar e)
+ (interp_ExprA lvar e0) H1);Intros.
+Rewrite (H0 (monom_remove e e1) lvar H3);Rewrite monom_remove_correct;Auto.
+Ring.
+Save.
+
+Lemma monom_simplif_correct:(e,a:ExprA)
+ (lvar:(listT (prodT AT nat)))~((interp_ExprA lvar a)=AzeroT)->
+ (interp_ExprA lvar (monom_simplif a e))=(interp_ExprA lvar e).
+Proof.
+Induction e;Intros;Auto.
+Simpl;Case (eqExprA a e0);Intros.
+Rewrite <-e2;Apply monom_simplif_rem_correct;Auto.
+Simpl;Trivial.
+Save.
+
+Lemma inverse_correct:
+ (e,a:ExprA)(lvar:(listT (prodT AT nat)))~((interp_ExprA lvar a)=AzeroT)->
+ (interp_ExprA lvar (inverse_simplif a e))=(interp_ExprA lvar e).
+Proof.
+Induction e;Intros;Auto.
+Simpl;Rewrite (H0 a lvar H1); Rewrite monom_simplif_correct ; Auto.
+Unfold inverse_simplif;Rewrite monom_simplif_correct ; Auto.
+Save.
+
+End Theory_of_fields.
diff --git a/contrib7/fourier/Fourier.v b/contrib7/fourier/Fourier.v
new file mode 100644
index 00000000..740bbef6
--- /dev/null
+++ b/contrib7/fourier/Fourier.v
@@ -0,0 +1,28 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Fourier.v,v 1.1.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+(* "Fourier's method to solve linear inequations/equations systems.".*)
+
+Declare ML Module "quote".
+Declare ML Module "ring".
+Declare ML Module "fourier".
+Declare ML Module "fourierR".
+Declare ML Module "field".
+
+Require Export Fourier_util.
+Require Export Field.
+Require Export DiscrR.
+
+Tactic Definition Fourier :=
+ Abstract (FourierZ;Field;DiscrR).
+
+Tactic Definition FourierEq :=
+ Apply Rge_ge_eq ; Fourier.
+
diff --git a/contrib7/fourier/Fourier_util.v b/contrib7/fourier/Fourier_util.v
new file mode 100644
index 00000000..be22e2ff
--- /dev/null
+++ b/contrib7/fourier/Fourier_util.v
@@ -0,0 +1,236 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Fourier_util.v,v 1.2.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+Require Export Rbase.
+Comments "Lemmas used by the tactic Fourier".
+
+Open Scope R_scope.
+
+Lemma Rfourier_lt:
+ (x1, y1, a : R) (Rlt x1 y1) -> (Rlt R0 a) -> (Rlt (Rmult a x1) (Rmult a y1)).
+Intros; Apply Rlt_monotony; Assumption.
+Qed.
+
+Lemma Rfourier_le:
+ (x1, y1, a : R) (Rle x1 y1) -> (Rlt R0 a) -> (Rle (Rmult a x1) (Rmult a y1)).
+Red.
+Intros.
+Case H; Auto with real.
+Qed.
+
+Lemma Rfourier_lt_lt:
+ (x1, y1, x2, y2, a : R)
+ (Rlt x1 y1) ->
+ (Rlt x2 y2) ->
+ (Rlt R0 a) -> (Rlt (Rplus x1 (Rmult a x2)) (Rplus y1 (Rmult a y2))).
+Intros x1 y1 x2 y2 a H H0 H1; Try Assumption.
+Apply Rplus_lt.
+Try Exact H.
+Apply Rfourier_lt.
+Try Exact H0.
+Try Exact H1.
+Qed.
+
+Lemma Rfourier_lt_le:
+ (x1, y1, x2, y2, a : R)
+ (Rlt x1 y1) ->
+ (Rle x2 y2) ->
+ (Rlt R0 a) -> (Rlt (Rplus x1 (Rmult a x2)) (Rplus y1 (Rmult a y2))).
+Intros x1 y1 x2 y2 a H H0 H1; Try Assumption.
+Case H0; Intros.
+Apply Rplus_lt.
+Try Exact H.
+Apply Rfourier_lt; Auto with real.
+Rewrite H2.
+Rewrite (Rplus_sym y1 (Rmult a y2)).
+Rewrite (Rplus_sym x1 (Rmult a y2)).
+Apply Rlt_compatibility.
+Try Exact H.
+Qed.
+
+Lemma Rfourier_le_lt:
+ (x1, y1, x2, y2, a : R)
+ (Rle x1 y1) ->
+ (Rlt x2 y2) ->
+ (Rlt R0 a) -> (Rlt (Rplus x1 (Rmult a x2)) (Rplus y1 (Rmult a y2))).
+Intros x1 y1 x2 y2 a H H0 H1; Try Assumption.
+Case H; Intros.
+Apply Rfourier_lt_le; Auto with real.
+Rewrite H2.
+Apply Rlt_compatibility.
+Apply Rfourier_lt; Auto with real.
+Qed.
+
+Lemma Rfourier_le_le:
+ (x1, y1, x2, y2, a : R)
+ (Rle x1 y1) ->
+ (Rle x2 y2) ->
+ (Rlt R0 a) -> (Rle (Rplus x1 (Rmult a x2)) (Rplus y1 (Rmult a y2))).
+Intros x1 y1 x2 y2 a H H0 H1; Try Assumption.
+Case H0; Intros.
+Red.
+Left; Try Assumption.
+Apply Rfourier_le_lt; Auto with real.
+Rewrite H2.
+Case H; Intros.
+Red.
+Left; Try Assumption.
+Rewrite (Rplus_sym x1 (Rmult a y2)).
+Rewrite (Rplus_sym y1 (Rmult a y2)).
+Apply Rlt_compatibility.
+Try Exact H3.
+Rewrite H3.
+Red.
+Right; Try Assumption.
+Auto with real.
+Qed.
+
+Lemma Rlt_zero_pos_plus1: (x : R) (Rlt R0 x) -> (Rlt R0 (Rplus R1 x)).
+Intros x H; Try Assumption.
+Rewrite Rplus_sym.
+Apply Rlt_r_plus_R1.
+Red; Auto with real.
+Qed.
+
+Lemma Rlt_mult_inv_pos:
+ (x, y : R) (Rlt R0 x) -> (Rlt R0 y) -> (Rlt R0 (Rmult x (Rinv y))).
+Intros x y H H0; Try Assumption.
+Replace R0 with (Rmult x R0).
+Apply Rlt_monotony; Auto with real.
+Ring.
+Qed.
+
+Lemma Rlt_zero_1: (Rlt R0 R1).
+Exact Rlt_R0_R1.
+Qed.
+
+Lemma Rle_zero_pos_plus1: (x : R) (Rle R0 x) -> (Rle R0 (Rplus R1 x)).
+Intros x H; Try Assumption.
+Case H; Intros.
+Red.
+Left; Try Assumption.
+Apply Rlt_zero_pos_plus1; Auto with real.
+Rewrite <- H0.
+Replace (Rplus R1 R0) with R1.
+Red; Left.
+Exact Rlt_zero_1.
+Ring.
+Qed.
+
+Lemma Rle_mult_inv_pos:
+ (x, y : R) (Rle R0 x) -> (Rlt R0 y) -> (Rle R0 (Rmult x (Rinv y))).
+Intros x y H H0; Try Assumption.
+Case H; Intros.
+Red; Left.
+Apply Rlt_mult_inv_pos; Auto with real.
+Rewrite <- H1.
+Red; Right; Ring.
+Qed.
+
+Lemma Rle_zero_1: (Rle R0 R1).
+Red; Left.
+Exact Rlt_zero_1.
+Qed.
+
+Lemma Rle_not_lt:
+ (n, d : R) (Rle R0 (Rmult n (Rinv d))) -> ~ (Rlt R0 (Rmult (Ropp n) (Rinv d))).
+Intros n d H; Red; Intros H0; Try Exact H0.
+Generalize (Rle_not R0 (Rmult n (Rinv d))).
+Intros H1; Elim H1; Try Assumption.
+Replace (Rmult n (Rinv d)) with (Ropp (Ropp (Rmult n (Rinv d)))).
+Replace R0 with (Ropp (Ropp R0)).
+Replace (Ropp (Rmult n (Rinv d))) with (Rmult (Ropp n) (Rinv d)).
+Replace (Ropp R0) with R0.
+Red.
+Apply Rgt_Ropp.
+Red.
+Exact H0.
+Ring.
+Ring.
+Ring.
+Ring.
+Qed.
+
+Lemma Rnot_lt0: (x : R) ~ (Rlt R0 (Rmult R0 x)).
+Intros x; Try Assumption.
+Replace (Rmult R0 x) with R0.
+Apply Rlt_antirefl.
+Ring.
+Qed.
+
+Lemma Rlt_not_le:
+ (n, d : R) (Rlt R0 (Rmult n (Rinv d))) -> ~ (Rle R0 (Rmult (Ropp n) (Rinv d))).
+Intros n d H; Try Assumption.
+Apply Rle_not.
+Replace R0 with (Ropp R0).
+Replace (Rmult (Ropp n) (Rinv d)) with (Ropp (Rmult n (Rinv d))).
+Apply Rlt_Ropp.
+Try Exact H.
+Ring.
+Ring.
+Qed.
+
+Lemma Rnot_lt_lt: (x, y : R) ~ (Rlt R0 (Rminus y x)) -> ~ (Rlt x y).
+Unfold not; Intros.
+Apply H.
+Apply Rlt_anti_compatibility with x.
+Replace (Rplus x R0) with x.
+Replace (Rplus x (Rminus y x)) with y.
+Try Exact H0.
+Ring.
+Ring.
+Qed.
+
+Lemma Rnot_le_le: (x, y : R) ~ (Rle R0 (Rminus y x)) -> ~ (Rle x y).
+Unfold not; Intros.
+Apply H.
+Case H0; Intros.
+Left.
+Apply Rlt_anti_compatibility with x.
+Replace (Rplus x R0) with x.
+Replace (Rplus x (Rminus y x)) with y.
+Try Exact H1.
+Ring.
+Ring.
+Right.
+Rewrite H1; Ring.
+Qed.
+
+Lemma Rfourier_gt_to_lt: (x, y : R) (Rgt y x) -> (Rlt x y).
+Unfold Rgt; Intros; Assumption.
+Qed.
+
+Lemma Rfourier_ge_to_le: (x, y : R) (Rge y x) -> (Rle x y).
+Intros x y; Exact (Rge_le y x).
+Qed.
+
+Lemma Rfourier_eqLR_to_le: (x, y : R) x == y -> (Rle x y).
+Exact eq_Rle.
+Qed.
+
+Lemma Rfourier_eqRL_to_le: (x, y : R) y == x -> (Rle x y).
+Exact eq_Rle_sym.
+Qed.
+
+Lemma Rfourier_not_ge_lt: (x, y : R) ((Rge x y) -> False) -> (Rlt x y).
+Exact not_Rge.
+Qed.
+
+Lemma Rfourier_not_gt_le: (x, y : R) ((Rgt x y) -> False) -> (Rle x y).
+Exact Rgt_not_le.
+Qed.
+
+Lemma Rfourier_not_le_gt: (x, y : R) ((Rle x y) -> False) -> (Rgt x y).
+Exact not_Rle.
+Qed.
+
+Lemma Rfourier_not_lt_ge: (x, y : R) ((Rlt x y) -> False) -> (Rge x y).
+Exact Rlt_not_ge.
+Qed.
diff --git a/contrib7/interface/AddDad.v b/contrib7/interface/AddDad.v
new file mode 100644
index 00000000..d22b7ed1
--- /dev/null
+++ b/contrib7/interface/AddDad.v
@@ -0,0 +1,19 @@
+Grammar vernac vernac : ast :=
+ add_dad_rule00 ["AddDadRule" stringarg($name) constrarg($pat)
+ "first_path" "second_path" tacarg($tac) "."] ->
+ [(AddDadRule $name $pat (NUMBERLIST) (NUMBERLIST) (TACTIC $tac))].
+Grammar vernac vernac:ast :=
+| add_dad_rule01 ["AddDadRule" stringarg($name) constrarg($pat)
+ "first_path" "second_path" ne_numarg_list($l) tacarg($tac) "."] ->
+ [(AddDadRule $name $pat (NUMBERLIST) (NUMBERLIST ($LIST $l)) (TACTIC $tac))]
+| add_dad_rule10 ["AddDadRule" stringarg($name) constrarg($pat)
+ "first_path" ne_numarg_list($l) "second_path" tacarg($tac) "."] ->
+ [(AddDadRule $name $pat (NUMBERLIST ($LIST $l))(NUMBERLIST) (TACTIC $tac))]
+| add_dad_rule11 ["AddDadRule" stringarg($name) constrarg($pat)
+ "first_path" ne_numarg_list($l) "second_path" ne_numarg_list($l1)
+ tacarg($tac) "."] ->
+ [(AddDadRule $name $pat (NUMBERLIST ($LIST $l))(NUMBERLIST ($LIST $l1))
+ (TACTIC $tac))].
+
+Grammar vernac vernac : ast :=
+ start_dad [ "StartDad" "."] -> [(StartDad)].
diff --git a/contrib7/interface/Centaur.v b/contrib7/interface/Centaur.v
new file mode 100644
index 00000000..d27929f8
--- /dev/null
+++ b/contrib7/interface/Centaur.v
@@ -0,0 +1,88 @@
+(*
+Declare ML Module "ctast".
+Declare ML Module "paths".
+Declare ML Module "name_to_ast".
+Declare ML Module "xlate".
+Declare ML Module "vtp".
+Declare ML Module "translate".
+Declare ML Module "pbp".
+Declare ML Module "blast".
+Declare ML Module "dad".
+Declare ML Module "showproof_ct".
+Declare ML Module "showproof".
+Declare ML Module "debug_tac".
+Declare ML Module "paths".
+Declare ML Module "history".
+Declare ML Module "centaur".
+(* Require Export Illustrations. *)
+(* Require Export AddDad. *)
+(*
+Grammar vernac vernac : ast :=
+ goal_cmd [ "Goal" "Cmd" numarg($n) "with" tacarg($tac) "." ] ->
+ [(GOAL_CMD $n (TACTIC $tac))]
+| kill_proof_after [ "Kill" "Proof" "after" numarg($n)"." ] -> [(KILL_NODE $n)]
+| kill_proof_at [ "Kill" "Proof" "at" numarg($n)"." ] -> [(KILL_NODE $n)]
+| kill_sub_proof [ "Kill" "SubProof" numarg($n) "." ] -> [(KILL_SUB_PROOF $n)]
+
+| print_past_goal [ "Print" "Past" "Goal" numarg($n) "." ] ->
+ [(PRINT_GOAL_AT $n) ]
+
+| check_in_goal [ "CHECK_IN_GOAL" numarg($n) constrarg($c) "." ] ->
+ [(CHECK_IN_GOAL "CHECK" $n $c)]
+| eval_in_goal [ "EVAL_IN_GOAL" numarg($n) constrarg($c) "." ] ->
+ [(CHECK_IN_GOAL "EVAL" $n $c)]
+| compute_in_goal [ "COMPUTE_IN_GOAL" numarg($n) constrarg($c) "." ] ->
+ [(CHECK_IN_GOAL "COMPUTE" $n $c)]
+| centaur_reset [ "Centaur" "Reset" identarg($id) "." ] -> [(Centaur_Reset $id)]
+(*| show_dad_rules [ "Show" "Dad" "Rules" "." ] -> [(Show_dad_rules)]*)
+| start_pcoq [ "Start" "Pcoq" "Mode" "." ] -> [ (START_PCOQ) ]
+| start_pcoq [ "Start" "Pcoq" "Debug" "Mode" "." ] -> [ (START_PCOQ_DEBUG) ].
+Grammar vernac ne_id_list : ast list :=
+ id_one [ identarg($id)] -> [$id]
+ | id_more [identarg($id) ne_id_list($others)] -> [$id ($LIST $others)].
+
+Grammar tactic ne_num_list : ast list :=
+ ne_last [ numarg($n) ] -> [ $n ]
+| ne_num_ste [ numarg($n) ne_num_list($ns) ] -> [ $n ($LIST $ns)].
+
+Grammar tactic two_numarg_list : ast list :=
+ two_single_and_ne [ numarg($n) "to" ne_num_list($ns)] ->
+ [$n (TACTIC (to)) ($LIST $ns)]
+| two_rec [ numarg($n) two_numarg_list($ns)] -> [ $n ($LIST $ns)].
+
+Grammar tactic simple_tactic : ast :=
+ pbp0 [ "Pbp" ] -> [(PcoqPbp)]
+| pbp1 [ "Pbp" ne_num_list($ns) ] ->
+ [ (PcoqPbp ($LIST $ns)) ]
+| pbp2 [ "Pbp" identarg($id) ] -> [ (PcoqPbp $id) ]
+| pbp3 [ "Pbp" identarg($id) ne_num_list($ns)] ->
+ [ (PcoqPbp $id ($LIST $ns)) ]
+| blast1 [ "Blast" ne_num_list($ns) ] ->
+ [ (PcoqBlast ($LIST $ns)) ]
+| dad00 [ "Dad" "to" ] -> [(Dad (TACTIC (to)))]
+| dad01 [ "Dad" "to" ne_num_list($ns) ] ->
+ [(Dad (TACTIC (to)) ($LIST $ns))]
+| dadnn [ "Dad" two_numarg_list($ns) ] -> [ (Dad ($LIST $ns)) ]
+| debug_tac [ "DebugTac" tactic($tac) ] ->
+ [(CtDebugTac (TACTIC $tac))]
+| on_then_empty [ "OnThen" tactic($tac1) tactic($tac2) ] ->
+ [(OnThen (TACTIC $tac1) (TACTIC $tac2))]
+| on_then_ne [ "OnThen" tactic($tac1) tactic($tac2) ne_num_list($l) ] ->
+ [(OnThen (TACTIC $tac1) (TACTIC $tac2) ($LIST $l))]
+| debug_tac2 [ "DebugTac2" tactic($tac) ] ->
+ [(CtDebugTac2 (TACTIC $tac))].
+
+
+(* Maybe we should have syntactic rules to make sur that syntax errors are
+ displayed with a readable syntax. It is not sure, since the error reporting
+ procedure changed from V6.1 and does not reprint the command anymore. *)
+Grammar vernac vernac : ast :=
+ text_proof_flag_on [ "Text" "Mode" "fr" "." ] ->
+ [(TEXT_MODE (AST "fr"))]
+| text_proof_flag_on [ "Text" "Mode" "en" "." ] ->
+ [(TEXT_MODE (AST "en"))]
+| text_proof_flag_on [ "Text" "Mode" "Off" "." ] ->
+ [(TEXT_MODE (AST "off"))].
+
+*)
+*)
diff --git a/contrib7/interface/vernacrc b/contrib7/interface/vernacrc
new file mode 100644
index 00000000..f95c4212
--- /dev/null
+++ b/contrib7/interface/vernacrc
@@ -0,0 +1,17 @@
+# $Id: vernacrc,v 1.1 2003/11/29 20:02:41 herbelin Exp $
+
+# This file is loaded initially by ./vernacparser.
+
+load_syntax_file 17 LogicSyntax
+load_syntax_file 36 SpecifSyntax
+load_syntax_file 18 Logic_TypeSyntax
+load_syntax_file 19 DatatypesSyntax
+load_syntax_file 21 Equality
+load_syntax_file 22 Inv
+load_syntax_file 26 Tauto
+load_syntax_file 34 Omega
+load_syntax_file 27 Ring
+quiet_parse_string
+Goal a.
+&& END--OF--DATA
+print_version
diff --git a/contrib7/omega/Omega.v b/contrib7/omega/Omega.v
new file mode 100644
index 00000000..76e37519
--- /dev/null
+++ b/contrib7/omega/Omega.v
@@ -0,0 +1,57 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+(**************************************************************************)
+(* *)
+(* Omega: a solver of quantifier-free problems in Presburger Arithmetic *)
+(* *)
+(* Pierre Crégut (CNET, Lannion, France) *)
+(* *)
+(**************************************************************************)
+
+(* $Id: Omega.v,v 1.1.2.1 2004/07/16 19:30:17 herbelin Exp $ *)
+
+(* We do not require [ZArith] anymore, but only what's necessary for Omega *)
+Require Export ZArith_base.
+Require Export OmegaLemmas.
+
+Hints Resolve Zle_n Zplus_sym Zplus_assoc Zmult_sym Zmult_assoc
+ Zero_left Zero_right Zmult_one Zplus_inverse_l Zplus_inverse_r
+ Zmult_plus_distr_l Zmult_plus_distr_r : zarith.
+
+Require Export Zhints.
+
+(*
+(* The constant minus is required in coq_omega.ml *)
+Require Minus.
+*)
+
+Hint eq_nat_Omega : zarith := Extern 10 (eq nat ? ?) Abstract Omega.
+Hint le_Omega : zarith := Extern 10 (le ? ?) Abstract Omega.
+Hint lt_Omega : zarith := Extern 10 (lt ? ?) Abstract Omega.
+Hint ge_Omega : zarith := Extern 10 (ge ? ?) Abstract Omega.
+Hint gt_Omega : zarith := Extern 10 (gt ? ?) Abstract Omega.
+
+Hint not_eq_nat_Omega : zarith := Extern 10 ~(eq nat ? ?) Abstract Omega.
+Hint not_le_Omega : zarith := Extern 10 ~(le ? ?) Abstract Omega.
+Hint not_lt_Omega : zarith := Extern 10 ~(lt ? ?) Abstract Omega.
+Hint not_ge_Omega : zarith := Extern 10 ~(ge ? ?) Abstract Omega.
+Hint not_gt_Omega : zarith := Extern 10 ~(gt ? ?) Abstract Omega.
+
+Hint eq_Z_Omega : zarith := Extern 10 (eq Z ? ?) Abstract Omega.
+Hint Zle_Omega : zarith := Extern 10 (Zle ? ?) Abstract Omega.
+Hint Zlt_Omega : zarith := Extern 10 (Zlt ? ?) Abstract Omega.
+Hint Zge_Omega : zarith := Extern 10 (Zge ? ?) Abstract Omega.
+Hint Zgt_Omega : zarith := Extern 10 (Zgt ? ?) Abstract Omega.
+
+Hint not_eq_nat_Omega : zarith := Extern 10 ~(eq Z ? ?) Abstract Omega.
+Hint not_Zle_Omega : zarith := Extern 10 ~(Zle ? ?) Abstract Omega.
+Hint not_Zlt_Omega : zarith := Extern 10 ~(Zlt ? ?) Abstract Omega.
+Hint not_Zge_Omega : zarith := Extern 10 ~(Zge ? ?) Abstract Omega.
+Hint not_Zgt_Omega : zarith := Extern 10 ~(Zgt ? ?) Abstract Omega.
+
+Hint false_Omega : zarith := Extern 10 False Abstract Omega.
diff --git a/contrib7/omega/OmegaLemmas.v b/contrib7/omega/OmegaLemmas.v
new file mode 100644
index 00000000..0d05fc3e
--- /dev/null
+++ b/contrib7/omega/OmegaLemmas.v
@@ -0,0 +1,399 @@
+(************************************************************************)
+(* 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 $Id: OmegaLemmas.v,v 1.1.2.1 2004/07/16 19:30:17 herbelin Exp $ i*)
+
+Require ZArith_base.
+
+(** These are specific variants of theorems dedicated for the Omega tactic *)
+
+Lemma new_var: (x:Z) (EX y:Z |(x=y)).
+Intros x; Exists x; Trivial with arith.
+Qed.
+
+Lemma OMEGA1 : (x,y:Z) (x=y) -> (Zle ZERO x) -> (Zle ZERO y).
+Intros x y H; Rewrite H; Auto with arith.
+Qed.
+
+Lemma OMEGA2 : (x,y:Z) (Zle ZERO x) -> (Zle ZERO y) -> (Zle ZERO (Zplus x y)).
+Exact Zle_0_plus.
+Qed.
+
+Lemma OMEGA3 :
+ (x,y,k:Z)(Zgt k ZERO)-> (x=(Zmult y k)) -> (x=ZERO) -> (y=ZERO).
+
+Intros x y k H1 H2 H3; Apply (Zmult_eq k); [
+ Unfold not ; Intros H4; Absurd (Zgt k ZERO); [
+ Rewrite H4; Unfold Zgt ; Simpl; Discriminate | Assumption]
+ | Rewrite <- H2; Assumption].
+Qed.
+
+Lemma OMEGA4 :
+ (x,y,z:Z)(Zgt x ZERO) -> (Zgt y x) -> ~(Zplus (Zmult z y) x) = ZERO.
+
+Unfold not ; Intros x y z H1 H2 H3; Cut (Zgt y ZERO); [
+ Intros H4; Cut (Zle ZERO (Zplus (Zmult z y) x)); [
+ Intros H5; Generalize (Zmult_le_approx y z x H4 H2 H5) ; Intros H6;
+ Absurd (Zgt (Zplus (Zmult z y) x) ZERO); [
+ Rewrite -> H3; Unfold Zgt ; Simpl; Discriminate
+ | Apply Zle_gt_trans with x ; [
+ Pattern 1 x ; Rewrite <- (Zero_left x); Apply Zle_reg_r;
+ Rewrite -> Zmult_sym; Generalize H4 ; Unfold Zgt;
+ Case y; [
+ Simpl; Intros H7; Discriminate H7
+ | Intros p H7; Rewrite <- (Zero_mult_right (POS p));
+ Unfold Zle ; Rewrite -> Zcompare_Zmult_compatible; Exact H6
+ | Simpl; Intros p H7; Discriminate H7]
+ | Assumption]]
+ | Rewrite -> H3; Unfold Zle ; Simpl; Discriminate]
+ | Apply Zgt_trans with x ; [ Assumption | Assumption]].
+Qed.
+
+Lemma OMEGA5: (x,y,z:Z)(x=ZERO) -> (y=ZERO) -> (Zplus x (Zmult y z)) = ZERO.
+
+Intros x y z H1 H2; Rewrite H1; Rewrite H2; Simpl; Trivial with arith.
+Qed.
+
+Lemma OMEGA6:
+ (x,y,z:Z)(Zle ZERO x) -> (y=ZERO) -> (Zle ZERO (Zplus x (Zmult y z))).
+
+Intros x y z H1 H2; Rewrite H2; Simpl; Rewrite Zero_right; Assumption.
+Qed.
+
+Lemma OMEGA7:
+ (x,y,z,t:Z)(Zgt z ZERO) -> (Zgt t ZERO) -> (Zle ZERO x) -> (Zle ZERO y) ->
+ (Zle ZERO (Zplus (Zmult x z) (Zmult y t))).
+
+Intros x y z t H1 H2 H3 H4; Rewrite <- (Zero_left ZERO);
+Apply Zle_plus_plus; Apply Zle_mult; Assumption.
+Qed.
+
+Lemma OMEGA8:
+ (x,y:Z) (Zle ZERO x) -> (Zle ZERO y) -> x = (Zopp y) -> x = ZERO.
+
+Intros x y H1 H2 H3; Elim (Zle_lt_or_eq ZERO x H1); [
+ Intros H4; Absurd (Zlt ZERO x); [
+ Change (Zge ZERO x); Apply Zle_ge; Apply Zsimpl_le_plus_l with y;
+ Rewrite -> H3; Rewrite Zplus_inverse_r; Rewrite Zero_right; Assumption
+ | Assumption]
+| Intros H4; Rewrite -> H4; Trivial with arith].
+Qed.
+
+Lemma OMEGA9:(x,y,z,t:Z) y=ZERO -> x = z ->
+ (Zplus y (Zmult (Zplus (Zopp x) z) t)) = ZERO.
+
+Intros x y z t H1 H2; Rewrite H2; Rewrite Zplus_inverse_l;
+Rewrite Zero_mult_left; Rewrite Zero_right; Assumption.
+Qed.
+
+Lemma OMEGA10:(v,c1,c2,l1,l2,k1,k2:Z)
+ (Zplus (Zmult (Zplus (Zmult v c1) l1) k1) (Zmult (Zplus (Zmult v c2) l2) k2))
+ = (Zplus (Zmult v (Zplus (Zmult c1 k1) (Zmult c2 k2)))
+ (Zplus (Zmult l1 k1) (Zmult l2 k2))).
+
+Intros; Repeat (Rewrite Zmult_plus_distr_l Orelse Rewrite Zmult_plus_distr_r);
+Repeat Rewrite Zmult_assoc; Repeat Elim Zplus_assoc;
+Rewrite (Zplus_permute (Zmult l1 k1) (Zmult (Zmult v c2) k2)); Trivial with arith.
+Qed.
+
+Lemma OMEGA11:(v1,c1,l1,l2,k1:Z)
+ (Zplus (Zmult (Zplus (Zmult v1 c1) l1) k1) l2)
+ = (Zplus (Zmult v1 (Zmult c1 k1)) (Zplus (Zmult l1 k1) l2)).
+
+Intros; Repeat (Rewrite Zmult_plus_distr_l Orelse Rewrite Zmult_plus_distr_r);
+Repeat Rewrite Zmult_assoc; Repeat Elim Zplus_assoc; Trivial with arith.
+Qed.
+
+Lemma OMEGA12:(v2,c2,l1,l2,k2:Z)
+ (Zplus l1 (Zmult (Zplus (Zmult v2 c2) l2) k2))
+ = (Zplus (Zmult v2 (Zmult c2 k2)) (Zplus l1 (Zmult l2 k2))).
+
+Intros; Repeat (Rewrite Zmult_plus_distr_l Orelse Rewrite Zmult_plus_distr_r);
+Repeat Rewrite Zmult_assoc; Repeat Elim Zplus_assoc; Rewrite Zplus_permute;
+Trivial with arith.
+Qed.
+
+Lemma OMEGA13:(v,l1,l2:Z)(x:positive)
+ (Zplus (Zplus (Zmult v (POS x)) l1) (Zplus (Zmult v (NEG x)) l2))
+ = (Zplus l1 l2).
+
+Intros; Rewrite Zplus_assoc; Rewrite (Zplus_sym (Zmult v (POS x)) l1);
+Rewrite (Zplus_assoc_r l1); Rewrite <- Zmult_plus_distr_r;
+Rewrite <- Zopp_NEG; Rewrite (Zplus_sym (Zopp (NEG x)) (NEG x));
+Rewrite Zplus_inverse_r; Rewrite Zero_mult_right; Rewrite Zero_right; Trivial with arith.
+Qed.
+
+Lemma OMEGA14:(v,l1,l2:Z)(x:positive)
+ (Zplus (Zplus (Zmult v (NEG x)) l1) (Zplus (Zmult v (POS x)) l2))
+ = (Zplus l1 l2).
+
+Intros; Rewrite Zplus_assoc; Rewrite (Zplus_sym (Zmult v (NEG x)) l1);
+Rewrite (Zplus_assoc_r l1); Rewrite <- Zmult_plus_distr_r;
+Rewrite <- Zopp_NEG; Rewrite Zplus_inverse_r; Rewrite Zero_mult_right;
+Rewrite Zero_right; Trivial with arith.
+Qed.
+Lemma OMEGA15:(v,c1,c2,l1,l2,k2:Z)
+ (Zplus (Zplus (Zmult v c1) l1) (Zmult (Zplus (Zmult v c2) l2) k2))
+ = (Zplus (Zmult v (Zplus c1 (Zmult c2 k2)))
+ (Zplus l1 (Zmult l2 k2))).
+
+Intros; Repeat (Rewrite Zmult_plus_distr_l Orelse Rewrite Zmult_plus_distr_r);
+Repeat Rewrite Zmult_assoc; Repeat Elim Zplus_assoc;
+Rewrite (Zplus_permute l1 (Zmult (Zmult v c2) k2)); Trivial with arith.
+Qed.
+
+Lemma OMEGA16:
+ (v,c,l,k:Z)
+ (Zmult (Zplus (Zmult v c) l) k) = (Zplus (Zmult v (Zmult c k)) (Zmult l k)).
+
+Intros; Repeat (Rewrite Zmult_plus_distr_l Orelse Rewrite Zmult_plus_distr_r);
+Repeat Rewrite Zmult_assoc; Repeat Elim Zplus_assoc; Trivial with arith.
+Qed.
+
+Lemma OMEGA17:
+ (x,y,z:Z)(Zne x ZERO) -> (y=ZERO) -> (Zne (Zplus x (Zmult y z)) ZERO).
+
+Unfold Zne not; Intros x y z H1 H2 H3; Apply H1;
+Apply Zsimpl_plus_l with (Zmult y z); Rewrite Zplus_sym; Rewrite H3;
+Rewrite H2; Auto with arith.
+Qed.
+
+Lemma OMEGA18:
+ (x,y,k:Z) x=(Zmult y k) -> (Zne x ZERO) -> (Zne y ZERO).
+
+Unfold Zne not; Intros x y k H1 H2 H3; Apply H2; Rewrite H1; Rewrite H3; Auto with arith.
+Qed.
+
+Lemma OMEGA19:
+ (x:Z) (Zne x ZERO) ->
+ (Zle ZERO (Zplus x (NEG xH))) \/ (Zle ZERO (Zplus (Zmult x (NEG xH)) (NEG xH))).
+
+Unfold Zne ; Intros x H; Elim (Zle_or_lt ZERO x); [
+ Intros H1; Elim Zle_lt_or_eq with 1:=H1; [
+ Intros H2; Left; Change (Zle ZERO (Zpred x)); Apply Zle_S_n;
+ Rewrite <- Zs_pred; Apply Zlt_le_S; Assumption
+ | Intros H2; Absurd x=ZERO; Auto with arith]
+| Intros H1; Right; Rewrite <- Zopp_one; Rewrite Zplus_sym;
+ Apply Zle_left; Apply Zle_S_n; Simpl; Apply Zlt_le_S; Auto with arith].
+Qed.
+
+Lemma OMEGA20:
+ (x,y,z:Z)(Zne x ZERO) -> (y=ZERO) -> (Zne (Zplus x (Zmult y z)) ZERO).
+
+Unfold Zne not; Intros x y z H1 H2 H3; Apply H1; Rewrite H2 in H3;
+Simpl in H3; Rewrite Zero_right in H3; Trivial with arith.
+Qed.
+
+Definition fast_Zplus_sym :=
+[x,y:Z][P:Z -> Prop][H: (P (Zplus y x))]
+ (eq_ind_r Z (Zplus y x) P H (Zplus x y) (Zplus_sym x y)).
+
+Definition fast_Zplus_assoc_r :=
+[n,m,p:Z][P:Z -> Prop][H : (P (Zplus n (Zplus m p)))]
+ (eq_ind_r Z (Zplus n (Zplus m p)) P H (Zplus (Zplus n m) p) (Zplus_assoc_r n m p)).
+
+Definition fast_Zplus_assoc_l :=
+[n,m,p:Z][P:Z -> Prop][H : (P (Zplus (Zplus n m) p))]
+ (eq_ind_r Z (Zplus (Zplus n m) p) P H (Zplus n (Zplus m p))
+ (Zplus_assoc_l n m p)).
+
+Definition fast_Zplus_permute :=
+[n,m,p:Z][P:Z -> Prop][H : (P (Zplus m (Zplus n p)))]
+ (eq_ind_r Z (Zplus m (Zplus n p)) P H (Zplus n (Zplus m p))
+ (Zplus_permute n m p)).
+
+Definition fast_OMEGA10 :=
+[v,c1,c2,l1,l2,k1,k2:Z][P:Z -> Prop]
+[H : (P (Zplus (Zmult v (Zplus (Zmult c1 k1) (Zmult c2 k2)))
+ (Zplus (Zmult l1 k1) (Zmult l2 k2))))]
+ (eq_ind_r Z
+ (Zplus (Zmult v (Zplus (Zmult c1 k1) (Zmult c2 k2)))
+ (Zplus (Zmult l1 k1) (Zmult l2 k2)))
+ P H
+ (Zplus (Zmult (Zplus (Zmult v c1) l1) k1)
+ (Zmult (Zplus (Zmult v c2) l2) k2))
+ (OMEGA10 v c1 c2 l1 l2 k1 k2)).
+
+Definition fast_OMEGA11 :=
+[v1,c1,l1,l2,k1:Z][P:Z -> Prop]
+[H : (P (Zplus (Zmult v1 (Zmult c1 k1)) (Zplus (Zmult l1 k1) l2)))]
+ (eq_ind_r Z
+ (Zplus (Zmult v1 (Zmult c1 k1)) (Zplus (Zmult l1 k1) l2))
+ P H
+ (Zplus (Zmult (Zplus (Zmult v1 c1) l1) k1) l2)
+ (OMEGA11 v1 c1 l1 l2 k1)).
+Definition fast_OMEGA12 :=
+[v2,c2,l1,l2,k2:Z][P:Z -> Prop]
+[H : (P (Zplus (Zmult v2 (Zmult c2 k2)) (Zplus l1 (Zmult l2 k2))))]
+ (eq_ind_r Z
+ (Zplus (Zmult v2 (Zmult c2 k2)) (Zplus l1 (Zmult l2 k2)))
+ P H
+ (Zplus l1 (Zmult (Zplus (Zmult v2 c2) l2) k2))
+ (OMEGA12 v2 c2 l1 l2 k2)).
+
+Definition fast_OMEGA15 :=
+[v,c1,c2,l1,l2,k2 :Z][P:Z -> Prop]
+[H : (P (Zplus (Zmult v (Zplus c1 (Zmult c2 k2))) (Zplus l1 (Zmult l2 k2))))]
+ (eq_ind_r Z
+ (Zplus (Zmult v (Zplus c1 (Zmult c2 k2))) (Zplus l1 (Zmult l2 k2)))
+ P H
+ (Zplus (Zplus (Zmult v c1) l1) (Zmult (Zplus (Zmult v c2) l2) k2))
+ (OMEGA15 v c1 c2 l1 l2 k2)).
+Definition fast_OMEGA16 :=
+[v,c,l,k :Z][P:Z -> Prop]
+[H : (P (Zplus (Zmult v (Zmult c k)) (Zmult l k)))]
+ (eq_ind_r Z
+ (Zplus (Zmult v (Zmult c k)) (Zmult l k))
+ P H
+ (Zmult (Zplus (Zmult v c) l) k)
+ (OMEGA16 v c l k)).
+
+Definition fast_OMEGA13 :=
+[v,l1,l2 :Z][x:positive][P:Z -> Prop]
+[H : (P (Zplus l1 l2))]
+ (eq_ind_r Z
+ (Zplus l1 l2)
+ P H
+ (Zplus (Zplus (Zmult v (POS x)) l1) (Zplus (Zmult v (NEG x)) l2))
+ (OMEGA13 v l1 l2 x )).
+
+Definition fast_OMEGA14 :=
+[v,l1,l2 :Z][x:positive][P:Z -> Prop]
+[H : (P (Zplus l1 l2))]
+ (eq_ind_r Z
+ (Zplus l1 l2)
+ P H
+ (Zplus (Zplus (Zmult v (NEG x)) l1) (Zplus (Zmult v (POS x)) l2))
+ (OMEGA14 v l1 l2 x )).
+Definition fast_Zred_factor0:=
+[x:Z][P:Z -> Prop]
+[H : (P (Zmult x (POS xH)) )]
+ (eq_ind_r Z
+ (Zmult x (POS xH))
+ P H
+ x
+ (Zred_factor0 x)).
+
+Definition fast_Zopp_one :=
+[x:Z][P:Z -> Prop]
+[H : (P (Zmult x (NEG xH)))]
+ (eq_ind_r Z
+ (Zmult x (NEG xH))
+ P H
+ (Zopp x)
+ (Zopp_one x)).
+
+Definition fast_Zmult_sym :=
+[x,y :Z][P:Z -> Prop]
+[H : (P (Zmult y x))]
+ (eq_ind_r Z
+(Zmult y x)
+ P H
+(Zmult x y)
+ (Zmult_sym x y )).
+
+Definition fast_Zopp_Zplus :=
+[x,y :Z][P:Z -> Prop]
+[H : (P (Zplus (Zopp x) (Zopp y)) )]
+ (eq_ind_r Z
+ (Zplus (Zopp x) (Zopp y))
+ P H
+ (Zopp (Zplus x y))
+ (Zopp_Zplus x y )).
+
+Definition fast_Zopp_Zopp :=
+[x:Z][P:Z -> Prop]
+[H : (P x )] (eq_ind_r Z x P H (Zopp (Zopp x)) (Zopp_Zopp x)).
+
+Definition fast_Zopp_Zmult_r :=
+[x,y:Z][P:Z -> Prop]
+[H : (P (Zmult x (Zopp y)))]
+ (eq_ind_r Z
+ (Zmult x (Zopp y))
+ P H
+ (Zopp (Zmult x y))
+ (Zopp_Zmult_r x y )).
+
+Definition fast_Zmult_plus_distr :=
+[n,m,p:Z][P:Z -> Prop]
+[H : (P (Zplus (Zmult n p) (Zmult m p)))]
+ (eq_ind_r Z
+ (Zplus (Zmult n p) (Zmult m p))
+ P H
+ (Zmult (Zplus n m) p)
+ (Zmult_plus_distr_l n m p)).
+Definition fast_Zmult_Zopp_left:=
+[x,y:Z][P:Z -> Prop]
+[H : (P (Zmult x (Zopp y)))]
+ (eq_ind_r Z
+ (Zmult x (Zopp y))
+ P H
+ (Zmult (Zopp x) y)
+ (Zmult_Zopp_left x y)).
+
+Definition fast_Zmult_assoc_r :=
+[n,m,p :Z][P:Z -> Prop]
+[H : (P (Zmult n (Zmult m p)))]
+ (eq_ind_r Z
+ (Zmult n (Zmult m p))
+ P H
+ (Zmult (Zmult n m) p)
+ (Zmult_assoc_r n m p)).
+
+Definition fast_Zred_factor1 :=
+[x:Z][P:Z -> Prop]
+[H : (P (Zmult x (POS (xO xH))) )]
+ (eq_ind_r Z
+ (Zmult x (POS (xO xH)))
+ P H
+ (Zplus x x)
+ (Zred_factor1 x)).
+
+Definition fast_Zred_factor2 :=
+[x,y:Z][P:Z -> Prop]
+[H : (P (Zmult x (Zplus (POS xH) y)))]
+ (eq_ind_r Z
+ (Zmult x (Zplus (POS xH) y))
+ P H
+ (Zplus x (Zmult x y))
+ (Zred_factor2 x y)).
+Definition fast_Zred_factor3 :=
+[x,y:Z][P:Z -> Prop]
+[H : (P (Zmult x (Zplus (POS xH) y)))]
+ (eq_ind_r Z
+ (Zmult x (Zplus (POS xH) y))
+ P H
+ (Zplus (Zmult x y) x)
+ (Zred_factor3 x y)).
+
+Definition fast_Zred_factor4 :=
+[x,y,z:Z][P:Z -> Prop]
+[H : (P (Zmult x (Zplus y z)))]
+ (eq_ind_r Z
+ (Zmult x (Zplus y z))
+ P H
+ (Zplus (Zmult x y) (Zmult x z))
+ (Zred_factor4 x y z)).
+
+Definition fast_Zred_factor5 :=
+[x,y:Z][P:Z -> Prop]
+[H : (P y)]
+ (eq_ind_r Z
+ y
+ P H
+ (Zplus (Zmult x ZERO) y)
+ (Zred_factor5 x y)).
+
+Definition fast_Zred_factor6 :=
+[x :Z][P:Z -> Prop]
+[H : (P(Zplus x ZERO) )]
+ (eq_ind_r Z
+ (Zplus x ZERO)
+ P H
+ x
+ (Zred_factor6 x )).
diff --git a/contrib7/ring/ArithRing.v b/contrib7/ring/ArithRing.v
new file mode 100644
index 00000000..c2abc4d1
--- /dev/null
+++ b/contrib7/ring/ArithRing.v
@@ -0,0 +1,81 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: ArithRing.v,v 1.1.2.1 2004/07/16 19:30:18 herbelin Exp $ *)
+
+(* Instantiation of the Ring tactic for the naturals of Arith $*)
+
+Require Export Ring.
+Require Export Arith.
+Require Eqdep_dec.
+
+V7only [Import nat_scope.].
+Open Local Scope nat_scope.
+
+Fixpoint nateq [n,m:nat] : bool :=
+ Cases n m of
+ | O O => true
+ | (S n') (S m') => (nateq n' m')
+ | _ _ => false
+ end.
+
+Lemma nateq_prop : (n,m:nat)(Is_true (nateq n m))->n==m.
+Proof.
+ Induction n; Induction m; Intros; Try Contradiction.
+ Trivial.
+ Unfold Is_true in H1.
+ Rewrite (H n1 H1).
+ Trivial.
+Save.
+
+Hints Resolve nateq_prop eq2eqT : arithring.
+
+Definition NatTheory : (Semi_Ring_Theory plus mult (1) (0) nateq).
+ Split; Intros; Auto with arith arithring.
+ Apply eq2eqT; Apply simpl_plus_l with n:=n.
+ Apply eqT2eq; Trivial.
+Defined.
+
+
+Add Semi Ring nat plus mult (1) (0) nateq NatTheory [O S].
+
+Goal (n:nat)(S n)=(plus (S O) n).
+Intro; Reflexivity.
+Save S_to_plus_one.
+
+(* Replace all occurrences of (S exp) by (plus (S O) exp), except when
+ exp is already O and only for those occurrences than can be reached by going
+ down plus and mult operations *)
+Recursive Meta Definition S_to_plus t :=
+ Match t With
+ | [(S O)] -> '(S O)
+ | [(S ?1)] -> Let t1 = (S_to_plus ?1) In
+ '(plus (S O) t1)
+ | [(plus ?1 ?2)] -> Let t1 = (S_to_plus ?1)
+ And t2 = (S_to_plus ?2) In
+ '(plus t1 t2)
+ | [(mult ?1 ?2)] -> Let t1 = (S_to_plus ?1)
+ And t2 = (S_to_plus ?2) In
+ '(mult t1 t2)
+ | [?] -> 't.
+
+(* Apply S_to_plus on both sides of an equality *)
+Tactic Definition S_to_plus_eq :=
+ Match Context With
+ | [ |- ?1 = ?2 ] ->
+ (**) Try (**)
+ Let t1 = (S_to_plus ?1)
+ And t2 = (S_to_plus ?2) In
+ Change t1=t2
+ | [ |- ?1 == ?2 ] ->
+ (**) Try (**)
+ Let t1 = (S_to_plus ?1)
+ And t2 = (S_to_plus ?2) In
+ Change (t1==t2).
+
+Tactic Definition NatRing := S_to_plus_eq;Ring.
diff --git a/contrib7/ring/NArithRing.v b/contrib7/ring/NArithRing.v
new file mode 100644
index 00000000..f4548bbb
--- /dev/null
+++ b/contrib7/ring/NArithRing.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 *)
+(************************************************************************)
+
+(* $Id: NArithRing.v,v 1.1.2.1 2004/07/16 19:30:18 herbelin Exp $ *)
+
+(* Instantiation of the Ring tactic for the binary natural numbers *)
+
+Require Export Ring.
+Require Export ZArith_base.
+Require NArith.
+Require Eqdep_dec.
+
+Definition Neq := [n,m:entier]
+ Cases (Ncompare n m) of
+ EGAL => true
+ | _ => false
+ end.
+
+Lemma Neq_prop : (n,m:entier)(Is_true (Neq n m)) -> n=m.
+ Intros n m H; Unfold Neq in H.
+ Apply Ncompare_Eq_eq.
+ NewDestruct (Ncompare n m); [Reflexivity | Contradiction | Contradiction ].
+Save.
+
+Definition NTheory : (Semi_Ring_Theory Nplus Nmult (Pos xH) Nul Neq).
+ Split.
+ Apply Nplus_comm.
+ Apply Nplus_assoc.
+ Apply Nmult_comm.
+ Apply Nmult_assoc.
+ Apply Nplus_0_l.
+ Apply Nmult_1_l.
+ Apply Nmult_0_l.
+ Apply Nmult_plus_distr_r.
+ Apply Nplus_reg_l.
+ Apply Neq_prop.
+Save.
+
+Add Semi Ring entier Nplus Nmult (Pos xH) Nul Neq NTheory [Pos Nul xO xI xH].
diff --git a/contrib7/ring/Quote.v b/contrib7/ring/Quote.v
new file mode 100644
index 00000000..12a51c9f
--- /dev/null
+++ b/contrib7/ring/Quote.v
@@ -0,0 +1,85 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Quote.v,v 1.1.2.1 2004/07/16 19:30:18 herbelin Exp $ *)
+
+(***********************************************************************
+ The "abstract" type index is defined to represent variables.
+
+ index : Set
+ index_eq : index -> bool
+ index_eq_prop: (n,m:index)(index_eq n m)=true -> n=m
+ index_lt : index -> bool
+ varmap : Type -> Type.
+ varmap_find : (A:Type)A -> index -> (varmap A) -> A.
+
+ The first arg. of varmap_find is the default value to take
+ if the object is not found in the varmap.
+
+ index_lt defines a total well-founded order, but we don't prove that.
+
+***********************************************************************)
+
+Set Implicit Arguments.
+
+Section variables_map.
+
+Variable A : Type.
+
+Inductive varmap : Type :=
+ Empty_vm : varmap
+| Node_vm : A->varmap->varmap->varmap.
+
+Inductive index : Set :=
+| Left_idx : index -> index
+| Right_idx : index -> index
+| End_idx : index
+.
+
+Fixpoint varmap_find [default_value:A; i:index; v:varmap] : A :=
+ Cases i v of
+ End_idx (Node_vm x _ _) => x
+ | (Right_idx i1) (Node_vm x v1 v2) => (varmap_find default_value i1 v2)
+ | (Left_idx i1) (Node_vm x v1 v2) => (varmap_find default_value i1 v1)
+ | _ _ => default_value
+ end.
+
+Fixpoint index_eq [n,m:index] : bool :=
+ Cases n m of
+ | End_idx End_idx => true
+ | (Left_idx n') (Left_idx m') => (index_eq n' m')
+ | (Right_idx n') (Right_idx m') => (index_eq n' m')
+ | _ _ => false
+ end.
+
+Fixpoint index_lt[n,m:index] : bool :=
+ Cases n m of
+ | End_idx (Left_idx _) => true
+ | End_idx (Right_idx _) => true
+ | (Left_idx n') (Right_idx m') => true
+ | (Right_idx n') (Right_idx m') => (index_lt n' m')
+ | (Left_idx n') (Left_idx m') => (index_lt n' m')
+ | _ _ => false
+ end.
+
+Lemma index_eq_prop : (n,m:index)(index_eq n m)=true -> n=m.
+ Induction n; Induction m; Simpl; Intros.
+ Rewrite (H i0 H1); Reflexivity.
+ Discriminate.
+ Discriminate.
+ Discriminate.
+ Rewrite (H i0 H1); Reflexivity.
+ Discriminate.
+ Discriminate.
+ Discriminate.
+ Reflexivity.
+Save.
+
+End variables_map.
+
+Unset Implicit Arguments.
diff --git a/contrib7/ring/Ring.v b/contrib7/ring/Ring.v
new file mode 100644
index 00000000..860dda13
--- /dev/null
+++ b/contrib7/ring/Ring.v
@@ -0,0 +1,34 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Ring.v,v 1.1.2.1 2004/07/16 19:30:18 herbelin Exp $ *)
+
+Require Export Bool.
+Require Export Ring_theory.
+Require Export Quote.
+Require Export Ring_normalize.
+Require Export Ring_abstract.
+
+(* As an example, we provide an instantation for bool. *)
+(* Other instatiations are given in ArithRing and ZArithRing in the
+ same directory *)
+
+Definition BoolTheory : (Ring_Theory xorb andb true false [b:bool]b eqb).
+Split; Simpl.
+NewDestruct n; NewDestruct m; Reflexivity.
+NewDestruct n; NewDestruct m; NewDestruct p; Reflexivity.
+NewDestruct n; NewDestruct m; Reflexivity.
+NewDestruct n; NewDestruct m; NewDestruct p; Reflexivity.
+NewDestruct n; Reflexivity.
+NewDestruct n; Reflexivity.
+NewDestruct n; Reflexivity.
+NewDestruct n; NewDestruct m; NewDestruct p; Reflexivity.
+NewDestruct x; NewDestruct y; Reflexivity Orelse Simpl; Tauto.
+Defined.
+
+Add Ring bool xorb andb true false [b:bool]b eqb BoolTheory [ true false ].
diff --git a/contrib7/ring/Ring_abstract.v b/contrib7/ring/Ring_abstract.v
new file mode 100644
index 00000000..55bb31da
--- /dev/null
+++ b/contrib7/ring/Ring_abstract.v
@@ -0,0 +1,699 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Ring_abstract.v,v 1.1.2.1 2004/07/16 19:30:18 herbelin Exp $ *)
+
+Require Ring_theory.
+Require Quote.
+Require Ring_normalize.
+
+Section abstract_semi_rings.
+
+Inductive Type aspolynomial :=
+ ASPvar : index -> aspolynomial
+| ASP0 : aspolynomial
+| ASP1 : aspolynomial
+| ASPplus : aspolynomial -> aspolynomial -> aspolynomial
+| ASPmult : aspolynomial -> aspolynomial -> aspolynomial
+.
+
+Inductive abstract_sum : Type :=
+| Nil_acs : abstract_sum
+| Cons_acs : varlist -> abstract_sum -> abstract_sum
+.
+
+Fixpoint abstract_sum_merge [s1:abstract_sum]
+ : abstract_sum -> abstract_sum :=
+Cases s1 of
+| (Cons_acs l1 t1) =>
+ Fix asm_aux{asm_aux[s2:abstract_sum] : abstract_sum :=
+ Cases s2 of
+ | (Cons_acs l2 t2) =>
+ if (varlist_lt l1 l2)
+ then (Cons_acs l1 (abstract_sum_merge t1 s2))
+ else (Cons_acs l2 (asm_aux t2))
+ | Nil_acs => s1
+ end}
+| Nil_acs => [s2]s2
+end.
+
+Fixpoint abstract_varlist_insert [l1:varlist; s2:abstract_sum]
+ : abstract_sum :=
+ Cases s2 of
+ | (Cons_acs l2 t2) =>
+ if (varlist_lt l1 l2)
+ then (Cons_acs l1 s2)
+ else (Cons_acs l2 (abstract_varlist_insert l1 t2))
+ | Nil_acs => (Cons_acs l1 Nil_acs)
+ end.
+
+Fixpoint abstract_sum_scalar [l1:varlist; s2:abstract_sum]
+ : abstract_sum :=
+ Cases s2 of
+ | (Cons_acs l2 t2) => (abstract_varlist_insert (varlist_merge l1 l2)
+ (abstract_sum_scalar l1 t2))
+ | Nil_acs => Nil_acs
+ end.
+
+Fixpoint abstract_sum_prod [s1:abstract_sum]
+ : abstract_sum -> abstract_sum :=
+ [s2]Cases s1 of
+ | (Cons_acs l1 t1) =>
+ (abstract_sum_merge (abstract_sum_scalar l1 s2)
+ (abstract_sum_prod t1 s2))
+ | Nil_acs => Nil_acs
+ end.
+
+Fixpoint aspolynomial_normalize[p:aspolynomial] : abstract_sum :=
+ Cases p of
+ | (ASPvar i) => (Cons_acs (Cons_var i Nil_var) Nil_acs)
+ | ASP1 => (Cons_acs Nil_var Nil_acs)
+ | ASP0 => Nil_acs
+ | (ASPplus l r) => (abstract_sum_merge (aspolynomial_normalize l)
+ (aspolynomial_normalize r))
+ | (ASPmult l r) => (abstract_sum_prod (aspolynomial_normalize l)
+ (aspolynomial_normalize r))
+ end.
+
+
+
+Variable A : Type.
+Variable Aplus : A -> A -> A.
+Variable Amult : A -> A -> A.
+Variable Aone : A.
+Variable Azero : A.
+Variable Aeq : A -> A -> bool.
+Variable vm : (varmap A).
+Variable T : (Semi_Ring_Theory Aplus Amult Aone Azero Aeq).
+
+Fixpoint interp_asp [p:aspolynomial] : A :=
+ Cases p of
+ | (ASPvar i) => (interp_var Azero vm i)
+ | ASP0 => Azero
+ | ASP1 => Aone
+ | (ASPplus l r) => (Aplus (interp_asp l) (interp_asp r))
+ | (ASPmult l r) => (Amult (interp_asp l) (interp_asp r))
+ end.
+
+(* Local *) Definition iacs_aux := Fix iacs_aux{iacs_aux [a:A; s:abstract_sum] : A :=
+ Cases s of
+ | Nil_acs => a
+ | (Cons_acs l t) => (Aplus a (iacs_aux (interp_vl Amult Aone Azero vm l) t))
+ end}.
+
+Definition interp_acs [s:abstract_sum] : A :=
+ Cases s of
+ | (Cons_acs l t) => (iacs_aux (interp_vl Amult Aone Azero vm l) t)
+ | Nil_acs => Azero
+ end.
+
+Hint SR_plus_sym_T := Resolve (SR_plus_sym T).
+Hint SR_plus_assoc_T := Resolve (SR_plus_assoc T).
+Hint SR_plus_assoc2_T := Resolve (SR_plus_assoc2 T).
+Hint SR_mult_sym_T := Resolve (SR_mult_sym T).
+Hint SR_mult_assoc_T := Resolve (SR_mult_assoc T).
+Hint SR_mult_assoc2_T := Resolve (SR_mult_assoc2 T).
+Hint SR_plus_zero_left_T := Resolve (SR_plus_zero_left T).
+Hint SR_plus_zero_left2_T := Resolve (SR_plus_zero_left2 T).
+Hint SR_mult_one_left_T := Resolve (SR_mult_one_left T).
+Hint SR_mult_one_left2_T := Resolve (SR_mult_one_left2 T).
+Hint SR_mult_zero_left_T := Resolve (SR_mult_zero_left T).
+Hint SR_mult_zero_left2_T := Resolve (SR_mult_zero_left2 T).
+Hint SR_distr_left_T := Resolve (SR_distr_left T).
+Hint SR_distr_left2_T := Resolve (SR_distr_left2 T).
+Hint SR_plus_reg_left_T := Resolve (SR_plus_reg_left T).
+Hint SR_plus_permute_T := Resolve (SR_plus_permute T).
+Hint SR_mult_permute_T := Resolve (SR_mult_permute T).
+Hint SR_distr_right_T := Resolve (SR_distr_right T).
+Hint SR_distr_right2_T := Resolve (SR_distr_right2 T).
+Hint SR_mult_zero_right_T := Resolve (SR_mult_zero_right T).
+Hint SR_mult_zero_right2_T := Resolve (SR_mult_zero_right2 T).
+Hint SR_plus_zero_right_T := Resolve (SR_plus_zero_right T).
+Hint SR_plus_zero_right2_T := Resolve (SR_plus_zero_right2 T).
+Hint SR_mult_one_right_T := Resolve (SR_mult_one_right T).
+Hint SR_mult_one_right2_T := Resolve (SR_mult_one_right2 T).
+Hint SR_plus_reg_right_T := Resolve (SR_plus_reg_right T).
+Hints Resolve refl_equal sym_equal trans_equal.
+(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hints Immediate T.
+
+Remark iacs_aux_ok : (x:A)(s:abstract_sum)
+ (iacs_aux x s)==(Aplus x (interp_acs s)).
+Proof.
+ Induction s; Simpl; Intros.
+ Trivial.
+ Reflexivity.
+Save.
+
+Hint rew_iacs_aux : core := Extern 10 (eqT A ? ?) Rewrite iacs_aux_ok.
+
+Lemma abstract_varlist_insert_ok : (l:varlist)(s:abstract_sum)
+ (interp_acs (abstract_varlist_insert l s))
+ ==(Aplus (interp_vl Amult Aone Azero vm l) (interp_acs s)).
+
+ Induction s.
+ Trivial.
+
+ Simpl; Intros.
+ Elim (varlist_lt l v); Simpl.
+ EAuto.
+ Rewrite iacs_aux_ok.
+ Rewrite H; Auto.
+
+Save.
+
+Lemma abstract_sum_merge_ok : (x,y:abstract_sum)
+ (interp_acs (abstract_sum_merge x y))
+ ==(Aplus (interp_acs x) (interp_acs y)).
+
+Proof.
+ Induction x.
+ Trivial.
+ Induction y; Intros.
+
+ Auto.
+
+ Simpl; Elim (varlist_lt v v0); Simpl.
+ Repeat Rewrite iacs_aux_ok.
+ Rewrite H; Simpl; Auto.
+
+ Simpl in H0.
+ Repeat Rewrite iacs_aux_ok.
+ Rewrite H0. Simpl; Auto.
+Save.
+
+Lemma abstract_sum_scalar_ok : (l:varlist)(s:abstract_sum)
+ (interp_acs (abstract_sum_scalar l s))
+ == (Amult (interp_vl Amult Aone Azero vm l) (interp_acs s)).
+Proof.
+ Induction s.
+ Simpl; EAuto.
+
+ Simpl; Intros.
+ Rewrite iacs_aux_ok.
+ Rewrite abstract_varlist_insert_ok.
+ Rewrite H.
+ Rewrite (varlist_merge_ok A Aplus Amult Aone Azero Aeq vm T).
+ Auto.
+Save.
+
+Lemma abstract_sum_prod_ok : (x,y:abstract_sum)
+ (interp_acs (abstract_sum_prod x y))
+ == (Amult (interp_acs x) (interp_acs y)).
+
+Proof.
+ Induction x.
+ Intros; Simpl; EAuto.
+
+ NewDestruct y; Intros.
+
+ Simpl; Rewrite H; EAuto.
+
+ Unfold abstract_sum_prod; Fold abstract_sum_prod.
+ Rewrite abstract_sum_merge_ok.
+ Rewrite abstract_sum_scalar_ok.
+ Rewrite H; Simpl; Auto.
+Save.
+
+Theorem aspolynomial_normalize_ok : (x:aspolynomial)
+ (interp_asp x)==(interp_acs (aspolynomial_normalize x)).
+Proof.
+ Induction x; Simpl; Intros; Trivial.
+ Rewrite abstract_sum_merge_ok.
+ Rewrite H; Rewrite H0; EAuto.
+ Rewrite abstract_sum_prod_ok.
+ Rewrite H; Rewrite H0; EAuto.
+Save.
+
+End abstract_semi_rings.
+
+Section abstract_rings.
+
+(* In abstract polynomials there is no constants other
+ than 0 and 1. An abstract ring is a ring whose operations plus,
+ and mult are not functions but constructors. In other words,
+ when c1 and c2 are closed, (plus c1 c2) doesn't reduce to a closed
+ term. "closed" mean here "without plus and mult". *)
+
+(* this section is not parametrized by a (semi-)ring.
+ Nevertheless, they are two different types for semi-rings and rings
+ and there will be 2 correction theorems *)
+
+Inductive Type apolynomial :=
+ APvar : index -> apolynomial
+| AP0 : apolynomial
+| AP1 : apolynomial
+| APplus : apolynomial -> apolynomial -> apolynomial
+| APmult : apolynomial -> apolynomial -> apolynomial
+| APopp : apolynomial -> apolynomial
+.
+
+(* A canonical "abstract" sum is a list of varlist with the sign "+" or "-".
+ Invariant : the list is sorted and there is no varlist is present
+ with both signs. +x +x +x -x is forbidden => the canonical form is +x+x *)
+
+Inductive signed_sum : Type :=
+| Nil_varlist : signed_sum
+| Plus_varlist : varlist -> signed_sum -> signed_sum
+| Minus_varlist : varlist -> signed_sum -> signed_sum
+.
+
+Fixpoint signed_sum_merge [s1:signed_sum]
+ : signed_sum -> signed_sum :=
+Cases s1 of
+| (Plus_varlist l1 t1) =>
+ Fix ssm_aux{ssm_aux[s2:signed_sum] : signed_sum :=
+ Cases s2 of
+ | (Plus_varlist l2 t2) =>
+ if (varlist_lt l1 l2)
+ then (Plus_varlist l1 (signed_sum_merge t1 s2))
+ else (Plus_varlist l2 (ssm_aux t2))
+ | (Minus_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (signed_sum_merge t1 t2)
+ else if (varlist_lt l1 l2)
+ then (Plus_varlist l1 (signed_sum_merge t1 s2))
+ else (Minus_varlist l2 (ssm_aux t2))
+ | Nil_varlist => s1
+ end}
+| (Minus_varlist l1 t1) =>
+ Fix ssm_aux2{ssm_aux2[s2:signed_sum] : signed_sum :=
+ Cases s2 of
+ | (Plus_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (signed_sum_merge t1 t2)
+ else if (varlist_lt l1 l2)
+ then (Minus_varlist l1 (signed_sum_merge t1 s2))
+ else (Plus_varlist l2 (ssm_aux2 t2))
+ | (Minus_varlist l2 t2) =>
+ if (varlist_lt l1 l2)
+ then (Minus_varlist l1 (signed_sum_merge t1 s2))
+ else (Minus_varlist l2 (ssm_aux2 t2))
+ | Nil_varlist => s1
+ end}
+| Nil_varlist => [s2]s2
+end.
+
+Fixpoint plus_varlist_insert [l1:varlist; s2:signed_sum]
+ : signed_sum :=
+ Cases s2 of
+ | (Plus_varlist l2 t2) =>
+ if (varlist_lt l1 l2)
+ then (Plus_varlist l1 s2)
+ else (Plus_varlist l2 (plus_varlist_insert l1 t2))
+ | (Minus_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then t2
+ else if (varlist_lt l1 l2)
+ then (Plus_varlist l1 s2)
+ else (Minus_varlist l2 (plus_varlist_insert l1 t2))
+ | Nil_varlist => (Plus_varlist l1 Nil_varlist)
+ end.
+
+Fixpoint minus_varlist_insert [l1:varlist; s2:signed_sum]
+ : signed_sum :=
+ Cases s2 of
+ | (Plus_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then t2
+ else if (varlist_lt l1 l2)
+ then (Minus_varlist l1 s2)
+ else (Plus_varlist l2 (minus_varlist_insert l1 t2))
+ | (Minus_varlist l2 t2) =>
+ if (varlist_lt l1 l2)
+ then (Minus_varlist l1 s2)
+ else (Minus_varlist l2 (minus_varlist_insert l1 t2))
+ | Nil_varlist => (Minus_varlist l1 Nil_varlist)
+ end.
+
+Fixpoint signed_sum_opp [s:signed_sum] : signed_sum :=
+ Cases s of
+ | (Plus_varlist l2 t2) => (Minus_varlist l2 (signed_sum_opp t2))
+ | (Minus_varlist l2 t2) => (Plus_varlist l2 (signed_sum_opp t2))
+ | Nil_varlist => Nil_varlist
+ end.
+
+
+Fixpoint plus_sum_scalar [l1:varlist; s2:signed_sum]
+ : signed_sum :=
+ Cases s2 of
+ | (Plus_varlist l2 t2) => (plus_varlist_insert (varlist_merge l1 l2)
+ (plus_sum_scalar l1 t2))
+ | (Minus_varlist l2 t2) => (minus_varlist_insert (varlist_merge l1 l2)
+ (plus_sum_scalar l1 t2))
+ | Nil_varlist => Nil_varlist
+ end.
+
+Fixpoint minus_sum_scalar [l1:varlist; s2:signed_sum]
+ : signed_sum :=
+ Cases s2 of
+ | (Plus_varlist l2 t2) => (minus_varlist_insert (varlist_merge l1 l2)
+ (minus_sum_scalar l1 t2))
+ | (Minus_varlist l2 t2) => (plus_varlist_insert (varlist_merge l1 l2)
+ (minus_sum_scalar l1 t2))
+ | Nil_varlist => Nil_varlist
+ end.
+
+Fixpoint signed_sum_prod [s1:signed_sum]
+ : signed_sum -> signed_sum :=
+ [s2]Cases s1 of
+ | (Plus_varlist l1 t1) =>
+ (signed_sum_merge (plus_sum_scalar l1 s2)
+ (signed_sum_prod t1 s2))
+ | (Minus_varlist l1 t1) =>
+ (signed_sum_merge (minus_sum_scalar l1 s2)
+ (signed_sum_prod t1 s2))
+ | Nil_varlist => Nil_varlist
+ end.
+
+Fixpoint apolynomial_normalize[p:apolynomial] : signed_sum :=
+ Cases p of
+ | (APvar i) => (Plus_varlist (Cons_var i Nil_var) Nil_varlist)
+ | AP1 => (Plus_varlist Nil_var Nil_varlist)
+ | AP0 => Nil_varlist
+ | (APplus l r) => (signed_sum_merge (apolynomial_normalize l)
+ (apolynomial_normalize r))
+ | (APmult l r) => (signed_sum_prod (apolynomial_normalize l)
+ (apolynomial_normalize r))
+ | (APopp q) => (signed_sum_opp (apolynomial_normalize q))
+ end.
+
+
+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 vm : (varmap A).
+Variable T : (Ring_Theory Aplus Amult Aone Azero Aopp Aeq).
+
+(* Local *) Definition isacs_aux := Fix isacs_aux{isacs_aux [a:A; s:signed_sum] : A :=
+ Cases s of
+ | Nil_varlist => a
+ | (Plus_varlist l t) =>
+ (Aplus a (isacs_aux (interp_vl Amult Aone Azero vm l) t))
+ | (Minus_varlist l t) =>
+ (Aplus a (isacs_aux (Aopp (interp_vl Amult Aone Azero vm l)) t))
+ end}.
+
+Definition interp_sacs [s:signed_sum] : A :=
+ Cases s of
+ | (Plus_varlist l t) => (isacs_aux (interp_vl Amult Aone Azero vm l) t)
+ | (Minus_varlist l t) =>
+ (isacs_aux (Aopp (interp_vl Amult Aone Azero vm l)) t)
+ | Nil_varlist => Azero
+ end.
+
+Fixpoint interp_ap [p:apolynomial] : A :=
+ Cases p of
+ | (APvar i) => (interp_var Azero vm i)
+ | AP0 => Azero
+ | AP1 => Aone
+ | (APplus l r) => (Aplus (interp_ap l) (interp_ap r))
+ | (APmult l r) => (Amult (interp_ap l) (interp_ap r))
+ | (APopp q) => (Aopp (interp_ap q))
+ end.
+
+Hint Th_plus_sym_T := Resolve (Th_plus_sym T).
+Hint Th_plus_assoc_T := Resolve (Th_plus_assoc T).
+Hint Th_plus_assoc2_T := Resolve (Th_plus_assoc2 T).
+Hint Th_mult_sym_T := Resolve (Th_mult_sym T).
+Hint Th_mult_assoc_T := Resolve (Th_mult_assoc T).
+Hint Th_mult_assoc2_T := Resolve (Th_mult_assoc2 T).
+Hint Th_plus_zero_left_T := Resolve (Th_plus_zero_left T).
+Hint Th_plus_zero_left2_T := Resolve (Th_plus_zero_left2 T).
+Hint Th_mult_one_left_T := Resolve (Th_mult_one_left T).
+Hint Th_mult_one_left2_T := Resolve (Th_mult_one_left2 T).
+Hint Th_mult_zero_left_T := Resolve (Th_mult_zero_left T).
+Hint Th_mult_zero_left2_T := Resolve (Th_mult_zero_left2 T).
+Hint Th_distr_left_T := Resolve (Th_distr_left T).
+Hint Th_distr_left2_T := Resolve (Th_distr_left2 T).
+Hint Th_plus_reg_left_T := Resolve (Th_plus_reg_left T).
+Hint Th_plus_permute_T := Resolve (Th_plus_permute T).
+Hint Th_mult_permute_T := Resolve (Th_mult_permute T).
+Hint Th_distr_right_T := Resolve (Th_distr_right T).
+Hint Th_distr_right2_T := Resolve (Th_distr_right2 T).
+Hint Th_mult_zero_right2_T := Resolve (Th_mult_zero_right2 T).
+Hint Th_plus_zero_right_T := Resolve (Th_plus_zero_right T).
+Hint Th_plus_zero_right2_T := Resolve (Th_plus_zero_right2 T).
+Hint Th_mult_one_right_T := Resolve (Th_mult_one_right T).
+Hint Th_mult_one_right2_T := Resolve (Th_mult_one_right2 T).
+Hint Th_plus_reg_right_T := Resolve (Th_plus_reg_right T).
+Hints Resolve refl_equal sym_equal trans_equal.
+(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hints Immediate T.
+
+Lemma isacs_aux_ok : (x:A)(s:signed_sum)
+ (isacs_aux x s)==(Aplus x (interp_sacs s)).
+Proof.
+ Induction s; Simpl; Intros.
+ Trivial.
+ Reflexivity.
+ Reflexivity.
+Save.
+
+Hint rew_isacs_aux : core := Extern 10 (eqT A ? ?) Rewrite isacs_aux_ok.
+
+Tactic Definition Solve1 v v0 H H0 :=
+ Simpl; Elim (varlist_lt v v0); Simpl; Rewrite isacs_aux_ok;
+ [Rewrite H; Simpl; Auto
+ |Simpl in H0; Rewrite H0; Auto ].
+
+Lemma signed_sum_merge_ok : (x,y:signed_sum)
+ (interp_sacs (signed_sum_merge x y))
+ ==(Aplus (interp_sacs x) (interp_sacs y)).
+
+ Induction x.
+ Intro; Simpl; Auto.
+
+ Induction y; Intros.
+
+ Auto.
+
+ Solve1 v v0 H H0.
+
+ Simpl; Generalize (varlist_eq_prop v v0).
+ Elim (varlist_eq v v0); Simpl.
+
+ Intro Heq; Rewrite (Heq I).
+ Rewrite H.
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite (Th_plus_permute T).
+ Repeat Rewrite (Th_plus_assoc T).
+ Rewrite (Th_plus_sym T (Aopp (interp_vl Amult Aone Azero vm v0))
+ (interp_vl Amult Aone Azero vm v0)).
+ Rewrite (Th_opp_def T).
+ Rewrite (Th_plus_zero_left T).
+ Reflexivity.
+
+ Solve1 v v0 H H0.
+
+ Induction y; Intros.
+
+ Auto.
+
+ Simpl; Generalize (varlist_eq_prop v v0).
+ Elim (varlist_eq v v0); Simpl.
+
+ Intro Heq; Rewrite (Heq I).
+ Rewrite H.
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite (Th_plus_permute T).
+ Repeat Rewrite (Th_plus_assoc T).
+ Rewrite (Th_opp_def T).
+ Rewrite (Th_plus_zero_left T).
+ Reflexivity.
+
+ Solve1 v v0 H H0.
+
+ Solve1 v v0 H H0.
+
+Save.
+
+Tactic Definition Solve2 l v H :=
+ Elim (varlist_lt l v); Simpl; Rewrite isacs_aux_ok;
+ [ Auto
+ | Rewrite H; Auto ].
+
+Lemma plus_varlist_insert_ok : (l:varlist)(s:signed_sum)
+ (interp_sacs (plus_varlist_insert l s))
+ == (Aplus (interp_vl Amult Aone Azero vm l) (interp_sacs s)).
+Proof.
+
+ Induction s.
+ Trivial.
+
+ Simpl; Intros.
+ Solve2 l v H.
+
+ Simpl; Intros.
+ Generalize (varlist_eq_prop l v).
+ Elim (varlist_eq l v); Simpl.
+
+ Intro Heq; Rewrite (Heq I).
+ Repeat Rewrite isacs_aux_ok.
+ Repeat Rewrite (Th_plus_assoc T).
+ Rewrite (Th_opp_def T).
+ Rewrite (Th_plus_zero_left T).
+ Reflexivity.
+
+ Solve2 l v H.
+
+Save.
+
+Lemma minus_varlist_insert_ok : (l:varlist)(s:signed_sum)
+ (interp_sacs (minus_varlist_insert l s))
+ == (Aplus (Aopp (interp_vl Amult Aone Azero vm l)) (interp_sacs s)).
+Proof.
+
+ Induction s.
+ Trivial.
+
+ Simpl; Intros.
+ Generalize (varlist_eq_prop l v).
+ Elim (varlist_eq l v); Simpl.
+
+ Intro Heq; Rewrite (Heq I).
+ Repeat Rewrite isacs_aux_ok.
+ Repeat Rewrite (Th_plus_assoc T).
+ Rewrite (Th_plus_sym T (Aopp (interp_vl Amult Aone Azero vm v))
+ (interp_vl Amult Aone Azero vm v)).
+ Rewrite (Th_opp_def T).
+ Auto.
+
+ Simpl; Intros.
+ Solve2 l v H.
+
+ Simpl; Intros; Solve2 l v H.
+
+Save.
+
+Lemma signed_sum_opp_ok : (s:signed_sum)
+ (interp_sacs (signed_sum_opp s))
+ == (Aopp (interp_sacs s)).
+Proof.
+
+ Induction s; Simpl; Intros.
+
+ Symmetry; Apply (Th_opp_zero T).
+
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite H.
+ Rewrite (Th_plus_opp_opp T).
+ Reflexivity.
+
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite H.
+ Rewrite <- (Th_plus_opp_opp T).
+ Rewrite (Th_opp_opp T).
+ Reflexivity.
+
+Save.
+
+Lemma plus_sum_scalar_ok : (l:varlist)(s:signed_sum)
+ (interp_sacs (plus_sum_scalar l s))
+ == (Amult (interp_vl Amult Aone Azero vm l) (interp_sacs s)).
+Proof.
+
+ Induction s.
+ Trivial.
+
+ Simpl; Intros.
+ Rewrite plus_varlist_insert_ok.
+ Rewrite (varlist_merge_ok A Aplus Amult Aone Azero Aeq vm T).
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite H.
+ Auto.
+
+ Simpl; Intros.
+ Rewrite minus_varlist_insert_ok.
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite (varlist_merge_ok A Aplus Amult Aone Azero Aeq vm T).
+ Rewrite H.
+ Rewrite (Th_distr_right T).
+ Rewrite <- (Th_opp_mult_right T).
+ Reflexivity.
+
+Save.
+
+Lemma minus_sum_scalar_ok : (l:varlist)(s:signed_sum)
+ (interp_sacs (minus_sum_scalar l s))
+ == (Aopp (Amult (interp_vl Amult Aone Azero vm l) (interp_sacs s))).
+Proof.
+
+ Induction s; Simpl; Intros.
+
+ Rewrite (Th_mult_zero_right T); Symmetry; Apply (Th_opp_zero T).
+
+ Simpl; Intros.
+ Rewrite minus_varlist_insert_ok.
+ Rewrite (varlist_merge_ok A Aplus Amult Aone Azero Aeq vm T).
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite H.
+ Rewrite (Th_distr_right T).
+ Rewrite (Th_plus_opp_opp T).
+ Reflexivity.
+
+ Simpl; Intros.
+ Rewrite plus_varlist_insert_ok.
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite (varlist_merge_ok A Aplus Amult Aone Azero Aeq vm T).
+ Rewrite H.
+ Rewrite (Th_distr_right T).
+ Rewrite <- (Th_opp_mult_right T).
+ Rewrite <- (Th_plus_opp_opp T).
+ Rewrite (Th_opp_opp T).
+ Reflexivity.
+
+Save.
+
+Lemma signed_sum_prod_ok : (x,y:signed_sum)
+ (interp_sacs (signed_sum_prod x y)) ==
+ (Amult (interp_sacs x) (interp_sacs y)).
+Proof.
+
+ Induction x.
+
+ Simpl; EAuto 1.
+
+ Intros; Simpl.
+ Rewrite signed_sum_merge_ok.
+ Rewrite plus_sum_scalar_ok.
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite H.
+ Auto.
+
+ Intros; Simpl.
+ Repeat Rewrite isacs_aux_ok.
+ Rewrite signed_sum_merge_ok.
+ Rewrite minus_sum_scalar_ok.
+ Rewrite H.
+ Rewrite (Th_distr_left T).
+ Rewrite (Th_opp_mult_left T).
+ Reflexivity.
+
+Save.
+
+Theorem apolynomial_normalize_ok : (p:apolynomial)
+ (interp_sacs (apolynomial_normalize p))==(interp_ap p).
+Proof.
+ Induction p; Simpl; Auto 1.
+ Intros.
+ Rewrite signed_sum_merge_ok.
+ Rewrite H; Rewrite H0; Reflexivity.
+ Intros.
+ Rewrite signed_sum_prod_ok.
+ Rewrite H; Rewrite H0; Reflexivity.
+ Intros.
+ Rewrite signed_sum_opp_ok.
+ Rewrite H; Reflexivity.
+Save.
+
+End abstract_rings.
diff --git a/contrib7/ring/Ring_normalize.v b/contrib7/ring/Ring_normalize.v
new file mode 100644
index 00000000..1dbd9d56
--- /dev/null
+++ b/contrib7/ring/Ring_normalize.v
@@ -0,0 +1,893 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Ring_normalize.v,v 1.1.2.1 2004/07/16 19:30:18 herbelin Exp $ *)
+
+Require Ring_theory.
+Require Quote.
+
+Set Implicit Arguments.
+
+Lemma index_eq_prop: (n,m:index)(Is_true (index_eq n m)) -> n=m.
+Proof.
+ Intros.
+ Apply Quote.index_eq_prop.
+ Generalize H.
+ Case (index_eq n m); Simpl; Trivial; Intros.
+ Contradiction.
+Save.
+
+Section semi_rings.
+
+Variable A : Type.
+Variable Aplus : A -> A -> A.
+Variable Amult : A -> A -> A.
+Variable Aone : A.
+Variable Azero : A.
+Variable Aeq : A -> A -> bool.
+
+(* Section definitions. *)
+
+
+(******************************************)
+(* Normal abtract Polynomials *)
+(******************************************)
+(* DEFINITIONS :
+- A varlist is a sorted product of one or more variables : x, x*y*z
+- A monom is a constant, a varlist or the product of a constant by a varlist
+ variables. 2*x*y, x*y*z, 3 are monoms : 2*3, x*3*y, 4*x*3 are NOT.
+- A canonical sum is either a monom or an ordered sum of monoms
+ (the order on monoms is defined later)
+- A normal polynomial it either a constant or a canonical sum or a constant
+ plus a canonical sum
+*)
+
+(* varlist is isomorphic to (list var), but we built a special inductive
+ for efficiency *)
+Inductive varlist : Type :=
+| Nil_var : varlist
+| Cons_var : index -> varlist -> varlist
+.
+
+Inductive canonical_sum : Type :=
+| Nil_monom : canonical_sum
+| Cons_monom : A -> varlist -> canonical_sum -> canonical_sum
+| Cons_varlist : varlist -> canonical_sum -> canonical_sum
+.
+
+(* Order on monoms *)
+
+(* That's the lexicographic order on varlist, extended by :
+ - A constant is less than every monom
+ - The relation between two varlist is preserved by multiplication by a
+ constant.
+
+ Examples :
+ 3 < x < y
+ x*y < x*y*y*z
+ 2*x*y < x*y*y*z
+ x*y < 54*x*y*y*z
+ 4*x*y < 59*x*y*y*z
+*)
+
+Fixpoint varlist_eq [x,y:varlist] : bool :=
+ Cases x y of
+ | Nil_var Nil_var => true
+ | (Cons_var i xrest) (Cons_var j yrest) =>
+ (andb (index_eq i j) (varlist_eq xrest yrest))
+ | _ _ => false
+ end.
+
+Fixpoint varlist_lt [x,y:varlist] : bool :=
+ Cases x y of
+ | Nil_var (Cons_var _ _) => true
+ | (Cons_var i xrest) (Cons_var j yrest) =>
+ if (index_lt i j) then true
+ else (andb (index_eq i j) (varlist_lt xrest yrest))
+ | _ _ => false
+ end.
+
+(* merges two variables lists *)
+Fixpoint varlist_merge [l1:varlist] : varlist -> varlist :=
+ Cases l1 of
+ | (Cons_var v1 t1) =>
+ Fix vm_aux {vm_aux [l2:varlist] : varlist :=
+ Cases l2 of
+ | (Cons_var v2 t2) =>
+ if (index_lt v1 v2)
+ then (Cons_var v1 (varlist_merge t1 l2))
+ else (Cons_var v2 (vm_aux t2))
+ | Nil_var => l1
+ end}
+ | Nil_var => [l2]l2
+ end.
+
+(* returns the sum of two canonical sums *)
+Fixpoint canonical_sum_merge [s1:canonical_sum]
+ : canonical_sum -> canonical_sum :=
+Cases s1 of
+| (Cons_monom c1 l1 t1) =>
+ Fix csm_aux{csm_aux[s2:canonical_sum] : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 c2) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 (canonical_sum_merge t1 s2))
+ else (Cons_monom c2 l2 (csm_aux t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 Aone) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 (canonical_sum_merge t1 s2))
+ else (Cons_varlist l2 (csm_aux t2))
+ | Nil_monom => s1
+ end}
+| (Cons_varlist l1 t1) =>
+ Fix csm_aux2{csm_aux2[s2:canonical_sum] : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone c2) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 (canonical_sum_merge t1 s2))
+ else (Cons_monom c2 l2 (csm_aux2 t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone Aone) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 (canonical_sum_merge t1 s2))
+ else (Cons_varlist l2 (csm_aux2 t2))
+ | Nil_monom => s1
+ end}
+| Nil_monom => [s2]s2
+end.
+
+(* Insertion of a monom in a canonical sum *)
+Fixpoint monom_insert [c1:A; l1:varlist; s2 : canonical_sum]
+ : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 c2) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 s2)
+ else (Cons_monom c2 l2 (monom_insert c1 l1 t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 Aone) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 s2)
+ else (Cons_varlist l2 (monom_insert c1 l1 t2))
+ | Nil_monom => (Cons_monom c1 l1 Nil_monom)
+ end.
+
+Fixpoint varlist_insert [l1:varlist; s2:canonical_sum]
+ : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone c2) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 s2)
+ else (Cons_monom c2 l2 (varlist_insert l1 t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone Aone) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 s2)
+ else (Cons_varlist l2 (varlist_insert l1 t2))
+ | Nil_monom => (Cons_varlist l1 Nil_monom)
+ end.
+
+(* Computes c0*s *)
+Fixpoint canonical_sum_scalar [c0:A; s:canonical_sum] : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ (Cons_monom (Amult c0 c) l (canonical_sum_scalar c0 t))
+ | (Cons_varlist l t) =>
+ (Cons_monom c0 l (canonical_sum_scalar c0 t))
+ | Nil_monom => Nil_monom
+ end.
+
+(* Computes l0*s *)
+Fixpoint canonical_sum_scalar2 [l0:varlist; s:canonical_sum]
+ : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ (monom_insert c (varlist_merge l0 l) (canonical_sum_scalar2 l0 t))
+ | (Cons_varlist l t) =>
+ (varlist_insert (varlist_merge l0 l) (canonical_sum_scalar2 l0 t))
+ | Nil_monom => Nil_monom
+ end.
+
+(* Computes c0*l0*s *)
+Fixpoint canonical_sum_scalar3 [c0:A;l0:varlist; s:canonical_sum]
+ : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ (monom_insert (Amult c0 c) (varlist_merge l0 l)
+ (canonical_sum_scalar3 c0 l0 t))
+ | (Cons_varlist l t) =>
+ (monom_insert c0 (varlist_merge l0 l)
+ (canonical_sum_scalar3 c0 l0 t))
+ | Nil_monom => Nil_monom
+ end.
+
+(* returns the product of two canonical sums *)
+Fixpoint canonical_sum_prod [s1:canonical_sum]
+ : canonical_sum -> canonical_sum :=
+ [s2]Cases s1 of
+ | (Cons_monom c1 l1 t1) =>
+ (canonical_sum_merge (canonical_sum_scalar3 c1 l1 s2)
+ (canonical_sum_prod t1 s2))
+ | (Cons_varlist l1 t1) =>
+ (canonical_sum_merge (canonical_sum_scalar2 l1 s2)
+ (canonical_sum_prod t1 s2))
+ | Nil_monom => Nil_monom
+ end.
+
+(* The type to represent concrete semi-ring polynomials *)
+Inductive Type spolynomial :=
+ SPvar : index -> spolynomial
+| SPconst : A -> spolynomial
+| SPplus : spolynomial -> spolynomial -> spolynomial
+| SPmult : spolynomial -> spolynomial -> spolynomial.
+
+Fixpoint spolynomial_normalize[p:spolynomial] : canonical_sum :=
+ Cases p of
+ | (SPvar i) => (Cons_varlist (Cons_var i Nil_var) Nil_monom)
+ | (SPconst c) => (Cons_monom c Nil_var Nil_monom)
+ | (SPplus l r) => (canonical_sum_merge (spolynomial_normalize l)
+ (spolynomial_normalize r))
+ | (SPmult l r) => (canonical_sum_prod (spolynomial_normalize l)
+ (spolynomial_normalize r))
+ end.
+
+(* Deletion of useless 0 and 1 in canonical sums *)
+Fixpoint canonical_sum_simplify [ s:canonical_sum] : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ if (Aeq c Azero)
+ then (canonical_sum_simplify t)
+ else if (Aeq c Aone)
+ then (Cons_varlist l (canonical_sum_simplify t))
+ else (Cons_monom c l (canonical_sum_simplify t))
+ | (Cons_varlist l t) => (Cons_varlist l (canonical_sum_simplify t))
+ | Nil_monom => Nil_monom
+ end.
+
+Definition spolynomial_simplify :=
+ [x:spolynomial](canonical_sum_simplify (spolynomial_normalize x)).
+
+(* End definitions. *)
+
+(* Section interpretation. *)
+
+(*** Here a variable map is defined and the interpetation of a spolynom
+ acording to a certain variables map. Once again the choosen definition
+ is generic and could be changed ****)
+
+Variable vm : (varmap A).
+
+(* Interpretation of list of variables
+ * [x1; ... ; xn ] is interpreted as (find v x1)* ... *(find v xn)
+ * The unbound variables are mapped to 0. Normally this case sould
+ * never occur. Since we want only to prove correctness theorems, which form
+ * is : for any varmap and any spolynom ... this is a safe and pain-saving
+ * choice *)
+Definition interp_var [i:index] := (varmap_find Azero i vm).
+
+(* Local *) Definition ivl_aux := Fix ivl_aux {ivl_aux[x:index; t:varlist] : A :=
+ Cases t of
+ | Nil_var => (interp_var x)
+ | (Cons_var x' t') => (Amult (interp_var x) (ivl_aux x' t'))
+ end}.
+
+Definition interp_vl := [l:varlist]
+ Cases l of
+ | Nil_var => Aone
+ | (Cons_var x t) => (ivl_aux x t)
+ end.
+
+(* Local *) Definition interp_m := [c:A][l:varlist]
+ Cases l of
+ | Nil_var => c
+ | (Cons_var x t) =>
+ (Amult c (ivl_aux x t))
+ end.
+
+(* Local *) Definition ics_aux := Fix ics_aux{ics_aux[a:A; s:canonical_sum] : A :=
+ Cases s of
+ | Nil_monom => a
+ | (Cons_varlist l t) => (Aplus a (ics_aux (interp_vl l) t))
+ | (Cons_monom c l t) => (Aplus a (ics_aux (interp_m c l) t))
+ end}.
+
+(* Interpretation of a canonical sum *)
+Definition interp_cs : canonical_sum -> A :=
+ [s]Cases s of
+ | Nil_monom => Azero
+ | (Cons_varlist l t) =>
+ (ics_aux (interp_vl l) t)
+ | (Cons_monom c l t) =>
+ (ics_aux (interp_m c l) t)
+ end.
+
+Fixpoint interp_sp [p:spolynomial] : A :=
+ Cases p of
+ (SPconst c) => c
+ | (SPvar i) => (interp_var i)
+ | (SPplus p1 p2) => (Aplus (interp_sp p1) (interp_sp p2))
+ | (SPmult p1 p2) => (Amult (interp_sp p1) (interp_sp p2))
+ end.
+
+
+(* End interpretation. *)
+
+Unset Implicit Arguments.
+
+(* Section properties. *)
+
+Variable T : (Semi_Ring_Theory Aplus Amult Aone Azero Aeq).
+
+Hint SR_plus_sym_T := Resolve (SR_plus_sym T).
+Hint SR_plus_assoc_T := Resolve (SR_plus_assoc T).
+Hint SR_plus_assoc2_T := Resolve (SR_plus_assoc2 T).
+Hint SR_mult_sym_T := Resolve (SR_mult_sym T).
+Hint SR_mult_assoc_T := Resolve (SR_mult_assoc T).
+Hint SR_mult_assoc2_T := Resolve (SR_mult_assoc2 T).
+Hint SR_plus_zero_left_T := Resolve (SR_plus_zero_left T).
+Hint SR_plus_zero_left2_T := Resolve (SR_plus_zero_left2 T).
+Hint SR_mult_one_left_T := Resolve (SR_mult_one_left T).
+Hint SR_mult_one_left2_T := Resolve (SR_mult_one_left2 T).
+Hint SR_mult_zero_left_T := Resolve (SR_mult_zero_left T).
+Hint SR_mult_zero_left2_T := Resolve (SR_mult_zero_left2 T).
+Hint SR_distr_left_T := Resolve (SR_distr_left T).
+Hint SR_distr_left2_T := Resolve (SR_distr_left2 T).
+Hint SR_plus_reg_left_T := Resolve (SR_plus_reg_left T).
+Hint SR_plus_permute_T := Resolve (SR_plus_permute T).
+Hint SR_mult_permute_T := Resolve (SR_mult_permute T).
+Hint SR_distr_right_T := Resolve (SR_distr_right T).
+Hint SR_distr_right2_T := Resolve (SR_distr_right2 T).
+Hint SR_mult_zero_right_T := Resolve (SR_mult_zero_right T).
+Hint SR_mult_zero_right2_T := Resolve (SR_mult_zero_right2 T).
+Hint SR_plus_zero_right_T := Resolve (SR_plus_zero_right T).
+Hint SR_plus_zero_right2_T := Resolve (SR_plus_zero_right2 T).
+Hint SR_mult_one_right_T := Resolve (SR_mult_one_right T).
+Hint SR_mult_one_right2_T := Resolve (SR_mult_one_right2 T).
+Hint SR_plus_reg_right_T := Resolve (SR_plus_reg_right T).
+Hints Resolve refl_equal sym_equal trans_equal.
+(* Hints Resolve refl_eqT sym_eqT trans_eqT. *)
+Hints Immediate T.
+
+Lemma varlist_eq_prop : (x,y:varlist)
+ (Is_true (varlist_eq x y))->x==y.
+Proof.
+ Induction x; Induction y; Contradiction Orelse Try Reflexivity.
+ Simpl; Intros.
+ Generalize (andb_prop2 ? ? H1); Intros; Elim H2; Intros.
+ Rewrite (index_eq_prop H3); Rewrite (H v0 H4); Reflexivity.
+Save.
+
+Remark ivl_aux_ok : (v:varlist)(i:index)
+ (ivl_aux i v)==(Amult (interp_var i) (interp_vl v)).
+Proof.
+ Induction v; Simpl; Intros.
+ Trivial.
+ Rewrite H; Trivial.
+Save.
+
+Lemma varlist_merge_ok : (x,y:varlist)
+ (interp_vl (varlist_merge x y))
+ ==(Amult (interp_vl x) (interp_vl y)).
+Proof.
+ Induction x.
+ Simpl; Trivial.
+ Induction y.
+ Simpl; Trivial.
+ Simpl; Intros.
+ Elim (index_lt i i0); Simpl; Intros.
+
+ Repeat Rewrite ivl_aux_ok.
+ Rewrite H. Simpl.
+ Rewrite ivl_aux_ok.
+ EAuto.
+
+ Repeat Rewrite ivl_aux_ok.
+ Rewrite H0.
+ Rewrite ivl_aux_ok.
+ EAuto.
+Save.
+
+Remark ics_aux_ok : (x:A)(s:canonical_sum)
+ (ics_aux x s)==(Aplus x (interp_cs s)).
+Proof.
+ Induction s; Simpl; Intros.
+ Trivial.
+ Reflexivity.
+ Reflexivity.
+Save.
+
+Remark interp_m_ok : (x:A)(l:varlist)
+ (interp_m x l)==(Amult x (interp_vl l)).
+Proof.
+ NewDestruct l.
+ Simpl; Trivial.
+ Reflexivity.
+Save.
+
+Lemma canonical_sum_merge_ok : (x,y:canonical_sum)
+ (interp_cs (canonical_sum_merge x y))
+ ==(Aplus (interp_cs x) (interp_cs y)).
+
+Induction x; Simpl.
+Trivial.
+
+Induction y; Simpl; Intros.
+(* monom and nil *)
+EAuto.
+
+(* monom and monom *)
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl; Repeat Rewrite ics_aux_ok; Rewrite H.
+Repeat Rewrite interp_m_ok.
+Rewrite (SR_distr_left T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Apply congr_eqT with f:=(Aplus (Amult a (interp_vl v0))).
+Trivial.
+
+Elim (varlist_lt v v0); Simpl.
+Repeat Rewrite ics_aux_ok.
+Rewrite H; Simpl; Rewrite ics_aux_ok; EAuto.
+
+Rewrite ics_aux_ok; Rewrite H0; Repeat Rewrite ics_aux_ok; Simpl; EAuto.
+
+(* monom and varlist *)
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl; Repeat Rewrite ics_aux_ok; Rewrite H.
+Repeat Rewrite interp_m_ok.
+Rewrite (SR_distr_left T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Apply congr_eqT with f:=(Aplus (Amult a (interp_vl v0))).
+Rewrite (SR_mult_one_left T).
+Trivial.
+
+Elim (varlist_lt v v0); Simpl.
+Repeat Rewrite ics_aux_ok.
+Rewrite H; Simpl; Rewrite ics_aux_ok; EAuto.
+Rewrite ics_aux_ok; Rewrite H0; Repeat Rewrite ics_aux_ok; Simpl; EAuto.
+
+Induction y; Simpl; Intros.
+(* varlist and nil *)
+Trivial.
+
+(* varlist and monom *)
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl; Repeat Rewrite ics_aux_ok; Rewrite H.
+Repeat Rewrite interp_m_ok.
+Rewrite (SR_distr_left T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Rewrite (SR_mult_one_left T).
+Apply congr_eqT with f:=(Aplus (interp_vl v0)).
+Trivial.
+
+Elim (varlist_lt v v0); Simpl.
+Repeat Rewrite ics_aux_ok.
+Rewrite H; Simpl; Rewrite ics_aux_ok; EAuto.
+Rewrite ics_aux_ok; Rewrite H0; Repeat Rewrite ics_aux_ok; Simpl; EAuto.
+
+(* varlist and varlist *)
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl; Repeat Rewrite ics_aux_ok; Rewrite H.
+Repeat Rewrite interp_m_ok.
+Rewrite (SR_distr_left T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Rewrite (SR_mult_one_left T).
+Apply congr_eqT with f:=(Aplus (interp_vl v0)).
+Trivial.
+
+Elim (varlist_lt v v0); Simpl.
+Repeat Rewrite ics_aux_ok.
+Rewrite H; Simpl; Rewrite ics_aux_ok; EAuto.
+Rewrite ics_aux_ok; Rewrite H0; Repeat Rewrite ics_aux_ok; Simpl; EAuto.
+Save.
+
+Lemma monom_insert_ok: (a:A)(l:varlist)(s:canonical_sum)
+ (interp_cs (monom_insert a l s))
+ == (Aplus (Amult a (interp_vl l)) (interp_cs s)).
+Intros; Generalize s; Induction s0.
+
+Simpl; Rewrite interp_m_ok; Trivial.
+
+Simpl; Intros.
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl; Rewrite interp_m_ok;
+ Repeat Rewrite ics_aux_ok; Rewrite interp_m_ok;
+ Rewrite (SR_distr_left T); EAuto.
+Elim (varlist_lt l v); Simpl;
+[ Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok; EAuto
+| Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok;
+ Rewrite H; Rewrite ics_aux_ok; EAuto].
+
+Simpl; Intros.
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl; Rewrite interp_m_ok;
+ Repeat Rewrite ics_aux_ok;
+ Rewrite (SR_distr_left T); Rewrite (SR_mult_one_left T); EAuto.
+Elim (varlist_lt l v); Simpl;
+[ Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok; EAuto
+| Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok;
+ Rewrite H; Rewrite ics_aux_ok; EAuto].
+Save.
+
+Lemma varlist_insert_ok :
+ (l:varlist)(s:canonical_sum)
+ (interp_cs (varlist_insert l s))
+ == (Aplus (interp_vl l) (interp_cs s)).
+Intros; Generalize s; Induction s0.
+
+Simpl; Trivial.
+
+Simpl; Intros.
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl; Rewrite interp_m_ok;
+ Repeat Rewrite ics_aux_ok; Rewrite interp_m_ok;
+ Rewrite (SR_distr_left T); Rewrite (SR_mult_one_left T); EAuto.
+Elim (varlist_lt l v); Simpl;
+[ Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok; EAuto
+| Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok;
+ Rewrite H; Rewrite ics_aux_ok; EAuto].
+
+Simpl; Intros.
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl; Rewrite interp_m_ok;
+ Repeat Rewrite ics_aux_ok;
+ Rewrite (SR_distr_left T); Rewrite (SR_mult_one_left T); EAuto.
+Elim (varlist_lt l v); Simpl;
+[ Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok; EAuto
+| Repeat Rewrite interp_m_ok; Rewrite ics_aux_ok;
+ Rewrite H; Rewrite ics_aux_ok; EAuto].
+Save.
+
+Lemma canonical_sum_scalar_ok : (a:A)(s:canonical_sum)
+ (interp_cs (canonical_sum_scalar a s))
+ ==(Amult a (interp_cs s)).
+Induction s.
+Simpl; EAuto.
+
+Simpl; Intros.
+Repeat Rewrite ics_aux_ok.
+Repeat Rewrite interp_m_ok.
+Rewrite H.
+Rewrite (SR_distr_right T).
+Repeat Rewrite <- (SR_mult_assoc T).
+Reflexivity.
+
+Simpl; Intros.
+Repeat Rewrite ics_aux_ok.
+Repeat Rewrite interp_m_ok.
+Rewrite H.
+Rewrite (SR_distr_right T).
+Repeat Rewrite <- (SR_mult_assoc T).
+Reflexivity.
+Save.
+
+Lemma canonical_sum_scalar2_ok : (l:varlist; s:canonical_sum)
+ (interp_cs (canonical_sum_scalar2 l s))
+ ==(Amult (interp_vl l) (interp_cs s)).
+Induction s.
+Simpl; Trivial.
+
+Simpl; Intros.
+Rewrite monom_insert_ok.
+Repeat Rewrite ics_aux_ok.
+Repeat Rewrite interp_m_ok.
+Rewrite H.
+Rewrite varlist_merge_ok.
+Repeat Rewrite (SR_distr_right T).
+Repeat Rewrite <- (SR_mult_assoc T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Rewrite (SR_mult_permute T a (interp_vl l) (interp_vl v)).
+Reflexivity.
+
+Simpl; Intros.
+Rewrite varlist_insert_ok.
+Repeat Rewrite ics_aux_ok.
+Repeat Rewrite interp_m_ok.
+Rewrite H.
+Rewrite varlist_merge_ok.
+Repeat Rewrite (SR_distr_right T).
+Repeat Rewrite <- (SR_mult_assoc T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Reflexivity.
+Save.
+
+Lemma canonical_sum_scalar3_ok : (c:A; l:varlist; s:canonical_sum)
+ (interp_cs (canonical_sum_scalar3 c l s))
+ ==(Amult c (Amult (interp_vl l) (interp_cs s))).
+Induction s.
+Simpl; Repeat Rewrite (SR_mult_zero_right T); Reflexivity.
+
+Simpl; Intros.
+Rewrite monom_insert_ok.
+Repeat Rewrite ics_aux_ok.
+Repeat Rewrite interp_m_ok.
+Rewrite H.
+Rewrite varlist_merge_ok.
+Repeat Rewrite (SR_distr_right T).
+Repeat Rewrite <- (SR_mult_assoc T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Rewrite (SR_mult_permute T a (interp_vl l) (interp_vl v)).
+Reflexivity.
+
+Simpl; Intros.
+Rewrite monom_insert_ok.
+Repeat Rewrite ics_aux_ok.
+Repeat Rewrite interp_m_ok.
+Rewrite H.
+Rewrite varlist_merge_ok.
+Repeat Rewrite (SR_distr_right T).
+Repeat Rewrite <- (SR_mult_assoc T).
+Repeat Rewrite <- (SR_plus_assoc T).
+Rewrite (SR_mult_permute T c (interp_vl l) (interp_vl v)).
+Reflexivity.
+Save.
+
+Lemma canonical_sum_prod_ok : (x,y:canonical_sum)
+ (interp_cs (canonical_sum_prod x y))
+ ==(Amult (interp_cs x) (interp_cs y)).
+Induction x; Simpl; Intros.
+Trivial.
+
+Rewrite canonical_sum_merge_ok.
+Rewrite canonical_sum_scalar3_ok.
+Rewrite ics_aux_ok.
+Rewrite interp_m_ok.
+Rewrite H.
+Rewrite (SR_mult_assoc T a (interp_vl v) (interp_cs y)).
+Symmetry.
+EAuto.
+
+Rewrite canonical_sum_merge_ok.
+Rewrite canonical_sum_scalar2_ok.
+Rewrite ics_aux_ok.
+Rewrite H.
+Trivial.
+Save.
+
+Theorem spolynomial_normalize_ok : (p:spolynomial)
+ (interp_cs (spolynomial_normalize p)) == (interp_sp p).
+Induction p; Simpl; Intros.
+
+Reflexivity.
+Reflexivity.
+
+Rewrite canonical_sum_merge_ok.
+Rewrite H; Rewrite H0.
+Reflexivity.
+
+Rewrite canonical_sum_prod_ok.
+Rewrite H; Rewrite H0.
+Reflexivity.
+Save.
+
+Lemma canonical_sum_simplify_ok : (s:canonical_sum)
+ (interp_cs (canonical_sum_simplify s)) == (interp_cs s).
+Induction s.
+
+Reflexivity.
+
+(* cons_monom *)
+Simpl; Intros.
+Generalize (SR_eq_prop T 8!a 9!Azero).
+Elim (Aeq a Azero).
+Intro Heq; Rewrite (Heq I).
+Rewrite H.
+Rewrite ics_aux_ok.
+Rewrite interp_m_ok.
+Rewrite (SR_mult_zero_left T).
+Trivial.
+
+Intros; Simpl.
+Generalize (SR_eq_prop T 8!a 9!Aone).
+Elim (Aeq a Aone).
+Intro Heq; Rewrite (Heq I).
+Simpl.
+Repeat Rewrite ics_aux_ok.
+Rewrite interp_m_ok.
+Rewrite H.
+Rewrite (SR_mult_one_left T).
+Reflexivity.
+
+Simpl.
+Repeat Rewrite ics_aux_ok.
+Rewrite interp_m_ok.
+Rewrite H.
+Reflexivity.
+
+(* cons_varlist *)
+Simpl; Intros.
+Repeat Rewrite ics_aux_ok.
+Rewrite H.
+Reflexivity.
+
+Save.
+
+Theorem spolynomial_simplify_ok : (p:spolynomial)
+ (interp_cs (spolynomial_simplify p)) == (interp_sp p).
+Intro.
+Unfold spolynomial_simplify.
+Rewrite canonical_sum_simplify_ok.
+Apply spolynomial_normalize_ok.
+Save.
+
+(* End properties. *)
+End semi_rings.
+
+Implicits Cons_varlist.
+Implicits Cons_monom.
+Implicits SPconst.
+Implicits SPplus.
+Implicits SPmult.
+
+Section rings.
+
+(* Here the coercion between Ring and Semi-Ring will be useful *)
+
+Set Implicit Arguments.
+
+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 vm : (varmap A).
+Variable T : (Ring_Theory Aplus Amult Aone Azero Aopp Aeq).
+
+Hint Th_plus_sym_T := Resolve (Th_plus_sym T).
+Hint Th_plus_assoc_T := Resolve (Th_plus_assoc T).
+Hint Th_plus_assoc2_T := Resolve (Th_plus_assoc2 T).
+Hint Th_mult_sym_T := Resolve (Th_mult_sym T).
+Hint Th_mult_assoc_T := Resolve (Th_mult_assoc T).
+Hint Th_mult_assoc2_T := Resolve (Th_mult_assoc2 T).
+Hint Th_plus_zero_left_T := Resolve (Th_plus_zero_left T).
+Hint Th_plus_zero_left2_T := Resolve (Th_plus_zero_left2 T).
+Hint Th_mult_one_left_T := Resolve (Th_mult_one_left T).
+Hint Th_mult_one_left2_T := Resolve (Th_mult_one_left2 T).
+Hint Th_mult_zero_left_T := Resolve (Th_mult_zero_left T).
+Hint Th_mult_zero_left2_T := Resolve (Th_mult_zero_left2 T).
+Hint Th_distr_left_T := Resolve (Th_distr_left T).
+Hint Th_distr_left2_T := Resolve (Th_distr_left2 T).
+Hint Th_plus_reg_left_T := Resolve (Th_plus_reg_left T).
+Hint Th_plus_permute_T := Resolve (Th_plus_permute T).
+Hint Th_mult_permute_T := Resolve (Th_mult_permute T).
+Hint Th_distr_right_T := Resolve (Th_distr_right T).
+Hint Th_distr_right2_T := Resolve (Th_distr_right2 T).
+Hint Th_mult_zero_right_T := Resolve (Th_mult_zero_right T).
+Hint Th_mult_zero_right2_T := Resolve (Th_mult_zero_right2 T).
+Hint Th_plus_zero_right_T := Resolve (Th_plus_zero_right T).
+Hint Th_plus_zero_right2_T := Resolve (Th_plus_zero_right2 T).
+Hint Th_mult_one_right_T := Resolve (Th_mult_one_right T).
+Hint Th_mult_one_right2_T := Resolve (Th_mult_one_right2 T).
+Hint Th_plus_reg_right_T := Resolve (Th_plus_reg_right T).
+Hints Resolve refl_equal sym_equal trans_equal.
+(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hints Immediate T.
+
+(*** Definitions *)
+
+Inductive Type polynomial :=
+ Pvar : index -> polynomial
+| Pconst : A -> polynomial
+| Pplus : polynomial -> polynomial -> polynomial
+| Pmult : polynomial -> polynomial -> polynomial
+| Popp : polynomial -> polynomial.
+
+Fixpoint polynomial_normalize [x:polynomial] : (canonical_sum A) :=
+ Cases x of
+ (Pplus l r) => (canonical_sum_merge Aplus Aone
+ (polynomial_normalize l)
+ (polynomial_normalize r))
+ | (Pmult l r) => (canonical_sum_prod Aplus Amult Aone
+ (polynomial_normalize l)
+ (polynomial_normalize r))
+ | (Pconst c) => (Cons_monom c Nil_var (Nil_monom A))
+ | (Pvar i) => (Cons_varlist (Cons_var i Nil_var) (Nil_monom A))
+ | (Popp p) => (canonical_sum_scalar3 Aplus Amult Aone
+ (Aopp Aone) Nil_var
+ (polynomial_normalize p))
+ end.
+
+Definition polynomial_simplify :=
+ [x:polynomial](canonical_sum_simplify Aone Azero Aeq
+ (polynomial_normalize x)).
+
+Fixpoint spolynomial_of [x:polynomial] : (spolynomial A) :=
+ Cases x of
+ (Pplus l r) => (SPplus (spolynomial_of l) (spolynomial_of r))
+ | (Pmult l r) => (SPmult (spolynomial_of l) (spolynomial_of r))
+ | (Pconst c) => (SPconst c)
+ | (Pvar i) => (SPvar A i)
+ | (Popp p) => (SPmult (SPconst (Aopp Aone)) (spolynomial_of p))
+ end.
+
+(*** Interpretation *)
+
+Fixpoint interp_p [p:polynomial] : A :=
+ Cases p of
+ (Pconst c) => c
+ | (Pvar i) => (varmap_find Azero i vm)
+ | (Pplus p1 p2) => (Aplus (interp_p p1) (interp_p p2))
+ | (Pmult p1 p2) => (Amult (interp_p p1) (interp_p p2))
+ | (Popp p1) => (Aopp (interp_p p1))
+ end.
+
+(*** Properties *)
+
+Unset Implicit Arguments.
+
+Lemma spolynomial_of_ok : (p:polynomial)
+ (interp_p p)==(interp_sp Aplus Amult Azero vm (spolynomial_of p)).
+Induction p; Reflexivity Orelse (Simpl; Intros).
+Rewrite H; Rewrite H0; Reflexivity.
+Rewrite H; Rewrite H0; Reflexivity.
+Rewrite H.
+Rewrite (Th_opp_mult_left2 T).
+Rewrite (Th_mult_one_left T).
+Reflexivity.
+Save.
+
+Theorem polynomial_normalize_ok : (p:polynomial)
+ (polynomial_normalize p)
+ ==(spolynomial_normalize Aplus Amult Aone (spolynomial_of p)).
+Induction p; Reflexivity Orelse (Simpl; Intros).
+Rewrite H; Rewrite H0; Reflexivity.
+Rewrite H; Rewrite H0; Reflexivity.
+Rewrite H; Simpl.
+Elim (canonical_sum_scalar3 Aplus Amult Aone (Aopp Aone) Nil_var
+ (spolynomial_normalize Aplus Amult Aone (spolynomial_of p0)));
+[ Reflexivity
+| Simpl; Intros; Rewrite H0; Reflexivity
+| Simpl; Intros; Rewrite H0; Reflexivity ].
+Save.
+
+Theorem polynomial_simplify_ok : (p:polynomial)
+ (interp_cs Aplus Amult Aone Azero vm (polynomial_simplify p))
+ ==(interp_p p).
+Intro.
+Unfold polynomial_simplify.
+Rewrite spolynomial_of_ok.
+Rewrite polynomial_normalize_ok.
+Rewrite (canonical_sum_simplify_ok A Aplus Amult Aone Azero Aeq vm T).
+Rewrite (spolynomial_normalize_ok A Aplus Amult Aone Azero Aeq vm T).
+Reflexivity.
+Save.
+
+End rings.
+
+V8Infix "+" Pplus : ring_scope.
+V8Infix "*" Pmult : ring_scope.
+V8Notation "- x" := (Popp x) : ring_scope.
+V8Notation "[ x ]" := (Pvar x) (at level 1) : ring_scope.
+
+Delimits Scope ring_scope with ring.
diff --git a/contrib7/ring/Ring_theory.v b/contrib7/ring/Ring_theory.v
new file mode 100644
index 00000000..85fb7f6c
--- /dev/null
+++ b/contrib7/ring/Ring_theory.v
@@ -0,0 +1,384 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Ring_theory.v,v 1.1.2.1 2004/07/16 19:30:19 herbelin Exp $ *)
+
+Require Export Bool.
+
+Set Implicit Arguments.
+
+Section Theory_of_semi_rings.
+
+Variable A : Type.
+Variable Aplus : A -> A -> A.
+Variable Amult : A -> A -> A.
+Variable Aone : A.
+Variable Azero : A.
+(* There is also a "weakly decidable" equality on A. That means
+ that if (A_eq x y)=true then x=y but x=y can arise when
+ (A_eq x y)=false. On an abstract ring the function [x,y:A]false
+ is a good choice. The proof of A_eq_prop is in this case easy. *)
+Variable Aeq : A -> A -> bool.
+
+Infix 4 "+" Aplus V8only 50 (left associativity).
+Infix 4 "*" Amult V8only 40 (left associativity).
+Notation "0" := Azero.
+Notation "1" := Aone.
+
+Record Semi_Ring_Theory : Prop :=
+{ SR_plus_sym : (n,m:A) n + m == m + n;
+ SR_plus_assoc : (n,m,p:A) n + (m + p) == (n + m) + p;
+ SR_mult_sym : (n,m:A) n*m == m*n;
+ SR_mult_assoc : (n,m,p:A) n*(m*p) == (n*m)*p;
+ SR_plus_zero_left :(n:A) 0 + n == n;
+ SR_mult_one_left : (n:A) 1*n == n;
+ SR_mult_zero_left : (n:A) 0*n == 0;
+ SR_distr_left : (n,m,p:A) (n + m)*p == n*p + m*p;
+ SR_plus_reg_left : (n,m,p:A) n + m == n + p -> m==p;
+ SR_eq_prop : (x,y:A) (Is_true (Aeq x y)) -> x==y
+}.
+
+Variable T : Semi_Ring_Theory.
+
+Local plus_sym := (SR_plus_sym T).
+Local plus_assoc := (SR_plus_assoc T).
+Local mult_sym := ( SR_mult_sym T).
+Local mult_assoc := (SR_mult_assoc T).
+Local plus_zero_left := (SR_plus_zero_left T).
+Local mult_one_left := (SR_mult_one_left T).
+Local mult_zero_left := (SR_mult_zero_left T).
+Local distr_left := (SR_distr_left T).
+Local plus_reg_left := (SR_plus_reg_left T).
+
+Hints Resolve plus_sym plus_assoc mult_sym mult_assoc
+ plus_zero_left mult_one_left mult_zero_left distr_left
+ plus_reg_left.
+
+(* Lemmas whose form is x=y are also provided in form y=x because Auto does
+ not symmetry *)
+Lemma SR_mult_assoc2 : (n,m,p:A) (n * m) * p == n * (m * p).
+Symmetry; EAuto. Qed.
+
+Lemma SR_plus_assoc2 : (n,m,p:A) (n + m) + p == n + (m + p).
+Symmetry; EAuto. Qed.
+
+Lemma SR_plus_zero_left2 : (n:A) n == 0 + n.
+Symmetry; EAuto. Qed.
+
+Lemma SR_mult_one_left2 : (n:A) n == 1*n.
+Symmetry; EAuto. Qed.
+
+Lemma SR_mult_zero_left2 : (n:A) 0 == 0*n.
+Symmetry; EAuto. Qed.
+
+Lemma SR_distr_left2 : (n,m,p:A) n*p + m*p == (n + m)*p.
+Symmetry; EAuto. Qed.
+
+Lemma SR_plus_permute : (n,m,p:A) n + (m + p) == m + (n + p).
+Intros.
+Rewrite -> plus_assoc.
+Elim (plus_sym m n).
+Rewrite <- plus_assoc.
+Reflexivity.
+Qed.
+
+Lemma SR_mult_permute : (n,m,p:A) n*(m*p) == m*(n*p).
+Intros.
+Rewrite -> mult_assoc.
+Elim (mult_sym m n).
+Rewrite <- mult_assoc.
+Reflexivity.
+Qed.
+
+Hints Resolve SR_plus_permute SR_mult_permute.
+
+Lemma SR_distr_right : (n,m,p:A) n*(m + p) == (n*m) + (n*p).
+Intros.
+Repeat Rewrite -> (mult_sym n).
+EAuto.
+Qed.
+
+Lemma SR_distr_right2 : (n,m,p:A) (n*m) + (n*p) == n*(m + p).
+Symmetry; Apply SR_distr_right. Qed.
+
+Lemma SR_mult_zero_right : (n:A) n*0 == 0.
+Intro; Rewrite mult_sym; EAuto.
+Qed.
+
+Lemma SR_mult_zero_right2 : (n:A) 0 == n*0.
+Intro; Rewrite mult_sym; EAuto.
+Qed.
+
+Lemma SR_plus_zero_right :(n:A) n + 0 == n.
+Intro; Rewrite plus_sym; EAuto.
+Qed.
+Lemma SR_plus_zero_right2 :(n:A) n == n + 0.
+Intro; Rewrite plus_sym; EAuto.
+Qed.
+
+Lemma SR_mult_one_right : (n:A) n*1 == n.
+Intro; Elim mult_sym; Auto.
+Qed.
+
+Lemma SR_mult_one_right2 : (n:A) n == n*1.
+Intro; Elim mult_sym; Auto.
+Qed.
+
+Lemma SR_plus_reg_right : (n,m,p:A) m + n == p + n -> m==p.
+Intros n m p; Rewrite (plus_sym m n); Rewrite (plus_sym p n); EAuto.
+Qed.
+
+End Theory_of_semi_rings.
+
+Section Theory_of_rings.
+
+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.
+
+Infix 4 "+" Aplus V8only 50 (left associativity).
+Infix 4 "*" Amult V8only 40 (left associativity).
+Notation "0" := Azero.
+Notation "1" := Aone.
+Notation "- x" := (Aopp x) (at level 0) V8only.
+
+Record Ring_Theory : Prop :=
+{ Th_plus_sym : (n,m:A) n + m == m + n;
+ Th_plus_assoc : (n,m,p:A) n + (m + p) == (n + m) + p;
+ Th_mult_sym : (n,m:A) n*m == m*n;
+ Th_mult_assoc : (n,m,p:A) n*(m*p) == (n*m)*p;
+ Th_plus_zero_left :(n:A) 0 + n == n;
+ Th_mult_one_left : (n:A) 1*n == n;
+ Th_opp_def : (n:A) n + (-n) == 0;
+ Th_distr_left : (n,m,p:A) (n + m)*p == n*p + m*p;
+ Th_eq_prop : (x,y:A) (Is_true (Aeq x y)) -> x==y
+}.
+
+Variable T : Ring_Theory.
+
+Local plus_sym := (Th_plus_sym T).
+Local plus_assoc := (Th_plus_assoc T).
+Local mult_sym := ( Th_mult_sym T).
+Local mult_assoc := (Th_mult_assoc T).
+Local plus_zero_left := (Th_plus_zero_left T).
+Local mult_one_left := (Th_mult_one_left T).
+Local opp_def := (Th_opp_def T).
+Local distr_left := (Th_distr_left T).
+
+Hints Resolve plus_sym plus_assoc mult_sym mult_assoc
+ plus_zero_left mult_one_left opp_def distr_left.
+
+(* Lemmas whose form is x=y are also provided in form y=x because Auto does
+ not symmetry *)
+Lemma Th_mult_assoc2 : (n,m,p:A) (n * m) * p == n * (m * p).
+Symmetry; EAuto. Qed.
+
+Lemma Th_plus_assoc2 : (n,m,p:A) (n + m) + p == n + (m + p).
+Symmetry; EAuto. Qed.
+
+Lemma Th_plus_zero_left2 : (n:A) n == 0 + n.
+Symmetry; EAuto. Qed.
+
+Lemma Th_mult_one_left2 : (n:A) n == 1*n.
+Symmetry; EAuto. Qed.
+
+Lemma Th_distr_left2 : (n,m,p:A) n*p + m*p == (n + m)*p.
+Symmetry; EAuto. Qed.
+
+Lemma Th_opp_def2 : (n:A) 0 == n + (-n).
+Symmetry; EAuto. Qed.
+
+Lemma Th_plus_permute : (n,m,p:A) n + (m + p) == m + (n + p).
+Intros.
+Rewrite -> plus_assoc.
+Elim (plus_sym m n).
+Rewrite <- plus_assoc.
+Reflexivity.
+Qed.
+
+Lemma Th_mult_permute : (n,m,p:A) n*(m*p) == m*(n*p).
+Intros.
+Rewrite -> mult_assoc.
+Elim (mult_sym m n).
+Rewrite <- mult_assoc.
+Reflexivity.
+Qed.
+
+Hints Resolve Th_plus_permute Th_mult_permute.
+
+Lemma aux1 : (a:A) a + a == a -> a == 0.
+Intros.
+Generalize (opp_def a).
+Pattern 1 a.
+Rewrite <- H.
+Rewrite <- plus_assoc.
+Rewrite -> opp_def.
+Elim plus_sym.
+Rewrite plus_zero_left.
+Trivial.
+Qed.
+
+Lemma Th_mult_zero_left :(n:A) 0*n == 0.
+Intros.
+Apply aux1.
+Rewrite <- distr_left.
+Rewrite plus_zero_left.
+Reflexivity.
+Qed.
+Hints Resolve Th_mult_zero_left.
+
+Lemma Th_mult_zero_left2 : (n:A) 0 == 0*n.
+Symmetry; EAuto. Qed.
+
+Lemma aux2 : (x,y,z:A) x+y==0 -> x+z==0 -> y==z.
+Intros.
+Rewrite <- (plus_zero_left y).
+Elim H0.
+Elim plus_assoc.
+Elim (plus_sym y z).
+Rewrite -> plus_assoc.
+Rewrite -> H.
+Rewrite plus_zero_left.
+Reflexivity.
+Qed.
+
+Lemma Th_opp_mult_left : (x,y:A) -(x*y) == (-x)*y.
+Intros.
+Apply (aux2 1!x*y);
+[ Apply opp_def
+| Rewrite <- distr_left;
+ Rewrite -> opp_def;
+ Auto].
+Qed.
+Hints Resolve Th_opp_mult_left.
+
+Lemma Th_opp_mult_left2 : (x,y:A) (-x)*y == -(x*y).
+Symmetry; EAuto. Qed.
+
+Lemma Th_mult_zero_right : (n:A) n*0 == 0.
+Intro; Elim mult_sym; EAuto.
+Qed.
+
+Lemma Th_mult_zero_right2 : (n:A) 0 == n*0.
+Intro; Elim mult_sym; EAuto.
+Qed.
+
+Lemma Th_plus_zero_right :(n:A) n + 0 == n.
+Intro; Rewrite plus_sym; EAuto.
+Qed.
+
+Lemma Th_plus_zero_right2 :(n:A) n == n + 0.
+Intro; Rewrite plus_sym; EAuto.
+Qed.
+
+Lemma Th_mult_one_right : (n:A) n*1 == n.
+Intro;Elim mult_sym; EAuto.
+Qed.
+
+Lemma Th_mult_one_right2 : (n:A) n == n*1.
+Intro;Elim mult_sym; EAuto.
+Qed.
+
+Lemma Th_opp_mult_right : (x,y:A) -(x*y) == x*(-y).
+Intros; Do 2 Rewrite -> (mult_sym x); Auto.
+Qed.
+
+Lemma Th_opp_mult_right2 : (x,y:A) x*(-y) == -(x*y).
+Intros; Do 2 Rewrite -> (mult_sym x); Auto.
+Qed.
+
+Lemma Th_plus_opp_opp : (x,y:A) (-x) + (-y) == -(x+y).
+Intros.
+Apply (aux2 1! x + y);
+[ Elim plus_assoc;
+ Rewrite -> (Th_plus_permute y (-x)); Rewrite -> plus_assoc;
+ Rewrite -> opp_def; Rewrite plus_zero_left; Auto
+| Auto ].
+Qed.
+
+Lemma Th_plus_permute_opp: (n,m,p:A) (-m)+(n+p) == n+((-m)+p).
+EAuto. Qed.
+
+Lemma Th_opp_opp : (n:A) -(-n) == n.
+Intro; Apply (aux2 1! -n);
+ [ Auto | Elim plus_sym; Auto ].
+Qed.
+Hints Resolve Th_opp_opp.
+
+Lemma Th_opp_opp2 : (n:A) n == -(-n).
+Symmetry; EAuto. Qed.
+
+Lemma Th_mult_opp_opp : (x,y:A) (-x)*(-y) == x*y.
+Intros; Rewrite <- Th_opp_mult_left; Rewrite <- Th_opp_mult_right; Auto.
+Qed.
+
+Lemma Th_mult_opp_opp2 : (x,y:A) x*y == (-x)*(-y).
+Symmetry; Apply Th_mult_opp_opp. Qed.
+
+Lemma Th_opp_zero : -0 == 0.
+Rewrite <- (plus_zero_left (-0)).
+Auto. Qed.
+
+Lemma Th_plus_reg_left : (n,m,p:A) n + m == n + p -> m==p.
+Intros; Generalize (congr_eqT ? ? [z] (-n)+z ? ? H).
+Repeat Rewrite plus_assoc.
+Rewrite (plus_sym (-n) n).
+Rewrite opp_def.
+Repeat Rewrite Th_plus_zero_left; EAuto.
+Qed.
+
+Lemma Th_plus_reg_right : (n,m,p:A) m + n == p + n -> m==p.
+Intros.
+EApply Th_plus_reg_left with n.
+Rewrite (plus_sym n m).
+Rewrite (plus_sym n p).
+Auto.
+Qed.
+
+Lemma Th_distr_right : (n,m,p:A) n*(m + p) == (n*m) + (n*p).
+Intros.
+Repeat Rewrite -> (mult_sym n).
+EAuto.
+Qed.
+
+Lemma Th_distr_right2 : (n,m,p:A) (n*m) + (n*p) == n*(m + p).
+Symmetry; Apply Th_distr_right.
+Qed.
+
+End Theory_of_rings.
+
+Hints Resolve Th_mult_zero_left Th_plus_reg_left : core.
+
+Unset Implicit Arguments.
+
+Definition Semi_Ring_Theory_of :
+ (A:Type)(Aplus : A -> A -> A)(Amult : A -> A -> A)(Aone : A)
+ (Azero : A)(Aopp : A -> A)(Aeq : A -> A -> bool)
+ (Ring_Theory Aplus Amult Aone Azero Aopp Aeq)
+ ->(Semi_Ring_Theory Aplus Amult Aone Azero Aeq).
+Intros until 1; Case H.
+Split; Intros; Simpl; EAuto.
+Defined.
+
+(* Every ring can be viewed as a semi-ring : this property will be used
+ in Abstract_polynom. *)
+Coercion Semi_Ring_Theory_of : Ring_Theory >-> Semi_Ring_Theory.
+
+
+Section product_ring.
+
+End product_ring.
+
+Section power_ring.
+
+End power_ring.
diff --git a/contrib7/ring/Setoid_ring.v b/contrib7/ring/Setoid_ring.v
new file mode 100644
index 00000000..222104e5
--- /dev/null
+++ b/contrib7/ring/Setoid_ring.v
@@ -0,0 +1,13 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Setoid_ring.v,v 1.1.2.1 2004/07/16 19:30:19 herbelin Exp $ *)
+
+Require Export Setoid_ring_theory.
+Require Export Quote.
+Require Export Setoid_ring_normalize.
diff --git a/contrib7/ring/Setoid_ring_normalize.v b/contrib7/ring/Setoid_ring_normalize.v
new file mode 100644
index 00000000..b6b79dae
--- /dev/null
+++ b/contrib7/ring/Setoid_ring_normalize.v
@@ -0,0 +1,1141 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Setoid_ring_normalize.v,v 1.1.2.1 2004/07/16 19:30:19 herbelin Exp $ *)
+
+Require Setoid_ring_theory.
+Require Quote.
+
+Set Implicit Arguments.
+
+Lemma index_eq_prop: (n,m:index)(Is_true (index_eq n m)) -> n=m.
+Proof.
+ Induction n; Induction m; Simpl; Try (Reflexivity Orelse Contradiction).
+ Intros; Rewrite (H i0); Trivial.
+ Intros; Rewrite (H i0); Trivial.
+Save.
+
+Section setoid.
+
+Variable A : Type.
+Variable Aequiv : A -> A -> Prop.
+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 S : (Setoid_Theory A Aequiv).
+
+Add Setoid A Aequiv S.
+
+Variable plus_morph : (a,a0,a1,a2:A)
+ (Aequiv a a0)->(Aequiv a1 a2)->(Aequiv (Aplus a a1) (Aplus a0 a2)).
+Variable mult_morph : (a,a0,a1,a2:A)
+ (Aequiv a a0)->(Aequiv a1 a2)->(Aequiv (Amult a a1) (Amult a0 a2)).
+Variable opp_morph : (a,a0:A)
+ (Aequiv a a0)->(Aequiv (Aopp a) (Aopp a0)).
+
+Add Morphism Aplus : Aplus_ext.
+Exact plus_morph.
+Save.
+
+Add Morphism Amult : Amult_ext.
+Exact mult_morph.
+Save.
+
+Add Morphism Aopp : Aopp_ext.
+Exact opp_morph.
+Save.
+
+Local equiv_refl := (Seq_refl A Aequiv S).
+Local equiv_sym := (Seq_sym A Aequiv S).
+Local equiv_trans := (Seq_trans A Aequiv S).
+
+Hints Resolve equiv_refl equiv_trans.
+Hints Immediate equiv_sym.
+
+Section semi_setoid_rings.
+
+(* Section definitions. *)
+
+
+(******************************************)
+(* Normal abtract Polynomials *)
+(******************************************)
+(* DEFINITIONS :
+- A varlist is a sorted product of one or more variables : x, x*y*z
+- A monom is a constant, a varlist or the product of a constant by a varlist
+ variables. 2*x*y, x*y*z, 3 are monoms : 2*3, x*3*y, 4*x*3 are NOT.
+- A canonical sum is either a monom or an ordered sum of monoms
+ (the order on monoms is defined later)
+- A normal polynomial it either a constant or a canonical sum or a constant
+ plus a canonical sum
+*)
+
+(* varlist is isomorphic to (list var), but we built a special inductive
+ for efficiency *)
+Inductive varlist : Type :=
+| Nil_var : varlist
+| Cons_var : index -> varlist -> varlist
+.
+
+Inductive canonical_sum : Type :=
+| Nil_monom : canonical_sum
+| Cons_monom : A -> varlist -> canonical_sum -> canonical_sum
+| Cons_varlist : varlist -> canonical_sum -> canonical_sum
+.
+
+(* Order on monoms *)
+
+(* That's the lexicographic order on varlist, extended by :
+ - A constant is less than every monom
+ - The relation between two varlist is preserved by multiplication by a
+ constant.
+
+ Examples :
+ 3 < x < y
+ x*y < x*y*y*z
+ 2*x*y < x*y*y*z
+ x*y < 54*x*y*y*z
+ 4*x*y < 59*x*y*y*z
+*)
+
+Fixpoint varlist_eq [x,y:varlist] : bool :=
+ Cases x y of
+ | Nil_var Nil_var => true
+ | (Cons_var i xrest) (Cons_var j yrest) =>
+ (andb (index_eq i j) (varlist_eq xrest yrest))
+ | _ _ => false
+ end.
+
+Fixpoint varlist_lt [x,y:varlist] : bool :=
+ Cases x y of
+ | Nil_var (Cons_var _ _) => true
+ | (Cons_var i xrest) (Cons_var j yrest) =>
+ if (index_lt i j) then true
+ else (andb (index_eq i j) (varlist_lt xrest yrest))
+ | _ _ => false
+ end.
+
+(* merges two variables lists *)
+Fixpoint varlist_merge [l1:varlist] : varlist -> varlist :=
+ Cases l1 of
+ | (Cons_var v1 t1) =>
+ Fix vm_aux {vm_aux [l2:varlist] : varlist :=
+ Cases l2 of
+ | (Cons_var v2 t2) =>
+ if (index_lt v1 v2)
+ then (Cons_var v1 (varlist_merge t1 l2))
+ else (Cons_var v2 (vm_aux t2))
+ | Nil_var => l1
+ end}
+ | Nil_var => [l2]l2
+ end.
+
+(* returns the sum of two canonical sums *)
+Fixpoint canonical_sum_merge [s1:canonical_sum]
+ : canonical_sum -> canonical_sum :=
+Cases s1 of
+| (Cons_monom c1 l1 t1) =>
+ Fix csm_aux{csm_aux[s2:canonical_sum] : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 c2) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 (canonical_sum_merge t1 s2))
+ else (Cons_monom c2 l2 (csm_aux t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 Aone) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 (canonical_sum_merge t1 s2))
+ else (Cons_varlist l2 (csm_aux t2))
+ | Nil_monom => s1
+ end}
+| (Cons_varlist l1 t1) =>
+ Fix csm_aux2{csm_aux2[s2:canonical_sum] : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone c2) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 (canonical_sum_merge t1 s2))
+ else (Cons_monom c2 l2 (csm_aux2 t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone Aone) l1
+ (canonical_sum_merge t1 t2))
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 (canonical_sum_merge t1 s2))
+ else (Cons_varlist l2 (csm_aux2 t2))
+ | Nil_monom => s1
+ end}
+| Nil_monom => [s2]s2
+end.
+
+(* Insertion of a monom in a canonical sum *)
+Fixpoint monom_insert [c1:A; l1:varlist; s2 : canonical_sum]
+ : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 c2) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 s2)
+ else (Cons_monom c2 l2 (monom_insert c1 l1 t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus c1 Aone) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_monom c1 l1 s2)
+ else (Cons_varlist l2 (monom_insert c1 l1 t2))
+ | Nil_monom => (Cons_monom c1 l1 Nil_monom)
+ end.
+
+Fixpoint varlist_insert [l1:varlist; s2:canonical_sum]
+ : canonical_sum :=
+ Cases s2 of
+ | (Cons_monom c2 l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone c2) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 s2)
+ else (Cons_monom c2 l2 (varlist_insert l1 t2))
+ | (Cons_varlist l2 t2) =>
+ if (varlist_eq l1 l2)
+ then (Cons_monom (Aplus Aone Aone) l1 t2)
+ else if (varlist_lt l1 l2)
+ then (Cons_varlist l1 s2)
+ else (Cons_varlist l2 (varlist_insert l1 t2))
+ | Nil_monom => (Cons_varlist l1 Nil_monom)
+ end.
+
+(* Computes c0*s *)
+Fixpoint canonical_sum_scalar [c0:A; s:canonical_sum] : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ (Cons_monom (Amult c0 c) l (canonical_sum_scalar c0 t))
+ | (Cons_varlist l t) =>
+ (Cons_monom c0 l (canonical_sum_scalar c0 t))
+ | Nil_monom => Nil_monom
+ end.
+
+(* Computes l0*s *)
+Fixpoint canonical_sum_scalar2 [l0:varlist; s:canonical_sum]
+ : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ (monom_insert c (varlist_merge l0 l) (canonical_sum_scalar2 l0 t))
+ | (Cons_varlist l t) =>
+ (varlist_insert (varlist_merge l0 l) (canonical_sum_scalar2 l0 t))
+ | Nil_monom => Nil_monom
+ end.
+
+(* Computes c0*l0*s *)
+Fixpoint canonical_sum_scalar3 [c0:A;l0:varlist; s:canonical_sum]
+ : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ (monom_insert (Amult c0 c) (varlist_merge l0 l)
+ (canonical_sum_scalar3 c0 l0 t))
+ | (Cons_varlist l t) =>
+ (monom_insert c0 (varlist_merge l0 l)
+ (canonical_sum_scalar3 c0 l0 t))
+ | Nil_monom => Nil_monom
+ end.
+
+(* returns the product of two canonical sums *)
+Fixpoint canonical_sum_prod [s1:canonical_sum]
+ : canonical_sum -> canonical_sum :=
+ [s2]Cases s1 of
+ | (Cons_monom c1 l1 t1) =>
+ (canonical_sum_merge (canonical_sum_scalar3 c1 l1 s2)
+ (canonical_sum_prod t1 s2))
+ | (Cons_varlist l1 t1) =>
+ (canonical_sum_merge (canonical_sum_scalar2 l1 s2)
+ (canonical_sum_prod t1 s2))
+ | Nil_monom => Nil_monom
+ end.
+
+(* The type to represent concrete semi-setoid-ring polynomials *)
+
+Inductive Type setspolynomial :=
+ SetSPvar : index -> setspolynomial
+| SetSPconst : A -> setspolynomial
+| SetSPplus : setspolynomial -> setspolynomial -> setspolynomial
+| SetSPmult : setspolynomial -> setspolynomial -> setspolynomial.
+
+Fixpoint setspolynomial_normalize [p:setspolynomial] : canonical_sum :=
+ Cases p of
+ | (SetSPplus l r) => (canonical_sum_merge (setspolynomial_normalize l) (setspolynomial_normalize r))
+ | (SetSPmult l r) => (canonical_sum_prod (setspolynomial_normalize l) (setspolynomial_normalize r))
+ | (SetSPconst c) => (Cons_monom c Nil_var Nil_monom)
+ | (SetSPvar i) => (Cons_varlist (Cons_var i Nil_var) Nil_monom)
+ end.
+
+Fixpoint canonical_sum_simplify [ s:canonical_sum] : canonical_sum :=
+ Cases s of
+ | (Cons_monom c l t) =>
+ if (Aeq c Azero)
+ then (canonical_sum_simplify t)
+ else if (Aeq c Aone)
+ then (Cons_varlist l (canonical_sum_simplify t))
+ else (Cons_monom c l (canonical_sum_simplify t))
+ | (Cons_varlist l t) => (Cons_varlist l (canonical_sum_simplify t))
+ | Nil_monom => Nil_monom
+ end.
+
+Definition setspolynomial_simplify :=
+ [x:setspolynomial] (canonical_sum_simplify (setspolynomial_normalize x)).
+
+Variable vm : (varmap A).
+
+Definition interp_var [i:index] := (varmap_find Azero i vm).
+
+Definition ivl_aux := Fix ivl_aux {ivl_aux[x:index; t:varlist] : A :=
+ Cases t of
+ | Nil_var => (interp_var x)
+ | (Cons_var x' t') => (Amult (interp_var x) (ivl_aux x' t'))
+ end}.
+
+Definition interp_vl := [l:varlist]
+ Cases l of
+ | Nil_var => Aone
+ | (Cons_var x t) => (ivl_aux x t)
+ end.
+
+Definition interp_m := [c:A][l:varlist]
+ Cases l of
+ | Nil_var => c
+ | (Cons_var x t) =>
+ (Amult c (ivl_aux x t))
+ end.
+
+Definition ics_aux := Fix ics_aux{ics_aux[a:A; s:canonical_sum] : A :=
+ Cases s of
+ | Nil_monom => a
+ | (Cons_varlist l t) => (Aplus a (ics_aux (interp_vl l) t))
+ | (Cons_monom c l t) => (Aplus a (ics_aux (interp_m c l) t))
+ end}.
+
+Definition interp_setcs : canonical_sum -> A :=
+ [s]Cases s of
+ | Nil_monom => Azero
+ | (Cons_varlist l t) =>
+ (ics_aux (interp_vl l) t)
+ | (Cons_monom c l t) =>
+ (ics_aux (interp_m c l) t)
+ end.
+
+Fixpoint interp_setsp [p:setspolynomial] : A :=
+ Cases p of
+ | (SetSPconst c) => c
+ | (SetSPvar i) => (interp_var i)
+ | (SetSPplus p1 p2) => (Aplus (interp_setsp p1) (interp_setsp p2))
+ | (SetSPmult p1 p2) => (Amult (interp_setsp p1) (interp_setsp p2))
+ end.
+
+(* End interpretation. *)
+
+Unset Implicit Arguments.
+
+(* Section properties. *)
+
+Variable T : (Semi_Setoid_Ring_Theory Aequiv Aplus Amult Aone Azero Aeq).
+
+Hint SSR_plus_sym_T := Resolve (SSR_plus_sym T).
+Hint SSR_plus_assoc_T := Resolve (SSR_plus_assoc T).
+Hint SSR_plus_assoc2_T := Resolve (SSR_plus_assoc2 S T).
+Hint SSR_mult_sym_T := Resolve (SSR_mult_sym T).
+Hint SSR_mult_assoc_T := Resolve (SSR_mult_assoc T).
+Hint SSR_mult_assoc2_T := Resolve (SSR_mult_assoc2 S T).
+Hint SSR_plus_zero_left_T := Resolve (SSR_plus_zero_left T).
+Hint SSR_plus_zero_left2_T := Resolve (SSR_plus_zero_left2 S T).
+Hint SSR_mult_one_left_T := Resolve (SSR_mult_one_left T).
+Hint SSR_mult_one_left2_T := Resolve (SSR_mult_one_left2 S T).
+Hint SSR_mult_zero_left_T := Resolve (SSR_mult_zero_left T).
+Hint SSR_mult_zero_left2_T := Resolve (SSR_mult_zero_left2 S T).
+Hint SSR_distr_left_T := Resolve (SSR_distr_left T).
+Hint SSR_distr_left2_T := Resolve (SSR_distr_left2 S T).
+Hint SSR_plus_reg_left_T := Resolve (SSR_plus_reg_left T).
+Hint SSR_plus_permute_T := Resolve (SSR_plus_permute S plus_morph T).
+Hint SSR_mult_permute_T := Resolve (SSR_mult_permute S mult_morph T).
+Hint SSR_distr_right_T := Resolve (SSR_distr_right S plus_morph T).
+Hint SSR_distr_right2_T := Resolve (SSR_distr_right2 S plus_morph T).
+Hint SSR_mult_zero_right_T := Resolve (SSR_mult_zero_right S T).
+Hint SSR_mult_zero_right2_T := Resolve (SSR_mult_zero_right2 S T).
+Hint SSR_plus_zero_right_T := Resolve (SSR_plus_zero_right S T).
+Hint SSR_plus_zero_right2_T := Resolve (SSR_plus_zero_right2 S T).
+Hint SSR_mult_one_right_T := Resolve (SSR_mult_one_right S T).
+Hint SSR_mult_one_right2_T := Resolve (SSR_mult_one_right2 S T).
+Hint SSR_plus_reg_right_T := Resolve (SSR_plus_reg_right S T).
+Hints Resolve refl_equal sym_equal trans_equal.
+(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hints Immediate T.
+
+Lemma varlist_eq_prop : (x,y:varlist)
+ (Is_true (varlist_eq x y))->x==y.
+Proof.
+ Induction x; Induction y; Contradiction Orelse Try Reflexivity.
+ Simpl; Intros.
+ Generalize (andb_prop2 ? ? H1); Intros; Elim H2; Intros.
+ Rewrite (index_eq_prop H3); Rewrite (H v0 H4); Reflexivity.
+Save.
+
+Remark ivl_aux_ok : (v:varlist)(i:index)
+ (Aequiv (ivl_aux i v) (Amult (interp_var i) (interp_vl v))).
+Proof.
+ Induction v; Simpl; Intros.
+ Trivial.
+ Rewrite (H i); Trivial.
+Save.
+
+Lemma varlist_merge_ok : (x,y:varlist)
+ (Aequiv (interp_vl (varlist_merge x y)) (Amult (interp_vl x) (interp_vl y))).
+Proof.
+ Induction x.
+ Simpl; Trivial.
+ Induction y.
+ Simpl; Trivial.
+ Simpl; Intros.
+ Elim (index_lt i i0); Simpl; Intros.
+
+ Rewrite (ivl_aux_ok v i).
+ Rewrite (ivl_aux_ok v0 i0).
+ Rewrite (ivl_aux_ok (varlist_merge v (Cons_var i0 v0)) i).
+ Rewrite (H (Cons_var i0 v0)).
+ Simpl.
+ Rewrite (ivl_aux_ok v0 i0).
+ EAuto.
+
+ Rewrite (ivl_aux_ok v i).
+ Rewrite (ivl_aux_ok v0 i0).
+ Rewrite (ivl_aux_ok
+ (Fix vm_aux
+ {vm_aux [l2:varlist] : varlist :=
+ Cases (l2) of
+ Nil_var => (Cons_var i v)
+ | (Cons_var v2 t2) =>
+ (if (index_lt i v2)
+ then (Cons_var i (varlist_merge v l2))
+ else (Cons_var v2 (vm_aux t2)))
+ end} v0) i0).
+ Rewrite H0.
+ Rewrite (ivl_aux_ok v i).
+ EAuto.
+Save.
+
+Remark ics_aux_ok : (x:A)(s:canonical_sum)
+ (Aequiv (ics_aux x s) (Aplus x (interp_setcs s))).
+Proof.
+ Induction s; Simpl; Intros;Trivial.
+Save.
+
+Remark interp_m_ok : (x:A)(l:varlist)
+ (Aequiv (interp_m x l) (Amult x (interp_vl l))).
+Proof.
+ NewDestruct l;Trivial.
+Save.
+
+Hint ivl_aux_ok_ := Resolve ivl_aux_ok.
+Hint ics_aux_ok_ := Resolve ics_aux_ok.
+Hint interp_m_ok_ := Resolve interp_m_ok.
+
+(* Hints Resolve ivl_aux_ok ics_aux_ok interp_m_ok. *)
+
+Lemma canonical_sum_merge_ok : (x,y:canonical_sum)
+ (Aequiv (interp_setcs (canonical_sum_merge x y))
+ (Aplus (interp_setcs x) (interp_setcs y))).
+Proof.
+Induction x; Simpl.
+Trivial.
+
+Induction y; Simpl; Intros.
+EAuto.
+
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl.
+Rewrite (ics_aux_ok (interp_m a v0) c).
+Rewrite (ics_aux_ok (interp_m a0 v0) c0).
+Rewrite (ics_aux_ok (interp_m (Aplus a a0) v0)
+ (canonical_sum_merge c c0)).
+Rewrite (H c0).
+Rewrite (interp_m_ok (Aplus a a0) v0).
+Rewrite (interp_m_ok a v0).
+Rewrite (interp_m_ok a0 v0).
+Setoid_replace (Amult (Aplus a a0) (interp_vl v0))
+ with (Aplus (Amult a (interp_vl v0)) (Amult a0 (interp_vl v0))).
+Setoid_replace (Aplus
+ (Aplus (Amult a (interp_vl v0))
+ (Amult a0 (interp_vl v0)))
+ (Aplus (interp_setcs c) (interp_setcs c0)))
+ with (Aplus (Amult a (interp_vl v0))
+ (Aplus (Amult a0 (interp_vl v0))
+ (Aplus (interp_setcs c) (interp_setcs c0)))).
+Setoid_replace (Aplus (Aplus (Amult a (interp_vl v0)) (interp_setcs c))
+ (Aplus (Amult a0 (interp_vl v0)) (interp_setcs c0)))
+ with (Aplus (Amult a (interp_vl v0))
+ (Aplus (interp_setcs c)
+ (Aplus (Amult a0 (interp_vl v0)) (interp_setcs c0)))).
+Auto.
+
+Elim (varlist_lt v v0); Simpl.
+Intro.
+Rewrite (ics_aux_ok (interp_m a v)
+ (canonical_sum_merge c (Cons_monom a0 v0 c0))).
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (ics_aux_ok (interp_m a0 v0) c0).
+Rewrite (H (Cons_monom a0 v0 c0)); Simpl.
+Rewrite (ics_aux_ok (interp_m a0 v0) c0); Auto.
+
+Intro.
+Rewrite (ics_aux_ok (interp_m a0 v0)
+ (Fix csm_aux
+ {csm_aux [s2:canonical_sum] : canonical_sum :=
+ Cases (s2) of
+ Nil_monom => (Cons_monom a v c)
+ | (Cons_monom c2 l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus a c2) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_monom a v
+ (canonical_sum_merge c s2))
+ else (Cons_monom c2 l2 (csm_aux t2))))
+ | (Cons_varlist l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus a Aone) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_monom a v
+ (canonical_sum_merge c s2))
+ else (Cons_varlist l2 (csm_aux t2))))
+ end} c0)).
+Rewrite H0.
+Rewrite (ics_aux_ok (interp_m a v) c);
+Rewrite (ics_aux_ok (interp_m a0 v0) c0); Simpl; Auto.
+
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus a Aone) v0)
+ (canonical_sum_merge c c0));
+Rewrite (ics_aux_ok (interp_m a v0) c);
+Rewrite (ics_aux_ok (interp_vl v0) c0).
+Rewrite (H c0).
+Rewrite (interp_m_ok (Aplus a Aone) v0).
+Rewrite (interp_m_ok a v0).
+Setoid_replace (Amult (Aplus a Aone) (interp_vl v0))
+ with (Aplus (Amult a (interp_vl v0)) (Amult Aone (interp_vl v0))).
+Setoid_replace (Aplus
+ (Aplus (Amult a (interp_vl v0))
+ (Amult Aone (interp_vl v0)))
+ (Aplus (interp_setcs c) (interp_setcs c0)))
+ with (Aplus (Amult a (interp_vl v0))
+ (Aplus (Amult Aone (interp_vl v0))
+ (Aplus (interp_setcs c) (interp_setcs c0)))).
+Setoid_replace (Aplus (Aplus (Amult a (interp_vl v0)) (interp_setcs c))
+ (Aplus (interp_vl v0) (interp_setcs c0)))
+ with (Aplus (Amult a (interp_vl v0))
+ (Aplus (interp_setcs c) (Aplus (interp_vl v0) (interp_setcs c0)))).
+Setoid_replace (Amult Aone (interp_vl v0)) with (interp_vl v0).
+Auto.
+
+Elim (varlist_lt v v0); Simpl.
+Intro.
+Rewrite (ics_aux_ok (interp_m a v)
+ (canonical_sum_merge c (Cons_varlist v0 c0)));
+Rewrite (ics_aux_ok (interp_m a v) c);
+Rewrite (ics_aux_ok (interp_vl v0) c0).
+Rewrite (H (Cons_varlist v0 c0)); Simpl.
+Rewrite (ics_aux_ok (interp_vl v0) c0).
+Auto.
+
+Intro.
+Rewrite (ics_aux_ok (interp_vl v0)
+ (Fix csm_aux
+ {csm_aux [s2:canonical_sum] : canonical_sum :=
+ Cases (s2) of
+ Nil_monom => (Cons_monom a v c)
+ | (Cons_monom c2 l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus a c2) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_monom a v
+ (canonical_sum_merge c s2))
+ else (Cons_monom c2 l2 (csm_aux t2))))
+ | (Cons_varlist l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus a Aone) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_monom a v
+ (canonical_sum_merge c s2))
+ else (Cons_varlist l2 (csm_aux t2))))
+ end} c0)); Rewrite H0.
+Rewrite (ics_aux_ok (interp_m a v) c);
+Rewrite (ics_aux_ok (interp_vl v0) c0); Simpl.
+Auto.
+
+Induction y; Simpl; Intros.
+Trivial.
+
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0).
+Intros; Rewrite (H1 I).
+Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus Aone a) v0)
+ (canonical_sum_merge c c0));
+Rewrite (ics_aux_ok (interp_vl v0) c);
+Rewrite (ics_aux_ok (interp_m a v0) c0); Rewrite (
+H c0).
+Rewrite (interp_m_ok (Aplus Aone a) v0);
+Rewrite (interp_m_ok a v0).
+Setoid_replace (Amult (Aplus Aone a) (interp_vl v0))
+ with (Aplus (Amult Aone (interp_vl v0)) (Amult a (interp_vl v0)));
+Setoid_replace (Aplus
+ (Aplus (Amult Aone (interp_vl v0))
+ (Amult a (interp_vl v0)))
+ (Aplus (interp_setcs c) (interp_setcs c0)))
+ with (Aplus (Amult Aone (interp_vl v0))
+ (Aplus (Amult a (interp_vl v0))
+ (Aplus (interp_setcs c) (interp_setcs c0))));
+Setoid_replace (Aplus (Aplus (interp_vl v0) (interp_setcs c))
+ (Aplus (Amult a (interp_vl v0)) (interp_setcs c0)))
+ with (Aplus (interp_vl v0)
+ (Aplus (interp_setcs c)
+ (Aplus (Amult a (interp_vl v0)) (interp_setcs c0)))).
+Auto.
+
+Elim (varlist_lt v v0); Simpl; Intros.
+Rewrite (ics_aux_ok (interp_vl v)
+ (canonical_sum_merge c (Cons_monom a v0 c0)));
+Rewrite (ics_aux_ok (interp_vl v) c);
+Rewrite (ics_aux_ok (interp_m a v0) c0).
+Rewrite (H (Cons_monom a v0 c0)); Simpl.
+Rewrite (ics_aux_ok (interp_m a v0) c0); Auto.
+
+Rewrite (ics_aux_ok (interp_m a v0)
+ (Fix csm_aux2
+ {csm_aux2 [s2:canonical_sum] : canonical_sum :=
+ Cases (s2) of
+ Nil_monom => (Cons_varlist v c)
+ | (Cons_monom c2 l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus Aone c2) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_varlist v
+ (canonical_sum_merge c s2))
+ else (Cons_monom c2 l2 (csm_aux2 t2))))
+ | (Cons_varlist l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus Aone Aone) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_varlist v
+ (canonical_sum_merge c s2))
+ else (Cons_varlist l2 (csm_aux2 t2))))
+ end} c0)); Rewrite H0.
+Rewrite (ics_aux_ok (interp_vl v) c);
+Rewrite (ics_aux_ok (interp_m a v0) c0); Simpl; Auto.
+
+Generalize (varlist_eq_prop v v0).
+Elim (varlist_eq v v0); Intros.
+Rewrite (H1 I); Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus Aone Aone) v0)
+ (canonical_sum_merge c c0));
+Rewrite (ics_aux_ok (interp_vl v0) c);
+Rewrite (ics_aux_ok (interp_vl v0) c0); Rewrite (
+H c0).
+Rewrite (interp_m_ok (Aplus Aone Aone) v0).
+Setoid_replace (Amult (Aplus Aone Aone) (interp_vl v0))
+ with (Aplus (Amult Aone (interp_vl v0)) (Amult Aone (interp_vl v0)));
+Setoid_replace (Aplus
+ (Aplus (Amult Aone (interp_vl v0))
+ (Amult Aone (interp_vl v0)))
+ (Aplus (interp_setcs c) (interp_setcs c0)))
+ with (Aplus (Amult Aone (interp_vl v0))
+ (Aplus (Amult Aone (interp_vl v0))
+ (Aplus (interp_setcs c) (interp_setcs c0))));
+Setoid_replace (Aplus (Aplus (interp_vl v0) (interp_setcs c))
+ (Aplus (interp_vl v0) (interp_setcs c0)))
+ with (Aplus (interp_vl v0)
+ (Aplus (interp_setcs c) (Aplus (interp_vl v0) (interp_setcs c0)))).
+Setoid_replace (Amult Aone (interp_vl v0)) with (interp_vl v0); Auto.
+
+Elim (varlist_lt v v0); Simpl.
+Rewrite (ics_aux_ok (interp_vl v)
+ (canonical_sum_merge c (Cons_varlist v0 c0)));
+Rewrite (ics_aux_ok (interp_vl v) c);
+Rewrite (ics_aux_ok (interp_vl v0) c0);
+Rewrite (H (Cons_varlist v0 c0)); Simpl.
+Rewrite (ics_aux_ok (interp_vl v0) c0); Auto.
+
+Rewrite (ics_aux_ok (interp_vl v0)
+ (Fix csm_aux2
+ {csm_aux2 [s2:canonical_sum] : canonical_sum :=
+ Cases (s2) of
+ Nil_monom => (Cons_varlist v c)
+ | (Cons_monom c2 l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus Aone c2) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_varlist v
+ (canonical_sum_merge c s2))
+ else (Cons_monom c2 l2 (csm_aux2 t2))))
+ | (Cons_varlist l2 t2) =>
+ (if (varlist_eq v l2)
+ then
+ (Cons_monom (Aplus Aone Aone) v
+ (canonical_sum_merge c t2))
+ else
+ (if (varlist_lt v l2)
+ then
+ (Cons_varlist v
+ (canonical_sum_merge c s2))
+ else (Cons_varlist l2 (csm_aux2 t2))))
+ end} c0)); Rewrite H0.
+Rewrite (ics_aux_ok (interp_vl v) c);
+Rewrite (ics_aux_ok (interp_vl v0) c0); Simpl; Auto.
+Save.
+
+Lemma monom_insert_ok: (a:A)(l:varlist)(s:canonical_sum)
+ (Aequiv (interp_setcs (monom_insert a l s))
+ (Aplus (Amult a (interp_vl l)) (interp_setcs s))).
+Proof.
+Induction s; Intros.
+Simpl; Rewrite (interp_m_ok a l); Trivial.
+
+Simpl; Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus a a0) v) c);
+Rewrite (ics_aux_ok (interp_m a0 v) c).
+Rewrite (interp_m_ok (Aplus a a0) v);
+Rewrite (interp_m_ok a0 v).
+Setoid_replace (Amult (Aplus a a0) (interp_vl v))
+ with (Aplus (Amult a (interp_vl v)) (Amult a0 (interp_vl v))).
+Auto.
+
+Elim (varlist_lt l v); Simpl; Intros.
+Rewrite (ics_aux_ok (interp_m a0 v) c).
+Rewrite (interp_m_ok a0 v); Rewrite (interp_m_ok a l).
+Auto.
+
+Rewrite (ics_aux_ok (interp_m a0 v) (monom_insert a l c));
+Rewrite (ics_aux_ok (interp_m a0 v) c); Rewrite H.
+Auto.
+
+Simpl.
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus a Aone) v) c);
+Rewrite (ics_aux_ok (interp_vl v) c).
+Rewrite (interp_m_ok (Aplus a Aone) v).
+Setoid_replace (Amult (Aplus a Aone) (interp_vl v))
+ with (Aplus (Amult a (interp_vl v)) (Amult Aone (interp_vl v))).
+Setoid_replace (Amult Aone (interp_vl v)) with (interp_vl v).
+Auto.
+
+Elim (varlist_lt l v); Simpl; Intros; Auto.
+Rewrite (ics_aux_ok (interp_vl v) (monom_insert a l c));
+Rewrite H.
+Rewrite (ics_aux_ok (interp_vl v) c); Auto.
+Save.
+
+Lemma varlist_insert_ok :
+ (l:varlist)(s:canonical_sum)
+ (Aequiv (interp_setcs (varlist_insert l s))
+ (Aplus (interp_vl l) (interp_setcs s))).
+Proof.
+Induction s; Simpl; Intros.
+Trivial.
+
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus Aone a) v) c);
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (interp_m_ok (Aplus Aone a) v);
+Rewrite (interp_m_ok a v).
+Setoid_replace (Amult (Aplus Aone a) (interp_vl v))
+ with (Aplus (Amult Aone (interp_vl v)) (Amult a (interp_vl v))).
+Setoid_replace (Amult Aone (interp_vl v)) with (interp_vl v); Auto.
+
+Elim (varlist_lt l v); Simpl; Intros; Auto.
+Rewrite (ics_aux_ok (interp_m a v) (varlist_insert l c));
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (interp_m_ok a v).
+Rewrite H; Auto.
+
+Generalize (varlist_eq_prop l v); Elim (varlist_eq l v).
+Intro Hr; Rewrite (Hr I); Simpl.
+Rewrite (ics_aux_ok (interp_m (Aplus Aone Aone) v) c);
+Rewrite (ics_aux_ok (interp_vl v) c).
+Rewrite (interp_m_ok (Aplus Aone Aone) v).
+Setoid_replace (Amult (Aplus Aone Aone) (interp_vl v))
+ with (Aplus (Amult Aone (interp_vl v)) (Amult Aone (interp_vl v))).
+Setoid_replace (Amult Aone (interp_vl v)) with (interp_vl v); Auto.
+
+Elim (varlist_lt l v); Simpl; Intros; Auto.
+Rewrite (ics_aux_ok (interp_vl v) (varlist_insert l c)).
+Rewrite H.
+Rewrite (ics_aux_ok (interp_vl v) c); Auto.
+Save.
+
+Lemma canonical_sum_scalar_ok : (a:A)(s:canonical_sum)
+ (Aequiv (interp_setcs (canonical_sum_scalar a s)) (Amult a (interp_setcs s))).
+Proof.
+Induction s; Simpl; Intros.
+Trivial.
+
+Rewrite (ics_aux_ok (interp_m (Amult a a0) v)
+ (canonical_sum_scalar a c));
+Rewrite (ics_aux_ok (interp_m a0 v) c).
+Rewrite (interp_m_ok (Amult a a0) v);
+Rewrite (interp_m_ok a0 v).
+Rewrite H.
+Setoid_replace (Amult a (Aplus (Amult a0 (interp_vl v)) (interp_setcs c)))
+ with (Aplus (Amult a (Amult a0 (interp_vl v))) (Amult a (interp_setcs c))).
+Auto.
+
+Rewrite (ics_aux_ok (interp_m a v) (canonical_sum_scalar a c));
+Rewrite (ics_aux_ok (interp_vl v) c); Rewrite H.
+Rewrite (interp_m_ok a v).
+Auto.
+Save.
+
+Lemma canonical_sum_scalar2_ok : (l:varlist; s:canonical_sum)
+ (Aequiv (interp_setcs (canonical_sum_scalar2 l s)) (Amult (interp_vl l) (interp_setcs s))).
+Proof.
+Induction s; Simpl; Intros; Auto.
+Rewrite (monom_insert_ok a (varlist_merge l v)
+ (canonical_sum_scalar2 l c)).
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (interp_m_ok a v).
+Rewrite H.
+Rewrite (varlist_merge_ok l v).
+Setoid_replace (Amult (interp_vl l)
+ (Aplus (Amult a (interp_vl v)) (interp_setcs c)))
+ with (Aplus (Amult (interp_vl l) (Amult a (interp_vl v)))
+ (Amult (interp_vl l) (interp_setcs c))).
+Auto.
+
+Rewrite (varlist_insert_ok (varlist_merge l v)
+ (canonical_sum_scalar2 l c)).
+Rewrite (ics_aux_ok (interp_vl v) c).
+Rewrite H.
+Rewrite (varlist_merge_ok l v).
+Auto.
+Save.
+
+Lemma canonical_sum_scalar3_ok : (c:A; l:varlist; s:canonical_sum)
+ (Aequiv (interp_setcs (canonical_sum_scalar3 c l s)) (Amult c (Amult (interp_vl l) (interp_setcs s)))).
+Proof.
+Induction s; Simpl; Intros.
+Rewrite (SSR_mult_zero_right S T (interp_vl l)).
+Auto.
+
+Rewrite (monom_insert_ok (Amult c a) (varlist_merge l v)
+ (canonical_sum_scalar3 c l c0)).
+Rewrite (ics_aux_ok (interp_m a v) c0).
+Rewrite (interp_m_ok a v).
+Rewrite H.
+Rewrite (varlist_merge_ok l v).
+Setoid_replace (Amult (interp_vl l)
+ (Aplus (Amult a (interp_vl v)) (interp_setcs c0)))
+ with (Aplus (Amult (interp_vl l) (Amult a (interp_vl v)))
+ (Amult (interp_vl l) (interp_setcs c0))).
+Setoid_replace (Amult c
+ (Aplus (Amult (interp_vl l) (Amult a (interp_vl v)))
+ (Amult (interp_vl l) (interp_setcs c0))))
+ with (Aplus (Amult c (Amult (interp_vl l) (Amult a (interp_vl v))))
+ (Amult c (Amult (interp_vl l) (interp_setcs c0)))).
+Setoid_replace (Amult (Amult c a) (Amult (interp_vl l) (interp_vl v)))
+ with (Amult c (Amult a (Amult (interp_vl l) (interp_vl v)))).
+Auto.
+
+Rewrite (monom_insert_ok c (varlist_merge l v)
+ (canonical_sum_scalar3 c l c0)).
+Rewrite (ics_aux_ok (interp_vl v) c0).
+Rewrite H.
+Rewrite (varlist_merge_ok l v).
+Setoid_replace (Aplus (Amult c (Amult (interp_vl l) (interp_vl v)))
+ (Amult c (Amult (interp_vl l) (interp_setcs c0))))
+ with (Amult c
+ (Aplus (Amult (interp_vl l) (interp_vl v))
+ (Amult (interp_vl l) (interp_setcs c0)))).
+Auto.
+Save.
+
+Lemma canonical_sum_prod_ok : (x,y:canonical_sum)
+ (Aequiv (interp_setcs (canonical_sum_prod x y)) (Amult (interp_setcs x) (interp_setcs y))).
+Proof.
+Induction x; Simpl; Intros.
+Trivial.
+
+Rewrite (canonical_sum_merge_ok (canonical_sum_scalar3 a v y)
+ (canonical_sum_prod c y)).
+Rewrite (canonical_sum_scalar3_ok a v y).
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (interp_m_ok a v).
+Rewrite (H y).
+Setoid_replace (Amult a (Amult (interp_vl v) (interp_setcs y)))
+ with (Amult (Amult a (interp_vl v)) (interp_setcs y)).
+Setoid_replace (Amult (Aplus (Amult a (interp_vl v)) (interp_setcs c))
+ (interp_setcs y))
+ with (Aplus (Amult (Amult a (interp_vl v)) (interp_setcs y))
+ (Amult (interp_setcs c) (interp_setcs y))).
+Trivial.
+
+Rewrite (canonical_sum_merge_ok (canonical_sum_scalar2 v y)
+ (canonical_sum_prod c y)).
+Rewrite (canonical_sum_scalar2_ok v y).
+Rewrite (ics_aux_ok (interp_vl v) c).
+Rewrite (H y).
+Trivial.
+Save.
+
+Theorem setspolynomial_normalize_ok : (p:setspolynomial)
+ (Aequiv (interp_setcs (setspolynomial_normalize p)) (interp_setsp p)).
+Proof.
+Induction p; Simpl; Intros; Trivial.
+Rewrite (canonical_sum_merge_ok (setspolynomial_normalize s)
+ (setspolynomial_normalize s0)).
+Rewrite H; Rewrite H0; Trivial.
+
+Rewrite (canonical_sum_prod_ok (setspolynomial_normalize s)
+ (setspolynomial_normalize s0)).
+Rewrite H; Rewrite H0; Trivial.
+Save.
+
+Lemma canonical_sum_simplify_ok : (s:canonical_sum)
+ (Aequiv (interp_setcs (canonical_sum_simplify s)) (interp_setcs s)).
+Proof.
+Induction s; Simpl; Intros.
+Trivial.
+
+Generalize (SSR_eq_prop T 9!a 10!Azero).
+Elim (Aeq a Azero).
+Simpl.
+Intros.
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (interp_m_ok a v).
+Rewrite (H0 I).
+Setoid_replace (Amult Azero (interp_vl v)) with Azero.
+Rewrite H.
+Trivial.
+
+Intros; Simpl.
+Generalize (SSR_eq_prop T 9!a 10!Aone).
+Elim (Aeq a Aone).
+Intros.
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite (interp_m_ok a v).
+Rewrite (H1 I).
+Simpl.
+Rewrite (ics_aux_ok (interp_vl v) (canonical_sum_simplify c)).
+Rewrite H.
+Auto.
+
+Simpl.
+Intros.
+Rewrite (ics_aux_ok (interp_m a v) (canonical_sum_simplify c)).
+Rewrite (ics_aux_ok (interp_m a v) c).
+Rewrite H; Trivial.
+
+Rewrite (ics_aux_ok (interp_vl v) (canonical_sum_simplify c)).
+Rewrite H.
+Auto.
+Save.
+
+Theorem setspolynomial_simplify_ok : (p:setspolynomial)
+ (Aequiv (interp_setcs (setspolynomial_simplify p)) (interp_setsp p)).
+Proof.
+Intro.
+Unfold setspolynomial_simplify.
+Rewrite (canonical_sum_simplify_ok (setspolynomial_normalize p)).
+Exact (setspolynomial_normalize_ok p).
+Save.
+
+End semi_setoid_rings.
+
+Implicits Cons_varlist.
+Implicits Cons_monom.
+Implicits SetSPconst.
+Implicits SetSPplus.
+Implicits SetSPmult.
+
+
+
+Section setoid_rings.
+
+Set Implicit Arguments.
+
+Variable vm : (varmap A).
+Variable T : (Setoid_Ring_Theory Aequiv Aplus Amult Aone Azero Aopp Aeq).
+
+Hint STh_plus_sym_T := Resolve (STh_plus_sym T).
+Hint STh_plus_assoc_T := Resolve (STh_plus_assoc T).
+Hint STh_plus_assoc2_T := Resolve (STh_plus_assoc2 S T).
+Hint STh_mult_sym_T := Resolve (STh_mult_sym T).
+Hint STh_mult_assoc_T := Resolve (STh_mult_assoc T).
+Hint STh_mult_assoc2_T := Resolve (STh_mult_assoc2 S T).
+Hint STh_plus_zero_left_T := Resolve (STh_plus_zero_left T).
+Hint STh_plus_zero_left2_T := Resolve (STh_plus_zero_left2 S T).
+Hint STh_mult_one_left_T := Resolve (STh_mult_one_left T).
+Hint STh_mult_one_left2_T := Resolve (STh_mult_one_left2 S T).
+Hint STh_mult_zero_left_T := Resolve (STh_mult_zero_left S plus_morph mult_morph T).
+Hint STh_mult_zero_left2_T := Resolve (STh_mult_zero_left2 S plus_morph mult_morph T).
+Hint STh_distr_left_T := Resolve (STh_distr_left T).
+Hint STh_distr_left2_T := Resolve (STh_distr_left2 S T).
+Hint STh_plus_reg_left_T := Resolve (STh_plus_reg_left S plus_morph T).
+Hint STh_plus_permute_T := Resolve (STh_plus_permute S plus_morph T).
+Hint STh_mult_permute_T := Resolve (STh_mult_permute S mult_morph T).
+Hint STh_distr_right_T := Resolve (STh_distr_right S plus_morph T).
+Hint STh_distr_right2_T := Resolve (STh_distr_right2 S plus_morph T).
+Hint STh_mult_zero_right_T := Resolve (STh_mult_zero_right S plus_morph mult_morph T).
+Hint STh_mult_zero_right2_T := Resolve (STh_mult_zero_right2 S plus_morph mult_morph T).
+Hint STh_plus_zero_right_T := Resolve (STh_plus_zero_right S T).
+Hint STh_plus_zero_right2_T := Resolve (STh_plus_zero_right2 S T).
+Hint STh_mult_one_right_T := Resolve (STh_mult_one_right S T).
+Hint STh_mult_one_right2_T := Resolve (STh_mult_one_right2 S T).
+Hint STh_plus_reg_right_T := Resolve (STh_plus_reg_right S plus_morph T).
+Hints Resolve refl_equal sym_equal trans_equal.
+(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hints Immediate T.
+
+
+(*** Definitions *)
+
+Inductive Type setpolynomial :=
+ SetPvar : index -> setpolynomial
+| SetPconst : A -> setpolynomial
+| SetPplus : setpolynomial -> setpolynomial -> setpolynomial
+| SetPmult : setpolynomial -> setpolynomial -> setpolynomial
+| SetPopp : setpolynomial -> setpolynomial.
+
+Fixpoint setpolynomial_normalize [x:setpolynomial] : canonical_sum :=
+ Cases x of
+ | (SetPplus l r) => (canonical_sum_merge
+ (setpolynomial_normalize l)
+ (setpolynomial_normalize r))
+ | (SetPmult l r) => (canonical_sum_prod
+ (setpolynomial_normalize l)
+ (setpolynomial_normalize r))
+ | (SetPconst c) => (Cons_monom c Nil_var Nil_monom)
+ | (SetPvar i) => (Cons_varlist (Cons_var i Nil_var) Nil_monom)
+ | (SetPopp p) => (canonical_sum_scalar3
+ (Aopp Aone) Nil_var
+ (setpolynomial_normalize p))
+ end.
+
+Definition setpolynomial_simplify :=
+ [x:setpolynomial](canonical_sum_simplify (setpolynomial_normalize x)).
+
+Fixpoint setspolynomial_of [x:setpolynomial] : setspolynomial :=
+ Cases x of
+ | (SetPplus l r) => (SetSPplus (setspolynomial_of l) (setspolynomial_of r))
+ | (SetPmult l r) => (SetSPmult (setspolynomial_of l) (setspolynomial_of r))
+ | (SetPconst c) => (SetSPconst c)
+ | (SetPvar i) => (SetSPvar i)
+ | (SetPopp p) => (SetSPmult (SetSPconst (Aopp Aone)) (setspolynomial_of p))
+ end.
+
+(*** Interpretation *)
+
+Fixpoint interp_setp [p:setpolynomial] : A :=
+ Cases p of
+ | (SetPconst c) => c
+ | (SetPvar i) => (varmap_find Azero i vm)
+ | (SetPplus p1 p2) => (Aplus (interp_setp p1) (interp_setp p2))
+ | (SetPmult p1 p2) => (Amult (interp_setp p1) (interp_setp p2))
+ | (SetPopp p1) => (Aopp (interp_setp p1))
+ end.
+
+(*** Properties *)
+
+Unset Implicit Arguments.
+
+Lemma setspolynomial_of_ok : (p:setpolynomial)
+ (Aequiv (interp_setp p) (interp_setsp vm (setspolynomial_of p))).
+Induction p; Trivial; Simpl; Intros.
+Rewrite H; Rewrite H0; Trivial.
+Rewrite H; Rewrite H0; Trivial.
+Rewrite H.
+Rewrite (STh_opp_mult_left2 S plus_morph mult_morph T Aone
+ (interp_setsp vm (setspolynomial_of s))).
+Rewrite (STh_mult_one_left T
+ (interp_setsp vm (setspolynomial_of s))).
+Trivial.
+Save.
+
+Theorem setpolynomial_normalize_ok : (p:setpolynomial)
+ (setpolynomial_normalize p)
+ ==(setspolynomial_normalize (setspolynomial_of p)).
+Induction p; Trivial; Simpl; Intros.
+Rewrite H; Rewrite H0; Reflexivity.
+Rewrite H; Rewrite H0; Reflexivity.
+Rewrite H; Simpl.
+Elim (canonical_sum_scalar3 (Aopp Aone) Nil_var
+ (setspolynomial_normalize (setspolynomial_of s)));
+ [ Reflexivity
+ | Simpl; Intros; Rewrite H0; Reflexivity
+ | Simpl; Intros; Rewrite H0; Reflexivity ].
+Save.
+
+Theorem setpolynomial_simplify_ok : (p:setpolynomial)
+ (Aequiv (interp_setcs vm (setpolynomial_simplify p)) (interp_setp p)).
+Intro.
+Unfold setpolynomial_simplify.
+Rewrite (setspolynomial_of_ok p).
+Rewrite setpolynomial_normalize_ok.
+Rewrite (canonical_sum_simplify_ok vm
+ (Semi_Setoid_Ring_Theory_of A Aequiv S Aplus Amult Aone Azero Aopp
+ Aeq plus_morph mult_morph T)
+ (setspolynomial_normalize (setspolynomial_of p))).
+Rewrite (setspolynomial_normalize_ok vm
+ (Semi_Setoid_Ring_Theory_of A Aequiv S Aplus Amult Aone Azero Aopp
+ Aeq plus_morph mult_morph T) (setspolynomial_of p)).
+Trivial.
+Save.
+
+End setoid_rings.
+
+End setoid.
diff --git a/contrib7/ring/Setoid_ring_theory.v b/contrib7/ring/Setoid_ring_theory.v
new file mode 100644
index 00000000..13afc5ee
--- /dev/null
+++ b/contrib7/ring/Setoid_ring_theory.v
@@ -0,0 +1,429 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: Setoid_ring_theory.v,v 1.1.2.1 2004/07/16 19:30:19 herbelin Exp $ *)
+
+Require Export Bool.
+Require Export Setoid.
+
+Set Implicit Arguments.
+
+Section Setoid_rings.
+
+Variable A : Type.
+Variable Aequiv : A -> A -> Prop.
+
+Infix Local "==" Aequiv (at level 5, no associativity).
+
+Variable S : (Setoid_Theory A Aequiv).
+
+Add Setoid A Aequiv S.
+
+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.
+
+Infix 4 "+" Aplus V8only 50 (left associativity).
+Infix 4 "*" Amult V8only 40 (left associativity).
+Notation "0" := Azero.
+Notation "1" := Aone.
+Notation "- x" := (Aopp x) (at level 0) V8only.
+
+Variable plus_morph : (a,a0,a1,a2:A) a == a0 -> a1 == a2 -> a+a1 == a0+a2.
+Variable mult_morph : (a,a0,a1,a2:A) a == a0 -> a1 == a2 -> a*a1 == a0*a2.
+Variable opp_morph : (a,a0:A) a == a0 -> -a == -a0.
+
+Add Morphism Aplus : Aplus_ext.
+Exact plus_morph.
+Save.
+
+Add Morphism Amult : Amult_ext.
+Exact mult_morph.
+Save.
+
+Add Morphism Aopp : Aopp_ext.
+Exact opp_morph.
+Save.
+
+Section Theory_of_semi_setoid_rings.
+
+Record Semi_Setoid_Ring_Theory : Prop :=
+{ SSR_plus_sym : (n,m:A) n + m == m + n;
+ SSR_plus_assoc : (n,m,p:A) n + (m + p) == (n + m) + p;
+ SSR_mult_sym : (n,m:A) n*m == m*n;
+ SSR_mult_assoc : (n,m,p:A) n*(m*p) == (n*m)*p;
+ SSR_plus_zero_left :(n:A) 0 + n == n;
+ SSR_mult_one_left : (n:A) 1*n == n;
+ SSR_mult_zero_left : (n:A) 0*n == 0;
+ SSR_distr_left : (n,m,p:A) (n + m)*p == n*p + m*p;
+ SSR_plus_reg_left : (n,m,p:A)n + m == n + p -> m == p;
+ SSR_eq_prop : (x,y:A) (Is_true (Aeq x y)) -> x == y
+}.
+
+Variable T : Semi_Setoid_Ring_Theory.
+
+Local plus_sym := (SSR_plus_sym T).
+Local plus_assoc := (SSR_plus_assoc T).
+Local mult_sym := ( SSR_mult_sym T).
+Local mult_assoc := (SSR_mult_assoc T).
+Local plus_zero_left := (SSR_plus_zero_left T).
+Local mult_one_left := (SSR_mult_one_left T).
+Local mult_zero_left := (SSR_mult_zero_left T).
+Local distr_left := (SSR_distr_left T).
+Local plus_reg_left := (SSR_plus_reg_left T).
+Local equiv_refl := (Seq_refl A Aequiv S).
+Local equiv_sym := (Seq_sym A Aequiv S).
+Local equiv_trans := (Seq_trans A Aequiv S).
+
+Hints Resolve plus_sym plus_assoc mult_sym mult_assoc
+ plus_zero_left mult_one_left mult_zero_left distr_left
+ plus_reg_left equiv_refl (*equiv_sym*).
+Hints Immediate equiv_sym.
+
+(* Lemmas whose form is x=y are also provided in form y=x because
+ Auto does not symmetry *)
+Lemma SSR_mult_assoc2 : (n,m,p:A) (n * m) * p == n * (m * p).
+Auto. Save.
+
+Lemma SSR_plus_assoc2 : (n,m,p:A) (n + m) + p == n + (m + p).
+Auto. Save.
+
+Lemma SSR_plus_zero_left2 : (n:A) n == 0 + n.
+Auto. Save.
+
+Lemma SSR_mult_one_left2 : (n:A) n == 1*n.
+Auto. Save.
+
+Lemma SSR_mult_zero_left2 : (n:A) 0 == 0*n.
+Auto. Save.
+
+Lemma SSR_distr_left2 : (n,m,p:A) n*p + m*p == (n + m)*p.
+Auto. Save.
+
+Lemma SSR_plus_permute : (n,m,p:A) n+(m+p) == m+(n+p).
+Intros.
+Rewrite (plus_assoc n m p).
+Rewrite (plus_sym n m).
+Rewrite <- (plus_assoc m n p).
+Trivial.
+Save.
+
+Lemma SSR_mult_permute : (n,m,p:A) n*(m*p) == m*(n*p).
+Intros.
+Rewrite (mult_assoc n m p).
+Rewrite (mult_sym n m).
+Rewrite <- (mult_assoc m n p).
+Trivial.
+Save.
+
+Hints Resolve SSR_plus_permute SSR_mult_permute.
+
+Lemma SSR_distr_right : (n,m,p:A) n*(m+p) == (n*m) + (n*p).
+Intros.
+Rewrite (mult_sym n (Aplus m p)).
+Rewrite (mult_sym n m).
+Rewrite (mult_sym n p).
+Auto.
+Save.
+
+Lemma SSR_distr_right2 : (n,m,p:A) (n*m) + (n*p) == n*(m + p).
+Intros.
+Apply equiv_sym.
+Apply SSR_distr_right.
+Save.
+
+Lemma SSR_mult_zero_right : (n:A) n*0 == 0.
+Intro; Rewrite (mult_sym n Azero); Auto.
+Save.
+
+Lemma SSR_mult_zero_right2 : (n:A) 0 == n*0.
+Intro; Rewrite (mult_sym n Azero); Auto.
+Save.
+
+Lemma SSR_plus_zero_right :(n:A) n + 0 == n.
+Intro; Rewrite (plus_sym n Azero); Auto.
+Save.
+
+Lemma SSR_plus_zero_right2 :(n:A) n == n + 0.
+Intro; Rewrite (plus_sym n Azero); Auto.
+Save.
+
+Lemma SSR_mult_one_right : (n:A) n*1 == n.
+Intro; Rewrite (mult_sym n Aone); Auto.
+Save.
+
+Lemma SSR_mult_one_right2 : (n:A) n == n*1.
+Intro; Rewrite (mult_sym n Aone); Auto.
+Save.
+
+Lemma SSR_plus_reg_right : (n,m,p:A) m+n == p+n -> m==p.
+Intros n m p; Rewrite (plus_sym m n); Rewrite (plus_sym p n).
+Intro; Apply plus_reg_left with n; Trivial.
+Save.
+
+End Theory_of_semi_setoid_rings.
+
+Section Theory_of_setoid_rings.
+
+Record Setoid_Ring_Theory : Prop :=
+{ STh_plus_sym : (n,m:A) n + m == m + n;
+ STh_plus_assoc : (n,m,p:A) n + (m + p) == (n + m) + p;
+ STh_mult_sym : (n,m:A) n*m == m*n;
+ STh_mult_assoc : (n,m,p:A) n*(m*p) == (n*m)*p;
+ STh_plus_zero_left :(n:A) 0 + n == n;
+ STh_mult_one_left : (n:A) 1*n == n;
+ STh_opp_def : (n:A) n + (-n) == 0;
+ STh_distr_left : (n,m,p:A) (n + m)*p == n*p + m*p;
+ STh_eq_prop : (x,y:A) (Is_true (Aeq x y)) -> x == y
+}.
+
+Variable T : Setoid_Ring_Theory.
+
+Local plus_sym := (STh_plus_sym T).
+Local plus_assoc := (STh_plus_assoc T).
+Local mult_sym := (STh_mult_sym T).
+Local mult_assoc := (STh_mult_assoc T).
+Local plus_zero_left := (STh_plus_zero_left T).
+Local mult_one_left := (STh_mult_one_left T).
+Local opp_def := (STh_opp_def T).
+Local distr_left := (STh_distr_left T).
+Local equiv_refl := (Seq_refl A Aequiv S).
+Local equiv_sym := (Seq_sym A Aequiv S).
+Local equiv_trans := (Seq_trans A Aequiv S).
+
+Hints Resolve plus_sym plus_assoc mult_sym mult_assoc
+ plus_zero_left mult_one_left opp_def distr_left
+ equiv_refl equiv_sym.
+
+(* Lemmas whose form is x=y are also provided in form y=x because Auto does
+ not symmetry *)
+
+Lemma STh_mult_assoc2 : (n,m,p:A) (n * m) * p == n * (m * p).
+Auto. Save.
+
+Lemma STh_plus_assoc2 : (n,m,p:A) (n + m) + p == n + (m + p).
+Auto. Save.
+
+Lemma STh_plus_zero_left2 : (n:A) n == 0 + n.
+Auto. Save.
+
+Lemma STh_mult_one_left2 : (n:A) n == 1*n.
+Auto. Save.
+
+Lemma STh_distr_left2 : (n,m,p:A) n*p + m*p == (n + m)*p.
+Auto. Save.
+
+Lemma STh_opp_def2 : (n:A) 0 == n + (-n).
+Auto. Save.
+
+Lemma STh_plus_permute : (n,m,p:A) n + (m + p) == m + (n + p).
+Intros.
+Rewrite (plus_assoc n m p).
+Rewrite (plus_sym n m).
+Rewrite <- (plus_assoc m n p).
+Trivial.
+Save.
+
+Lemma STh_mult_permute : (n,m,p:A) n*(m*p) == m*(n*p).
+Intros.
+Rewrite (mult_assoc n m p).
+Rewrite (mult_sym n m).
+Rewrite <- (mult_assoc m n p).
+Trivial.
+Save.
+
+Hints Resolve STh_plus_permute STh_mult_permute.
+
+Lemma Saux1 : (a:A) a + a == a -> a == 0.
+Intros.
+Rewrite <- (plus_zero_left a).
+Rewrite (plus_sym Azero a).
+Setoid_replace (Aplus a Azero) with (Aplus a (Aplus a (Aopp a))); Auto.
+Rewrite (plus_assoc a a (Aopp a)).
+Rewrite H.
+Apply opp_def.
+Save.
+
+Lemma STh_mult_zero_left :(n:A) 0*n == 0.
+Intros.
+Apply Saux1.
+Rewrite <- (distr_left Azero Azero n).
+Rewrite (plus_zero_left Azero).
+Trivial.
+Save.
+Hints Resolve STh_mult_zero_left.
+
+Lemma STh_mult_zero_left2 : (n:A) 0 == 0*n.
+Auto.
+Save.
+
+Lemma Saux2 : (x,y,z:A) x+y==0 -> x+z==0 -> y == z.
+Intros.
+Rewrite <- (plus_zero_left y).
+Rewrite <- H0.
+Rewrite <- (plus_assoc x z y).
+Rewrite (plus_sym z y).
+Rewrite (plus_assoc x y z).
+Rewrite H.
+Auto.
+Save.
+
+Lemma STh_opp_mult_left : (x,y:A) -(x*y) == (-x)*y.
+Intros.
+Apply Saux2 with (Amult x y); Auto.
+Rewrite <- (distr_left x (Aopp x) y).
+Rewrite (opp_def x).
+Auto.
+Save.
+Hints Resolve STh_opp_mult_left.
+
+Lemma STh_opp_mult_left2 : (x,y:A) (-x)*y == -(x*y) .
+Auto.
+Save.
+
+Lemma STh_mult_zero_right : (n:A) n*0 == 0.
+Intro; Rewrite (mult_sym n Azero); Auto.
+Save.
+
+Lemma STh_mult_zero_right2 : (n:A) 0 == n*0.
+Intro; Rewrite (mult_sym n Azero); Auto.
+Save.
+
+Lemma STh_plus_zero_right :(n:A) n + 0 == n.
+Intro; Rewrite (plus_sym n Azero); Auto.
+Save.
+
+Lemma STh_plus_zero_right2 :(n:A) n == n + 0.
+Intro; Rewrite (plus_sym n Azero); Auto.
+Save.
+
+Lemma STh_mult_one_right : (n:A) n*1 == n.
+Intro; Rewrite (mult_sym n Aone); Auto.
+Save.
+
+Lemma STh_mult_one_right2 : (n:A) n == n*1.
+Intro; Rewrite (mult_sym n Aone); Auto.
+Save.
+
+Lemma STh_opp_mult_right : (x,y:A) -(x*y) == x*(-y).
+Intros.
+Rewrite (mult_sym x y).
+Rewrite (mult_sym x (Aopp y)).
+Auto.
+Save.
+
+Lemma STh_opp_mult_right2 : (x,y:A) x*(-y) == -(x*y).
+Intros.
+Rewrite (mult_sym x y).
+Rewrite (mult_sym x (Aopp y)).
+Auto.
+Save.
+
+Lemma STh_plus_opp_opp : (x,y:A) (-x) + (-y) == -(x+y).
+Intros.
+Apply Saux2 with (Aplus x y); Auto.
+Rewrite (STh_plus_permute (Aplus x y) (Aopp x) (Aopp y)).
+Rewrite <- (plus_assoc x y (Aopp y)).
+Rewrite (opp_def y); Rewrite (STh_plus_zero_right x).
+Rewrite (STh_opp_def2 x); Trivial.
+Save.
+
+Lemma STh_plus_permute_opp: (n,m,p:A) (-m)+(n+p) == n+((-m)+p).
+Auto.
+Save.
+
+Lemma STh_opp_opp : (n:A) -(-n) == n.
+Intro.
+Apply Saux2 with (Aopp n); Auto.
+Rewrite (plus_sym (Aopp n) n); Auto.
+Save.
+Hints Resolve STh_opp_opp.
+
+Lemma STh_opp_opp2 : (n:A) n == -(-n).
+Auto.
+Save.
+
+Lemma STh_mult_opp_opp : (x,y:A) (-x)*(-y) == x*y.
+Intros.
+Rewrite (STh_opp_mult_left2 x (Aopp y)).
+Rewrite (STh_opp_mult_right2 x y).
+Trivial.
+Save.
+
+Lemma STh_mult_opp_opp2 : (x,y:A) x*y == (-x)*(-y).
+Intros.
+Apply equiv_sym.
+Apply STh_mult_opp_opp.
+Save.
+
+Lemma STh_opp_zero : -0 == 0.
+Rewrite <- (plus_zero_left (Aopp Azero)).
+Trivial.
+Save.
+
+Lemma STh_plus_reg_left : (n,m,p:A) n+m == n+p -> m==p.
+Intros.
+Rewrite <- (plus_zero_left m).
+Rewrite <- (plus_zero_left p).
+Rewrite <- (opp_def n).
+Rewrite (plus_sym n (Aopp n)).
+Rewrite <- (plus_assoc (Aopp n) n m).
+Rewrite <- (plus_assoc (Aopp n) n p).
+Auto.
+Save.
+
+Lemma STh_plus_reg_right : (n,m,p:A) m+n == p+n -> m==p.
+Intros.
+Apply STh_plus_reg_left with n.
+Rewrite (plus_sym n m); Rewrite (plus_sym n p);
+Assumption.
+Save.
+
+Lemma STh_distr_right : (n,m,p:A) n*(m+p) == (n*m)+(n*p).
+Intros.
+Rewrite (mult_sym n (Aplus m p)).
+Rewrite (mult_sym n m).
+Rewrite (mult_sym n p).
+Trivial.
+Save.
+
+Lemma STh_distr_right2 : (n,m,p:A) (n*m)+(n*p) == n*(m+p).
+Intros.
+Apply equiv_sym.
+Apply STh_distr_right.
+Save.
+
+End Theory_of_setoid_rings.
+
+Hints Resolve STh_mult_zero_left STh_plus_reg_left : core.
+
+Unset Implicit Arguments.
+
+Definition Semi_Setoid_Ring_Theory_of :
+ Setoid_Ring_Theory -> Semi_Setoid_Ring_Theory.
+Intros until 1; Case H.
+Split; Intros; Simpl; EAuto.
+Defined.
+
+Coercion Semi_Setoid_Ring_Theory_of :
+ Setoid_Ring_Theory >-> Semi_Setoid_Ring_Theory.
+
+
+
+Section product_ring.
+
+End product_ring.
+
+Section power_ring.
+
+End power_ring.
+
+End Setoid_rings.
diff --git a/contrib7/ring/ZArithRing.v b/contrib7/ring/ZArithRing.v
new file mode 100644
index 00000000..fc7ef29f
--- /dev/null
+++ b/contrib7/ring/ZArithRing.v
@@ -0,0 +1,35 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(* $Id: ZArithRing.v,v 1.1.2.1 2004/07/16 19:30:19 herbelin Exp $ *)
+
+(* Instantiation of the Ring tactic for the binary integers of ZArith *)
+
+Require Export ArithRing.
+Require Export ZArith_base.
+Require Eqdep_dec.
+
+Definition Zeq := [x,y:Z]
+ Cases `x ?= y ` of
+ EGAL => true
+ | _ => false
+ end.
+
+Lemma Zeq_prop : (x,y:Z)(Is_true (Zeq x y)) -> x==y.
+ Intros x y H; Unfold Zeq in H.
+ Apply Zcompare_EGAL_eq.
+ NewDestruct (Zcompare x y); [Reflexivity | Contradiction | Contradiction ].
+Save.
+
+Definition ZTheory : (Ring_Theory Zplus Zmult `1` `0` Zopp Zeq).
+ Split; Intros; Apply eq2eqT; EAuto with zarith.
+ Apply eqT2eq; Apply Zeq_prop; Assumption.
+Save.
+
+(* NatConstants and NatTheory are defined in Ring_theory.v *)
+Add Ring Z Zplus Zmult `1` `0` Zopp Zeq ZTheory [POS NEG ZERO xO xI xH].
diff --git a/contrib7/romega/ROmega.v b/contrib7/romega/ROmega.v
new file mode 100644
index 00000000..7ee246c7
--- /dev/null
+++ b/contrib7/romega/ROmega.v
@@ -0,0 +1,12 @@
+(*************************************************************************
+
+ PROJET RNRT Calife - 2001
+ Author: Pierre Crégut - France Télécom R&D
+ Licence : LGPL version 2.1
+
+ *************************************************************************)
+
+Require Omega.
+Require ReflOmegaCore.
+
+
diff --git a/contrib7/romega/ReflOmegaCore.v b/contrib7/romega/ReflOmegaCore.v
new file mode 100644
index 00000000..81baa8d9
--- /dev/null
+++ b/contrib7/romega/ReflOmegaCore.v
@@ -0,0 +1,2602 @@
+(*************************************************************************
+
+ PROJET RNRT Calife - 2001
+ Author: Pierre Crégut - France Télécom R&D
+ Licence du projet : LGPL version 2.1
+
+ *************************************************************************)
+
+Require Arith.
+Require PolyList.
+Require Bool.
+Require ZArith.
+Require Import OmegaLemmas.
+
+(* \subsection{Definition of basic types} *)
+
+(* \subsubsection{Environment of propositions (lists) *)
+Inductive PropList : Type :=
+ Pnil : PropList | Pcons : Prop -> PropList -> PropList.
+
+(* Access function for the environment with a default *)
+Fixpoint nthProp [n:nat; l:PropList] : Prop -> Prop :=
+ [default]Cases n l of
+ O (Pcons x l') => x
+ | O other => default
+ | (S m) Pnil => default
+ | (S m) (Pcons x t) => (nthProp m t default)
+ end.
+
+(* \subsubsection{Définition of reified integer expressions}
+ Terms are either:
+ \begin{itemize}
+ \item integers [Tint]
+ \item variables [Tvar]
+ \item operation over integers (addition, product, opposite, subtraction)
+ The last two are translated in additions and products. *)
+
+Inductive term : Set :=
+ Tint : Z -> term
+ | Tplus : term -> term -> term
+ | Tmult : term -> term -> term
+ | Tminus : term -> term -> term
+ | Topp : term -> term
+ | Tvar : nat -> term
+.
+
+(* \subsubsection{Definition of reified goals} *)
+(* Very restricted definition of handled predicates that should be extended
+ to cover a wider set of operations.
+ Taking care of negations and disequations require solving more than a
+ goal in parallel. This is a major improvement over previous versions. *)
+
+Inductive proposition : Set :=
+ EqTerm : term -> term -> proposition (* egalité entre termes *)
+| LeqTerm : term -> term -> proposition (* plus petit ou egal *)
+| TrueTerm : proposition (* vrai *)
+| FalseTerm : proposition (* faux *)
+| Tnot : proposition -> proposition (* négation *)
+| GeqTerm : term -> term -> proposition
+| GtTerm : term -> term -> proposition
+| LtTerm : term -> term -> proposition
+| NeqTerm: term -> term -> proposition
+| Tor : proposition -> proposition -> proposition
+| Tand : proposition -> proposition -> proposition
+| Timp : proposition -> proposition -> proposition
+| Tprop : nat -> proposition
+.
+
+(* Definition of goals as a list of hypothesis *)
+Syntactic Definition hyps := (list proposition).
+
+(* Definition of lists of subgoals (set of open goals) *)
+Syntactic Definition lhyps := (list hyps).
+
+(* a syngle goal packed in a subgoal list *)
+Syntactic Definition singleton := [a: hyps] (cons a (nil hyps)).
+
+(* an absurd goal *)
+Definition absurd := (cons FalseTerm (nil proposition)).
+
+(* \subsubsection{Traces for merging equations}
+ This inductive type describes how the monomial of two equations should be
+ merged when the equations are added.
+
+ For [F_equal], both equations have the same head variable and coefficient
+ must be added, furthermore if coefficients are opposite, [F_cancel] should
+ be used to collapse the term. [F_left] and [F_right] indicate which monomial
+ should be put first in the result *)
+
+Inductive t_fusion : Set :=
+ F_equal : t_fusion | F_cancel : t_fusion
+ | F_left : t_fusion | F_right : t_fusion.
+
+(* \subsubsection{Rewriting steps to normalize terms} *)
+Inductive step : Set :=
+ (* apply the rewriting steps to both subterms of an operation *)
+ | C_DO_BOTH : step -> step -> step
+ (* apply the rewriting step to the first branch *)
+ | C_LEFT : step -> step
+ (* apply the rewriting step to the second branch *)
+ | C_RIGHT : step -> step
+ (* apply two steps consecutively to a term *)
+ | C_SEQ : step -> step -> step
+ (* empty step *)
+ | C_NOP : step
+ (* the following operations correspond to actual rewriting *)
+ | C_OPP_PLUS : step
+ | C_OPP_OPP : step
+ | C_OPP_MULT_R : step
+ | C_OPP_ONE : step
+ (* This is a special step that reduces the term (computation) *)
+ | C_REDUCE : step
+ | C_MULT_PLUS_DISTR : step
+ | C_MULT_OPP_LEFT : step
+ | C_MULT_ASSOC_R : step
+ | C_PLUS_ASSOC_R : step
+ | C_PLUS_ASSOC_L : step
+ | C_PLUS_PERMUTE : step
+ | C_PLUS_SYM : step
+ | C_RED0 : step
+ | C_RED1 : step
+ | C_RED2 : step
+ | C_RED3 : step
+ | C_RED4 : step
+ | C_RED5 : step
+ | C_RED6 : step
+ | C_MULT_ASSOC_REDUCED : step
+ | C_MINUS :step
+ | C_MULT_SYM : step
+.
+
+(* \subsubsection{Omega steps} *)
+(* The following inductive type describes steps as they can be found in
+ the trace coming from the decision procedure Omega. *)
+
+Inductive t_omega : Set :=
+ (* n = 0 n!= 0 *)
+ | O_CONSTANT_NOT_NUL : nat -> t_omega
+ | O_CONSTANT_NEG : nat -> t_omega
+ (* division et approximation of an equation *)
+ | O_DIV_APPROX : Z -> Z -> term -> nat -> t_omega -> nat -> t_omega
+ (* no solution because no exact division *)
+ | O_NOT_EXACT_DIVIDE : Z -> Z -> term -> nat -> nat -> t_omega
+ (* exact division *)
+ | O_EXACT_DIVIDE : Z -> term -> nat -> t_omega -> nat -> t_omega
+ | O_SUM : Z -> nat -> Z -> nat -> (list t_fusion) -> t_omega -> t_omega
+ | O_CONTRADICTION : nat -> nat -> nat -> t_omega
+ | O_MERGE_EQ : nat -> nat -> nat -> t_omega -> t_omega
+ | O_SPLIT_INEQ : nat -> nat -> t_omega -> t_omega -> t_omega
+ | O_CONSTANT_NUL : nat -> t_omega
+ | O_NEGATE_CONTRADICT : nat -> nat -> t_omega
+ | O_NEGATE_CONTRADICT_INV : nat -> nat -> nat -> t_omega
+ | O_STATE : Z -> step -> nat -> nat -> t_omega -> t_omega.
+
+(* \subsubsection{Règles pour normaliser les hypothèses} *)
+(* Ces règles indiquent comment normaliser les propositions utiles
+ de chaque hypothèse utile avant la décomposition des hypothèses et
+ incluent l'étape d'inversion pour la suppression des négations *)
+Inductive p_step : Set :=
+ P_LEFT : p_step -> p_step
+| P_RIGHT : p_step -> p_step
+| P_INVERT : step -> p_step
+| P_STEP : step -> p_step
+| P_NOP : p_step
+.
+(* Liste des normalisations a effectuer : avec un constructeur dans le
+ type [p_step] permettant
+ de parcourir à la fois les branches gauches et droit, on pourrait n'avoir
+ qu'une normalisation par hypothèse. Et comme toutes les hypothèses sont
+ utiles (sinon on ne les incluerait pas), on pourrait remplacer [h_step]
+ par une simple liste *)
+
+Inductive h_step : Set := pair_step : nat -> p_step -> h_step.
+
+(* \subsubsection{Règles pour décomposer les hypothèses} *)
+(* Ce type permet de se diriger dans les constructeurs logiques formant les
+ prédicats des hypothèses pour aller les décomposer. Ils permettent
+ en particulier d'extraire une hypothèse d'une conjonction avec
+ éventuellement le bon niveau de négations. *)
+
+Inductive direction : Set :=
+ D_left : direction
+ | D_right : direction
+ | D_mono : direction.
+
+(* Ce type permet d'extraire les composants utiles des hypothèses : que ce
+ soit des hypothèses générées par éclatement d'une disjonction, ou
+ des équations. Le constructeur terminal indique comment résoudre le système
+ obtenu en recourrant au type de trace d'Omega [t_omega] *)
+
+Inductive e_step : Set :=
+ E_SPLIT : nat -> (list direction) -> e_step -> e_step -> e_step
+ | E_EXTRACT : nat -> (list direction) -> e_step -> e_step
+ | E_SOLVE : t_omega -> e_step.
+
+(* \subsection{Egalité décidable efficace} *)
+(* Pour chaque type de donnée réifié, on calcule un test d'égalité efficace.
+ Ce n'est pas le cas de celui rendu par [Decide Equality].
+
+ Puis on prouve deux théorèmes permettant d'éliminer de telles égalités :
+ \begin{verbatim}
+ (t1,t2: typ) (eq_typ t1 t2) = true -> t1 = t2.
+ (t1,t2: typ) (eq_typ t1 t2) = false -> ~ t1 = t2.
+ \end{verbatim} *)
+
+(* Ces deux tactiques permettent de résoudre pas mal de cas. L'une pour
+ les théorèmes positifs, l'autre pour les théorèmes négatifs *)
+
+Tactic Definition absurd_case := Simpl; Intros; Discriminate.
+Tactic Definition trivial_case := Unfold not; Intros; Discriminate.
+
+(* \subsubsection{Entiers naturels} *)
+
+Fixpoint eq_nat [t1,t2: nat] : bool :=
+ Cases t1 of
+ O => Cases t2 of O => true | _ => false end
+ | (S n1)=> Cases t2 of O => false | (S n2) => (eq_nat n1 n2) end
+ end.
+
+Theorem eq_nat_true : (t1,t2: nat) (eq_nat t1 t2) = true -> t1 = t2.
+
+Induction t1; [
+ Intro t2; Case t2; [ Trivial | absurd_case ]
+| Intros n H t2; Case t2;
+ [ absurd_case | Simpl; Intros; Rewrite (H n0); [ Trivial | Assumption]]].
+
+Save.
+
+Theorem eq_nat_false : (t1,t2: nat) (eq_nat t1 t2) = false -> ~t1 = t2.
+
+Induction t1; [
+ Intro t2; Case t2;
+ [ Simpl;Intros; Discriminate | trivial_case ]
+| Intros n H t2; Case t2; Simpl; Unfold not; Intros; [
+ Discriminate
+ | Elim (H n0 H0); Simplify_eq H1; Trivial]].
+
+Save.
+
+
+(* \subsubsection{Entiers positifs} *)
+
+Fixpoint eq_pos [p1,p2 : positive] : bool :=
+ Cases p1 of
+ (xI n1) => Cases p2 of (xI n2) => (eq_pos n1 n2) | _ => false end
+ | (xO n1) => Cases p2 of (xO n2) => (eq_pos n1 n2) | _ => false end
+ | xH => Cases p2 of xH => true | _ => false end
+ end.
+
+Theorem eq_pos_true : (t1,t2: positive) (eq_pos t1 t2) = true -> t1 = t2.
+
+Induction t1; [
+ Intros p H t2; Case t2; [
+ Simpl; Intros; Rewrite (H p0 H0); Trivial | absurd_case | absurd_case ]
+| Intros p H t2; Case t2; [
+ absurd_case | Simpl; Intros; Rewrite (H p0 H0); Trivial | absurd_case ]
+| Intro t2; Case t2; [ absurd_case | absurd_case | Auto ]].
+
+Save.
+
+Theorem eq_pos_false : (t1,t2: positive) (eq_pos t1 t2) = false -> ~t1 = t2.
+
+Induction t1; [
+ Intros p H t2; Case t2; [
+ Simpl; Unfold not; Intros; Elim (H p0 H0); Simplify_eq H1; Auto
+ | trivial_case | trivial_case ]
+| Intros p H t2; Case t2; [
+ trivial_case
+ | Simpl; Unfold not; Intros; Elim (H p0 H0); Simplify_eq H1; Auto
+ | trivial_case ]
+| Intros t2; Case t2; [ trivial_case | trivial_case | absurd_case ]].
+Save.
+
+(* \subsubsection{Entiers relatifs} *)
+
+Definition eq_Z [z1,z2: Z] : bool :=
+ Cases z1 of
+ ZERO => Cases z2 of ZERO => true | _ => false end
+ | (POS p1) => Cases z2 of (POS p2) => (eq_pos p1 p2) | _ => false end
+ | (NEG p1) => Cases z2 of (NEG p2) => (eq_pos p1 p2) | _ => false end
+ end.
+
+Theorem eq_Z_true : (t1,t2: Z) (eq_Z t1 t2) = true -> t1 = t2.
+
+Induction t1; [
+ Intros t2; Case t2; [ Auto | absurd_case | absurd_case ]
+| Intros p t2; Case t2; [
+ absurd_case | Simpl; Intros; Rewrite (eq_pos_true p p0 H); Trivial
+ | absurd_case ]
+| Intros p t2; Case t2; [
+ absurd_case | absurd_case
+ | Simpl; Intros; Rewrite (eq_pos_true p p0 H); Trivial ]].
+
+Save.
+
+Theorem eq_Z_false : (t1,t2: Z) (eq_Z t1 t2) = false -> ~(t1 = t2).
+
+Induction t1; [
+ Intros t2; Case t2; [ absurd_case | trivial_case | trivial_case ]
+| Intros p t2; Case t2; [
+ absurd_case
+ | Simpl; Unfold not; Intros; Elim (eq_pos_false p p0 H); Simplify_eq H0; Auto
+ | trivial_case ]
+| Intros p t2; Case t2; [
+ absurd_case | trivial_case
+ | Simpl; Unfold not; Intros; Elim (eq_pos_false p p0 H);
+ Simplify_eq H0; Auto]].
+Save.
+
+(* \subsubsection{Termes réifiés} *)
+
+Fixpoint eq_term [t1,t2: term] : bool :=
+ Cases t1 of
+ (Tint st1) =>
+ Cases t2 of (Tint st2) => (eq_Z st1 st2) | _ => false end
+ | (Tplus st11 st12) =>
+ Cases t2 of
+ (Tplus st21 st22) =>
+ (andb (eq_term st11 st21) (eq_term st12 st22))
+ | _ => false
+ end
+ | (Tmult st11 st12) =>
+ Cases t2 of
+ (Tmult st21 st22) =>
+ (andb (eq_term st11 st21) (eq_term st12 st22))
+ | _ => false
+ end
+ | (Tminus st11 st12) =>
+ Cases t2 of
+ (Tminus st21 st22) =>
+ (andb (eq_term st11 st21) (eq_term st12 st22))
+ | _ => false
+ end
+ | (Topp st1) =>
+ Cases t2 of (Topp st2) => (eq_term st1 st2) | _ => false end
+ | (Tvar st1) =>
+ Cases t2 of (Tvar st2) => (eq_nat st1 st2) | _ => false end
+ end.
+
+Theorem eq_term_true : (t1,t2: term) (eq_term t1 t2) = true -> t1 = t2.
+
+
+Induction t1; Intros until t2; Case t2; Try absurd_case; Simpl; [
+ Intros; Elim eq_Z_true with 1 := H; Trivial
+| Intros t21 t22 H3; Elim andb_prop with 1:= H3; Intros H4 H5;
+ Elim H with 1 := H4; Elim H0 with 1 := H5; Trivial
+| Intros t21 t22 H3; Elim andb_prop with 1:= H3; Intros H4 H5;
+ Elim H with 1 := H4; Elim H0 with 1 := H5; Trivial
+| Intros t21 t22 H3; Elim andb_prop with 1:= H3; Intros H4 H5;
+ Elim H with 1 := H4; Elim H0 with 1 := H5; Trivial
+| Intros t21 H3; Elim H with 1 := H3; Trivial
+| Intros; Elim eq_nat_true with 1 := H; Trivial ].
+
+Save.
+
+Theorem eq_term_false : (t1,t2: term) (eq_term t1 t2) = false -> ~(t1 = t2).
+
+Induction t1; [
+ Intros z t2; Case t2; Try trivial_case; Simpl; Unfold not; Intros;
+ Elim eq_Z_false with 1:=H; Simplify_eq H0; Auto
+| Intros t11 H1 t12 H2 t2; Case t2; Try trivial_case; Simpl; Intros t21 t22 H3;
+ Unfold not; Intro H4; Elim andb_false_elim with 1:= H3; Intros H5;
+ [ Elim H1 with 1 := H5; Simplify_eq H4; Auto |
+ Elim H2 with 1 := H5; Simplify_eq H4; Auto ]
+| Intros t11 H1 t12 H2 t2; Case t2; Try trivial_case; Simpl; Intros t21 t22 H3;
+ Unfold not; Intro H4; Elim andb_false_elim with 1:= H3; Intros H5;
+ [ Elim H1 with 1 := H5; Simplify_eq H4; Auto |
+ Elim H2 with 1 := H5; Simplify_eq H4; Auto ]
+| Intros t11 H1 t12 H2 t2; Case t2; Try trivial_case; Simpl; Intros t21 t22 H3;
+ Unfold not; Intro H4; Elim andb_false_elim with 1:= H3; Intros H5;
+ [ Elim H1 with 1 := H5; Simplify_eq H4; Auto |
+ Elim H2 with 1 := H5; Simplify_eq H4; Auto ]
+| Intros t11 H1 t2; Case t2; Try trivial_case; Simpl; Intros t21 H3;
+ Unfold not; Intro H4; Elim H1 with 1 := H3; Simplify_eq H4; Auto
+| Intros n t2; Case t2; Try trivial_case; Simpl; Unfold not; Intros;
+ Elim eq_nat_false with 1:=H; Simplify_eq H0; Auto ].
+
+Save.
+
+(* \subsubsection{Tactiques pour éliminer ces tests}
+
+ Si on se contente de faire un [Case (eq_typ t1 t2)] on perd
+ totalement dans chaque branche le fait que [t1=t2] ou [~t1=t2].
+
+ Initialement, les développements avaient été réalisés avec les
+ tests rendus par [Decide Equality], c'est à dire un test rendant
+ des termes du type [{t1=t2}+{~t1=t2}]. Faire une élimination sur un
+ tel test préserve bien l'information voulue mais calculatoirement de
+ telles fonctions sont trop lentes. *)
+
+(* Le théorème suivant permet de garder dans les hypothèses la valeur
+ du booléen lors de l'élimination. *)
+
+Theorem bool_ind2 :
+ (P:(bool->Prop)) (b:bool)
+ (b = true -> (P true))->
+ (b = false -> (P false)) -> (P b).
+
+Induction b; Auto.
+Save.
+
+(* Les tactiques définies si après se comportent exactement comme si on
+ avait utilisé le test précédent et fait une elimination dessus. *)
+
+Tactic Definition Elim_eq_term t1 t2 :=
+ Pattern (eq_term t1 t2); Apply bool_ind2; Intro Aux; [
+ Generalize (eq_term_true t1 t2 Aux); Clear Aux
+ | Generalize (eq_term_false t1 t2 Aux); Clear Aux ].
+
+Tactic Definition Elim_eq_Z t1 t2 :=
+ Pattern (eq_Z t1 t2); Apply bool_ind2; Intro Aux; [
+ Generalize (eq_Z_true t1 t2 Aux); Clear Aux
+ | Generalize (eq_Z_false t1 t2 Aux); Clear Aux ].
+
+Tactic Definition Elim_eq_pos t1 t2 :=
+ Pattern (eq_pos t1 t2); Apply bool_ind2; Intro Aux; [
+ Generalize (eq_pos_true t1 t2 Aux); Clear Aux
+ | Generalize (eq_pos_false t1 t2 Aux); Clear Aux ].
+
+(* \subsubsection{Comparaison sur Z} *)
+
+(* Sujet très lié au précédent : on introduit la tactique d'élimination
+ avec son théorème *)
+
+Theorem relation_ind2 :
+ (P:(relation->Prop)) (b:relation)
+ (b = EGAL -> (P EGAL))->
+ (b = INFERIEUR -> (P INFERIEUR))->
+ (b = SUPERIEUR -> (P SUPERIEUR)) -> (P b).
+
+Induction b; Auto.
+Save.
+
+Tactic Definition Elim_Zcompare t1 t2 :=
+ Pattern (Zcompare t1 t2); Apply relation_ind2.
+
+(* \subsection{Interprétations}
+ \subsubsection{Interprétation des termes dans Z} *)
+
+Fixpoint interp_term [env:(list Z); t:term] : Z :=
+ Cases t of
+ (Tint x) => x
+ | (Tplus t1 t2) => (Zplus (interp_term env t1) (interp_term env t2))
+ | (Tmult t1 t2) => (Zmult (interp_term env t1) (interp_term env t2))
+ | (Tminus t1 t2) => (Zminus (interp_term env t1) (interp_term env t2))
+ | (Topp t) => (Zopp (interp_term env t))
+ | (Tvar n) => (nth n env ZERO)
+ end.
+
+(* \subsubsection{Interprétation des prédicats} *)
+Fixpoint interp_proposition
+ [envp : PropList; env: (list Z); p:proposition] : Prop :=
+ Cases p of
+ (EqTerm t1 t2) => ((interp_term env t1) = (interp_term env t2))
+ | (LeqTerm t1 t2) => `(interp_term env t1) <= (interp_term env t2)`
+ | TrueTerm => True
+ | FalseTerm => False
+ | (Tnot p') => ~(interp_proposition envp env p')
+ | (GeqTerm t1 t2) => `(interp_term env t1) >= (interp_term env t2)`
+ | (GtTerm t1 t2) => `(interp_term env t1) > (interp_term env t2)`
+ | (LtTerm t1 t2) => `(interp_term env t1) < (interp_term env t2)`
+ | (NeqTerm t1 t2) => `(Zne (interp_term env t1) (interp_term env t2))`
+
+ | (Tor p1 p2) =>
+ (interp_proposition envp env p1) \/ (interp_proposition envp env p2)
+ | (Tand p1 p2) =>
+ (interp_proposition envp env p1) /\ (interp_proposition envp env p2)
+ | (Timp p1 p2) =>
+ (interp_proposition envp env p1) -> (interp_proposition envp env p2)
+ | (Tprop n) => (nthProp n envp True)
+ end.
+
+(* \subsubsection{Inteprétation des listes d'hypothèses}
+ \paragraph{Sous forme de conjonction}
+ Interprétation sous forme d'une conjonction d'hypothèses plus faciles
+ à manipuler individuellement *)
+
+Fixpoint interp_hyps [envp: PropList; env : (list Z); l: hyps] : Prop :=
+ Cases l of
+ nil => True
+ | (cons p' l') =>
+ (interp_proposition envp env p') /\ (interp_hyps envp env l')
+ end.
+
+(* \paragraph{sous forme de but}
+ C'est cette interpétation que l'on utilise sur le but (car on utilise
+ [Generalize] et qu'une conjonction est forcément lourde (répétition des
+ types dans les conjonctions intermédiaires) *)
+
+Fixpoint interp_goal_concl [envp: PropList;env : (list Z); c: proposition; l: hyps] : Prop :=
+ Cases l of
+ nil => (interp_proposition envp env c)
+ | (cons p' l') =>
+ (interp_proposition envp env p') -> (interp_goal_concl envp env c l')
+ end.
+
+Syntactic Definition interp_goal :=
+ [envp: PropList;env : (list Z); l: hyps]
+ (interp_goal_concl envp env FalseTerm l).
+
+(* Les théorèmes qui suivent assurent la correspondance entre les deux
+ interprétations. *)
+
+Theorem goal_to_hyps :
+ (envp: PropList; env : (list Z); l: hyps)
+ ((interp_hyps envp env l) -> False) -> (interp_goal envp env l).
+
+Induction l; [
+ Simpl; Auto
+| Simpl; Intros a l1 H1 H2 H3; Apply H1; Intro H4; Apply H2; Auto ].
+Save.
+
+Theorem hyps_to_goal :
+ (envp: PropList; env : (list Z); l: hyps)
+ (interp_goal envp env l) -> ((interp_hyps envp env l) -> False).
+
+Induction l; Simpl; [
+ Auto
+| Intros; Apply H; Elim H1; Auto ].
+Save.
+
+(* \subsection{Manipulations sur les hypothèses} *)
+
+(* \subsubsection{Définitions de base de stabilité pour la réflexion} *)
+(* Une opération laisse un terme stable si l'égalité est préservée *)
+Definition term_stable [f: term -> term] :=
+ (e: (list Z); t:term) (interp_term e t) = (interp_term e (f t)).
+
+(* Une opération est valide sur une hypothèse, si l'hypothèse implique le
+ résultat de l'opération. \emph{Attention : cela ne concerne que des
+ opérations sur les hypothèses et non sur les buts (contravariance)}.
+ On définit la validité pour une opération prenant une ou deux propositions
+ en argument (cela suffit pour omega). *)
+
+Definition valid1 [f: proposition -> proposition] :=
+ (ep : PropList; e: (list Z)) (p1: proposition)
+ (interp_proposition ep e p1) -> (interp_proposition ep e (f p1)).
+
+Definition valid2 [f: proposition -> proposition -> proposition] :=
+ (ep : PropList; e: (list Z)) (p1,p2: proposition)
+ (interp_proposition ep e p1) -> (interp_proposition ep e p2) ->
+ (interp_proposition ep e (f p1 p2)).
+
+(* Dans cette notion de validité, la fonction prend directement une
+ liste de propositions et rend une nouvelle liste de proposition.
+ On reste contravariant *)
+
+Definition valid_hyps [f: hyps -> hyps] :=
+ (ep : PropList; e : (list Z))
+ (lp: hyps) (interp_hyps ep e lp) -> (interp_hyps ep e (f lp)).
+
+(* Enfin ce théorème élimine la contravariance et nous ramène à une
+ opération sur les buts *)
+
+ Theorem valid_goal :
+ (ep: PropList; env : (list Z); l: hyps; a : hyps -> hyps)
+ (valid_hyps a) -> (interp_goal ep env (a l)) -> (interp_goal ep env l).
+
+Intros; Simpl; Apply goal_to_hyps; Intro H1;
+Apply (hyps_to_goal ep env (a l) H0); Apply H; Assumption.
+Save.
+
+(* \subsubsection{Généralisation a des listes de buts (disjonctions)} *)
+
+
+Fixpoint interp_list_hyps [envp: PropList; env: (list Z); l : lhyps] : Prop :=
+ Cases l of
+ nil => False
+ | (cons h l') => (interp_hyps envp env h) \/ (interp_list_hyps envp env l')
+ end.
+
+Fixpoint interp_list_goal [envp: PropList; env: (list Z);l : lhyps] : Prop :=
+ Cases l of
+ nil => True
+ | (cons h l') => (interp_goal envp env h) /\ (interp_list_goal envp env l')
+ end.
+
+Theorem list_goal_to_hyps :
+ (envp: PropList; env: (list Z); l: lhyps)
+ ((interp_list_hyps envp env l) -> False) -> (interp_list_goal envp env l).
+
+Induction l; Simpl; [
+ Auto
+| Intros h1 l1 H H1; Split; [
+ Apply goal_to_hyps; Intro H2; Apply H1; Auto
+ | Apply H; Intro H2; Apply H1; Auto ]].
+Save.
+
+Theorem list_hyps_to_goal :
+ (envp: PropList; env: (list Z); l: lhyps)
+ (interp_list_goal envp env l) -> ((interp_list_hyps envp env l) -> False).
+
+Induction l; Simpl; [
+ Auto
+| Intros h1 l1 H (H1,H2) H3; Elim H3; Intro H4; [
+ Apply hyps_to_goal with 1 := H1; Assumption
+ | Auto ]].
+Save.
+
+Definition valid_list_hyps [f: hyps -> lhyps] :=
+ (ep : PropList; e : (list Z)) (lp: hyps)
+ (interp_hyps ep e lp) -> (interp_list_hyps ep e (f lp)).
+
+Definition valid_list_goal [f: hyps -> lhyps] :=
+ (ep : PropList; e : (list Z)) (lp: hyps)
+ (interp_list_goal ep e (f lp)) -> (interp_goal ep e lp) .
+
+Theorem goal_valid :
+ (f: hyps -> lhyps) (valid_list_hyps f) -> (valid_list_goal f).
+
+Unfold valid_list_goal; Intros f H ep e lp H1; Apply goal_to_hyps;
+Intro H2; Apply list_hyps_to_goal with 1:=H1; Apply (H ep e lp); Assumption.
+Save.
+
+Theorem append_valid :
+ (ep: PropList; e: (list Z)) (l1,l2:lhyps)
+ (interp_list_hyps ep e l1) \/ (interp_list_hyps ep e l2) ->
+ (interp_list_hyps ep e (app l1 l2)).
+
+Intros ep e; Induction l1; [
+ Simpl; Intros l2 [H | H]; [ Contradiction | Trivial ]
+| Simpl; Intros h1 t1 HR l2 [[H | H] | H] ;[
+ Auto
+ | Right; Apply (HR l2); Left; Trivial
+ | Right; Apply (HR l2); Right; Trivial ]].
+
+Save.
+
+(* \subsubsection{Opérateurs valides sur les hypothèses} *)
+
+(* Extraire une hypothèse de la liste *)
+Definition nth_hyps [n:nat; l: hyps] := (nth n l TrueTerm).
+
+Theorem nth_valid :
+ (ep: PropList; e: (list Z); i:nat; l: hyps)
+ (interp_hyps ep e l) -> (interp_proposition ep e (nth_hyps i l)).
+
+Unfold nth_hyps; Induction i; [
+ Induction l; Simpl; [ Auto | Intros; Elim H0; Auto ]
+| Intros n H; Induction l;
+ [ Simpl; Trivial | Intros; Simpl; Apply H; Elim H1; Auto ]].
+Save.
+
+(* Appliquer une opération (valide) sur deux hypothèses extraites de
+ la liste et ajouter le résultat à la liste. *)
+Definition apply_oper_2
+ [i,j : nat; f : proposition -> proposition -> proposition ] :=
+ [l: hyps] (cons (f (nth_hyps i l) (nth_hyps j l)) l).
+
+Theorem apply_oper_2_valid :
+ (i,j : nat; f : proposition -> proposition -> proposition )
+ (valid2 f) -> (valid_hyps (apply_oper_2 i j f)).
+
+Intros i j f Hf; Unfold apply_oper_2 valid_hyps; Simpl; Intros lp Hlp; Split;
+ [ Apply Hf; Apply nth_valid; Assumption | Assumption].
+Save.
+
+(* Modifier une hypothèse par application d'une opération valide *)
+
+Fixpoint apply_oper_1 [i:nat] : (proposition -> proposition) -> hyps -> hyps :=
+ [f : (proposition -> proposition); l : hyps]
+ Cases l of
+ nil => (nil proposition)
+ | (cons p l') =>
+ Cases i of
+ O => (cons (f p) l')
+ | (S j) => (cons p (apply_oper_1 j f l'))
+ end
+ end.
+
+Theorem apply_oper_1_valid :
+ (i : nat; f : proposition -> proposition )
+ (valid1 f) -> (valid_hyps (apply_oper_1 i f)).
+
+Unfold valid_hyps; Intros i f Hf ep e; Elim i; [
+ Intro lp; Case lp; [
+ Simpl; Trivial
+ | Simpl; Intros p l' (H1, H2); Split; [ Apply Hf with 1:=H1 | Assumption ]]
+| Intros n Hrec lp; Case lp; [
+ Simpl; Auto
+ | Simpl; Intros p l' (H1, H2);
+ Split; [ Assumption | Apply Hrec; Assumption ]]].
+
+Save.
+
+(* \subsubsection{Manipulations de termes} *)
+(* Les fonctions suivantes permettent d'appliquer une fonction de
+ réécriture sur un sous terme du terme principal. Avec la composition,
+ cela permet de construire des réécritures complexes proches des
+ tactiques de conversion *)
+
+Definition apply_left [f: term -> term; t : term]:=
+ Cases t of
+ (Tplus x y) => (Tplus (f x) y)
+ | (Tmult x y) => (Tmult (f x) y)
+ | (Topp x) => (Topp (f x))
+ | x => x
+ end.
+
+Definition apply_right [f: term -> term; t : term]:=
+ Cases t of
+ (Tplus x y) => (Tplus x (f y))
+ | (Tmult x y) => (Tmult x (f y))
+ | x => x
+ end.
+
+Definition apply_both [f,g: term -> term; t : term]:=
+ Cases t of
+ (Tplus x y) => (Tplus (f x) (g y))
+ | (Tmult x y) => (Tmult (f x) (g y))
+ | x => x
+ end.
+
+(* Les théorèmes suivants montrent la stabilité (conditionnée) des
+ fonctions. *)
+
+Theorem apply_left_stable :
+ (f: term -> term) (term_stable f) -> (term_stable (apply_left f)).
+
+Unfold term_stable; Intros f H e t; Case t; Auto; Simpl;
+Intros; Elim H; Trivial.
+Save.
+
+Theorem apply_right_stable :
+ (f: term -> term) (term_stable f) -> (term_stable (apply_right f)).
+
+Unfold term_stable; Intros f H e t; Case t; Auto; Simpl;
+Intros t0 t1; Elim H; Trivial.
+Save.
+
+Theorem apply_both_stable :
+ (f,g: term -> term) (term_stable f) -> (term_stable g) ->
+ (term_stable (apply_both f g)).
+
+Unfold term_stable; Intros f g H1 H2 e t; Case t; Auto; Simpl;
+Intros t0 t1; Elim H1; Elim H2; Trivial.
+Save.
+
+Theorem compose_term_stable :
+ (f,g: term -> term) (term_stable f) -> (term_stable g) ->
+ (term_stable [t: term](f (g t))).
+
+Unfold term_stable; Intros f g Hf Hg e t; Elim Hf; Apply Hg.
+Save.
+
+(* \subsection{Les règles de réécriture} *)
+(* Chacune des règles de réécriture est accompagnée par sa preuve de
+ stabilité. Toutes ces preuves ont la même forme : il faut analyser
+ suivant la forme du terme (élimination de chaque Case). On a besoin d'une
+ élimination uniquement dans les cas d'utilisation d'égalité décidable.
+
+ Cette tactique itère la décomposition des Case. Elle est
+ constituée de deux fonctions s'appelant mutuellement :
+ \begin{itemize}
+ \item une fonction d'enrobage qui lance la recherche sur le but,
+ \item une fonction récursive qui décompose ce but. Quand elle a trouvé un
+ Case, elle l'élimine.
+ \end{itemize}
+ Les motifs sur les cas sont très imparfaits et dans certains cas, il
+ semble que cela ne marche pas. On aimerait plutot un motif de la
+ forme [ Case (?1 :: T) of _ end ] permettant de s'assurer que l'on
+ utilise le bon type.
+
+ Chaque élimination introduit correctement exactement le nombre d'hypothèses
+ nécessaires et conserve dans le cas d'une égalité la connaissance du
+ résultat du test en faisant la réécriture. Pour un test de comparaison,
+ on conserve simplement le résultat.
+
+ Cette fonction déborde très largement la résolution des réécritures
+ simples et fait une bonne partie des preuves des pas de Omega.
+*)
+
+(* \subsubsection{La tactique pour prouver la stabilité} *)
+
+Recursive Tactic Definition loop t := (
+ Match t With
+ (* Global *)
+ [(?1 = ?2)] -> (loop ?1) Orelse (loop ?2)
+ | [ ? -> ?1 ] -> (loop ?1)
+ (* Interpretations *)
+ | [ (interp_hyps ? ? ?1) ] -> (loop ?1)
+ | [ (interp_list_hyps ? ? ?1) ] -> (loop ?1)
+ | [ (interp_proposition ? ? ?1) ] -> (loop ?1)
+ | [ (interp_term ? ?1) ] -> (loop ?1)
+ (* Propositions *)
+ | [(EqTerm ?1 ?2)] -> (loop ?1) Orelse (loop ?2)
+ | [(LeqTerm ?1 ?2)] -> (loop ?1) Orelse (loop ?2)
+ (* Termes *)
+ | [(Tplus ?1 ?2)] -> (loop ?1) Orelse (loop ?2)
+ | [(Tminus ?1 ?2)] -> (loop ?1) Orelse (loop ?2)
+ | [(Tmult ?1 ?2)] -> (loop ?1) Orelse (loop ?2)
+ | [(Topp ?1)] -> (loop ?1)
+ | [(Tint ?1)] -> (loop ?1)
+ (* Eliminations *)
+ | [(Cases ?1 of
+ | (EqTerm _ _) => ?
+ | (LeqTerm _ _) => ?
+ | TrueTerm => ?
+ | FalseTerm => ?
+ | (Tnot _) => ?
+ | (GeqTerm _ _) => ?
+ | (GtTerm _ _) => ?
+ | (LtTerm _ _) => ?
+ | (NeqTerm _ _) => ?
+ | (Tor _ _) => ?
+ | (Tand _ _) => ?
+ | (Timp _ _) => ?
+ | (Tprop _) => ?
+ end)] ->
+ (Case ?1; [ Intro; Intro | Intro; Intro | Idtac | Idtac
+ | Intro | Intro; Intro | Intro; Intro | Intro; Intro
+ | Intro; Intro
+ | Intro;Intro | Intro;Intro | Intro;Intro | Intro ]);
+ Auto; Simplify
+ | [(Cases ?1 of
+ (Tint _) => ?
+ | (Tplus _ _) => ?
+ | (Tmult _ _) => ?
+ | (Tminus _ _) => ?
+ | (Topp _) => ?
+ | (Tvar _) => ?
+ end)] ->
+ (Case ?1; [ Intro | Intro; Intro | Intro; Intro | Intro; Intro |
+ Intro | Intro ]); Auto; Simplify
+ | [(Cases (Zcompare ?1 ?2) of
+ EGAL => ?
+ | INFERIEUR => ?
+ | SUPERIEUR => ?
+ end)] ->
+ (Elim_Zcompare ?1 ?2) ; Intro ; Auto; Simplify
+ | [(Cases ?1 of ZERO => ? | (POS _) => ? | (NEG _) => ? end)] ->
+ (Case ?1; [ Idtac | Intro | Intro ]); Auto; Simplify
+ | [(if (eq_Z ?1 ?2) then ? else ?)] ->
+ ((Elim_eq_Z ?1 ?2); Intro H; [Rewrite H; Clear H | Clear H]);
+ Simpl; Auto; Simplify
+ | [(if (eq_term ?1 ?2) then ? else ?)] ->
+ ((Elim_eq_term ?1 ?2); Intro H; [Rewrite H; Clear H | Clear H]);
+ Simpl; Auto; Simplify
+ | [(if (eq_pos ?1 ?2) then ? else ?)] ->
+ ((Elim_eq_pos ?1 ?2); Intro H; [Rewrite H; Clear H | Clear H]);
+ Simpl; Auto; Simplify
+ | _ -> Fail)
+And Simplify := (
+ Match Context With [|- ?1 ] -> Try (loop ?1) | _ -> Idtac).
+
+
+Tactic Definition ProveStable x th :=
+ (Match x With [?1] -> Unfold term_stable ?1; Intros; Simplify; Simpl; Apply th).
+
+(* \subsubsection{Les règles elle mêmes} *)
+Definition Tplus_assoc_l [t: term] :=
+ Cases t of
+ (Tplus n (Tplus m p)) => (Tplus (Tplus n m) p)
+ | _ => t
+ end.
+
+Theorem Tplus_assoc_l_stable : (term_stable Tplus_assoc_l).
+
+(ProveStable Tplus_assoc_l Zplus_assoc_l).
+Save.
+
+Definition Tplus_assoc_r [t: term] :=
+ Cases t of
+ (Tplus (Tplus n m) p) => (Tplus n (Tplus m p))
+ | _ => t
+ end.
+
+Theorem Tplus_assoc_r_stable : (term_stable Tplus_assoc_r).
+
+(ProveStable Tplus_assoc_r Zplus_assoc_r).
+Save.
+
+Definition Tmult_assoc_r [t: term] :=
+ Cases t of
+ (Tmult (Tmult n m) p) => (Tmult n (Tmult m p))
+ | _ => t
+ end.
+
+Theorem Tmult_assoc_r_stable : (term_stable Tmult_assoc_r).
+
+(ProveStable Tmult_assoc_r Zmult_assoc_r).
+Save.
+
+Definition Tplus_permute [t: term] :=
+ Cases t of
+ (Tplus n (Tplus m p)) => (Tplus m (Tplus n p))
+ | _ => t
+ end.
+
+Theorem Tplus_permute_stable : (term_stable Tplus_permute).
+
+(ProveStable Tplus_permute Zplus_permute).
+Save.
+
+Definition Tplus_sym [t: term] :=
+ Cases t of
+ (Tplus x y) => (Tplus y x)
+ | _ => t
+ end.
+
+Theorem Tplus_sym_stable : (term_stable Tplus_sym).
+
+(ProveStable Tplus_sym Zplus_sym).
+Save.
+
+Definition Tmult_sym [t: term] :=
+ Cases t of
+ (Tmult x y) => (Tmult y x)
+ | _ => t
+ end.
+
+Theorem Tmult_sym_stable : (term_stable Tmult_sym).
+
+(ProveStable Tmult_sym Zmult_sym).
+Save.
+
+Definition T_OMEGA10 [t: term] :=
+ Cases t of
+ (Tplus (Tmult (Tplus (Tmult v (Tint c1)) l1) (Tint k1))
+ (Tmult (Tplus (Tmult v' (Tint c2)) l2) (Tint k2))) =>
+ Case (eq_term v v') of
+ (Tplus (Tmult v (Tint (Zplus (Zmult c1 k1) (Zmult c2 k2))))
+ (Tplus (Tmult l1 (Tint k1)) (Tmult l2 (Tint k2))))
+ t
+ end
+ | _ => t
+ end.
+
+Theorem T_OMEGA10_stable : (term_stable T_OMEGA10).
+
+(ProveStable T_OMEGA10 OMEGA10).
+Save.
+
+Definition T_OMEGA11 [t: term] :=
+ Cases t of
+ (Tplus (Tmult (Tplus (Tmult v1 (Tint c1)) l1) (Tint k1)) l2) =>
+ (Tplus (Tmult v1 (Tint (Zmult c1 k1))) (Tplus (Tmult l1 (Tint k1)) l2))
+ | _ => t
+ end.
+
+Theorem T_OMEGA11_stable : (term_stable T_OMEGA11).
+
+(ProveStable T_OMEGA11 OMEGA11).
+Save.
+
+Definition T_OMEGA12 [t: term] :=
+ Cases t of
+ (Tplus l1 (Tmult (Tplus (Tmult v2 (Tint c2)) l2) (Tint k2))) =>
+ (Tplus (Tmult v2 (Tint (Zmult c2 k2))) (Tplus l1 (Tmult l2 (Tint k2))))
+ | _ => t
+ end.
+
+Theorem T_OMEGA12_stable : (term_stable T_OMEGA12).
+
+(ProveStable T_OMEGA12 OMEGA12).
+Save.
+
+Definition T_OMEGA13 [t: term] :=
+ Cases t of
+ (Tplus (Tplus (Tmult v (Tint (POS x))) l1)
+ (Tplus (Tmult v' (Tint (NEG x'))) l2)) =>
+ Case (eq_term v v') of
+ Case (eq_pos x x') of
+ (Tplus l1 l2)
+ t
+ end
+ t
+ end
+ | (Tplus (Tplus (Tmult v (Tint (NEG x))) l1)
+ (Tplus (Tmult v' (Tint (POS x'))) l2)) =>
+ Case (eq_term v v') of
+ Case (eq_pos x x') of
+ (Tplus l1 l2)
+ t
+ end
+ t
+ end
+
+ | _ => t
+ end.
+
+Theorem T_OMEGA13_stable : (term_stable T_OMEGA13).
+
+Unfold term_stable T_OMEGA13; Intros; Simplify; Simpl;
+ [ Apply OMEGA13 | Apply OMEGA14 ].
+Save.
+
+Definition T_OMEGA15 [t: term] :=
+ Cases t of
+ (Tplus (Tplus (Tmult v (Tint c1)) l1)
+ (Tmult (Tplus (Tmult v' (Tint c2)) l2) (Tint k2))) =>
+ Case (eq_term v v') of
+ (Tplus (Tmult v (Tint (Zplus c1 (Zmult c2 k2))))
+ (Tplus l1 (Tmult l2 (Tint k2))))
+ t
+ end
+ | _ => t
+ end.
+
+Theorem T_OMEGA15_stable : (term_stable T_OMEGA15).
+
+(ProveStable T_OMEGA15 OMEGA15).
+Save.
+
+Definition T_OMEGA16 [t: term] :=
+ Cases t of
+ (Tmult (Tplus (Tmult v (Tint c)) l) (Tint k)) =>
+ (Tplus (Tmult v (Tint (Zmult c k))) (Tmult l (Tint k)))
+ | _ => t
+ end.
+
+
+Theorem T_OMEGA16_stable : (term_stable T_OMEGA16).
+
+(ProveStable T_OMEGA16 OMEGA16).
+Save.
+
+Definition Tred_factor5 [t: term] :=
+ Cases t of
+ (Tplus (Tmult x (Tint ZERO)) y) => y
+ | _ => t
+ end.
+
+Theorem Tred_factor5_stable : (term_stable Tred_factor5).
+
+
+(ProveStable Tred_factor5 Zred_factor5).
+Save.
+
+Definition Topp_plus [t: term] :=
+ Cases t of
+ (Topp (Tplus x y)) => (Tplus (Topp x) (Topp y))
+ | _ => t
+ end.
+
+Theorem Topp_plus_stable : (term_stable Topp_plus).
+
+(ProveStable Topp_plus Zopp_Zplus).
+Save.
+
+
+Definition Topp_opp [t: term] :=
+ Cases t of
+ (Topp (Topp x)) => x
+ | _ => t
+ end.
+
+Theorem Topp_opp_stable : (term_stable Topp_opp).
+
+(ProveStable Topp_opp Zopp_Zopp).
+Save.
+
+Definition Topp_mult_r [t: term] :=
+ Cases t of
+ (Topp (Tmult x (Tint k))) => (Tmult x (Tint (Zopp k)))
+ | _ => t
+ end.
+
+Theorem Topp_mult_r_stable : (term_stable Topp_mult_r).
+
+(ProveStable Topp_mult_r Zopp_Zmult_r).
+Save.
+
+Definition Topp_one [t: term] :=
+ Cases t of
+ (Topp x) => (Tmult x (Tint `-1`))
+ | _ => t
+ end.
+
+Theorem Topp_one_stable : (term_stable Topp_one).
+
+(ProveStable Topp_one Zopp_one).
+Save.
+
+Definition Tmult_plus_distr [t: term] :=
+ Cases t of
+ (Tmult (Tplus n m) p) => (Tplus (Tmult n p) (Tmult m p))
+ | _ => t
+ end.
+
+Theorem Tmult_plus_distr_stable : (term_stable Tmult_plus_distr).
+
+(ProveStable Tmult_plus_distr Zmult_plus_distr).
+Save.
+
+Definition Tmult_opp_left [t: term] :=
+ Cases t of
+ (Tmult (Topp x) (Tint y)) => (Tmult x (Tint (Zopp y)))
+ | _ => t
+ end.
+
+Theorem Tmult_opp_left_stable : (term_stable Tmult_opp_left).
+
+(ProveStable Tmult_opp_left Zmult_Zopp_left).
+Save.
+
+Definition Tmult_assoc_reduced [t: term] :=
+ Cases t of
+ (Tmult (Tmult n (Tint m)) (Tint p)) => (Tmult n (Tint (Zmult m p)))
+ | _ => t
+ end.
+
+Theorem Tmult_assoc_reduced_stable : (term_stable Tmult_assoc_reduced).
+
+(ProveStable Tmult_assoc_reduced Zmult_assoc_r).
+Save.
+
+Definition Tred_factor0 [t: term] := (Tmult t (Tint `1`)).
+
+Theorem Tred_factor0_stable : (term_stable Tred_factor0).
+
+(ProveStable Tred_factor0 Zred_factor0).
+Save.
+
+Definition Tred_factor1 [t: term] :=
+ Cases t of
+ (Tplus x y) =>
+ Case (eq_term x y) of
+ (Tmult x (Tint `2`))
+ t
+ end
+ | _ => t
+ end.
+
+Theorem Tred_factor1_stable : (term_stable Tred_factor1).
+
+(ProveStable Tred_factor1 Zred_factor1).
+Save.
+
+Definition Tred_factor2 [t: term] :=
+ Cases t of
+ (Tplus x (Tmult y (Tint k))) =>
+ Case (eq_term x y) of
+ (Tmult x (Tint (Zplus `1` k)))
+ t
+ end
+ | _ => t
+ end.
+
+(* Attention : il faut rendre opaque [Zplus] pour éviter que la tactique
+ de simplification n'aille trop loin et défasse [Zplus 1 k] *)
+
+Opaque Zplus.
+
+Theorem Tred_factor2_stable : (term_stable Tred_factor2).
+(ProveStable Tred_factor2 Zred_factor2).
+Save.
+
+Definition Tred_factor3 [t: term] :=
+ Cases t of
+ (Tplus (Tmult x (Tint k)) y) =>
+ Case (eq_term x y) of
+ (Tmult x (Tint `1+k`))
+ t
+ end
+ | _ => t
+ end.
+
+Theorem Tred_factor3_stable : (term_stable Tred_factor3).
+
+(ProveStable Tred_factor3 Zred_factor3).
+Save.
+
+
+Definition Tred_factor4 [t: term] :=
+ Cases t of
+ (Tplus (Tmult x (Tint k1)) (Tmult y (Tint k2))) =>
+ Case (eq_term x y) of
+ (Tmult x (Tint `k1+k2`))
+ t
+ end
+ | _ => t
+ end.
+
+Theorem Tred_factor4_stable : (term_stable Tred_factor4).
+
+(ProveStable Tred_factor4 Zred_factor4).
+Save.
+
+Definition Tred_factor6 [t: term] := (Tplus t (Tint `0`)).
+
+Theorem Tred_factor6_stable : (term_stable Tred_factor6).
+
+(ProveStable Tred_factor6 Zred_factor6).
+Save.
+
+Transparent Zplus.
+
+Definition Tminus_def [t:term] :=
+ Cases t of
+ (Tminus x y) => (Tplus x (Topp y))
+ | _ => t
+ end.
+
+Theorem Tminus_def_stable : (term_stable Tminus_def).
+
+(* Le théorème ne sert à rien. Le but est prouvé avant. *)
+(ProveStable Tminus_def False).
+Save.
+
+(* \subsection{Fonctions de réécriture complexes} *)
+
+(* \subsubsection{Fonction de réduction} *)
+(* Cette fonction réduit un terme dont la forme normale est un entier. Il
+ suffit pour cela d'échanger le constructeur [Tint] avec les opérateurs
+ réifiés. La réduction est ``gratuite''. *)
+
+Fixpoint reduce [t:term] : term :=
+ Cases t of
+ (Tplus x y) =>
+ Cases (reduce x) of
+ (Tint x') =>
+ Cases (reduce y) of
+ (Tint y') => (Tint (Zplus x' y'))
+ | y' => (Tplus (Tint x') y')
+ end
+ | x' => (Tplus x' (reduce y))
+ end
+ | (Tmult x y) =>
+ Cases (reduce x) of
+ (Tint x') =>
+ Cases (reduce y) of
+ (Tint y') => (Tint (Zmult x' y'))
+ | y' => (Tmult (Tint x') y')
+ end
+ | x' => (Tmult x' (reduce y))
+ end
+ | (Tminus x y) =>
+ Cases (reduce x) of
+ (Tint x') =>
+ Cases (reduce y) of
+ (Tint y') => (Tint (Zminus x' y'))
+ | y' => (Tminus (Tint x') y')
+ end
+ | x' => (Tminus x' (reduce y))
+ end
+ | (Topp x) =>
+ Cases (reduce x) of
+ (Tint x') => (Tint (Zopp x'))
+ | x' => (Topp x')
+ end
+ | _ => t
+ end.
+
+Theorem reduce_stable : (term_stable reduce).
+
+Unfold term_stable; Intros e t; Elim t; Auto;
+Try (Intros t0 H0 t1 H1; Simpl; Rewrite H0; Rewrite H1; (
+ Case (reduce t0); [
+ Intro z0; Case (reduce t1); Intros; Auto
+ | Intros; Auto
+ | Intros; Auto
+ | Intros; Auto
+ | Intros; Auto
+ | Intros; Auto ]));
+Intros t0 H0; Simpl; Rewrite H0; Case (reduce t0); Intros; Auto.
+Save.
+
+(* \subsubsection{Fusions}
+ \paragraph{Fusion de deux équations} *)
+(* On donne une somme de deux équations qui sont supposées normalisées.
+ Cette fonction prend une trace de fusion en argument et transforme
+ le terme en une équation normalisée. C'est une version très simplifiée
+ du moteur de réécriture [rewrite]. *)
+
+Fixpoint fusion [trace : (list t_fusion)] : term -> term := [t: term]
+ Cases trace of
+ nil => (reduce t)
+ | (cons step trace') =>
+ Cases step of
+ | F_equal =>
+ (apply_right (fusion trace') (T_OMEGA10 t))
+ | F_cancel =>
+ (fusion trace' (Tred_factor5 (T_OMEGA10 t)))
+ | F_left =>
+ (apply_right (fusion trace') (T_OMEGA11 t))
+ | F_right =>
+ (apply_right (fusion trace') (T_OMEGA12 t))
+ end
+ end.
+
+Theorem fusion_stable : (t : (list t_fusion)) (term_stable (fusion t)).
+
+Induction t; Simpl; [
+ Exact reduce_stable
+| Intros stp l H; Case stp; [
+ Apply compose_term_stable;
+ [ Apply apply_right_stable; Assumption | Exact T_OMEGA10_stable ]
+ | Unfold term_stable; Intros e t1; Rewrite T_OMEGA10_stable;
+ Rewrite Tred_factor5_stable; Apply H
+ | Apply compose_term_stable;
+ [ Apply apply_right_stable; Assumption | Exact T_OMEGA11_stable ]
+ | Apply compose_term_stable;
+ [ Apply apply_right_stable; Assumption | Exact T_OMEGA12_stable ]]].
+
+Save.
+
+(* \paragraph{Fusion de deux équations dont une sans coefficient} *)
+
+Definition fusion_right [trace : (list t_fusion)] : term -> term := [t: term]
+ Cases trace of
+ nil => (reduce t) (* Il faut mettre un compute *)
+ | (cons step trace') =>
+ Cases step of
+ | F_equal =>
+ (apply_right (fusion trace') (T_OMEGA15 t))
+ | F_cancel =>
+ (fusion trace' (Tred_factor5 (T_OMEGA15 t)))
+ | F_left =>
+ (apply_right (fusion trace') (Tplus_assoc_r t))
+ | F_right =>
+ (apply_right (fusion trace') (T_OMEGA12 t))
+ end
+ end.
+
+(* \paragraph{Fusion avec anihilation} *)
+(* Normalement le résultat est une constante *)
+
+Fixpoint fusion_cancel [trace:nat] : term -> term := [t:term]
+ Cases trace of
+ O => (reduce t)
+ | (S trace') => (fusion_cancel trace' (T_OMEGA13 t))
+ end.
+
+Theorem fusion_cancel_stable : (t:nat) (term_stable (fusion_cancel t)).
+
+Unfold term_stable fusion_cancel; Intros trace e; Elim trace; [
+ Exact (reduce_stable e)
+| Intros n H t; Elim H; Exact (T_OMEGA13_stable e t) ].
+Save.
+
+(* \subsubsection{Opérations afines sur une équation} *)
+(* \paragraph{Multiplication scalaire et somme d'une constante} *)
+
+Fixpoint scalar_norm_add [trace:nat] : term -> term := [t: term]
+ Cases trace of
+ O => (reduce t)
+ | (S trace') => (apply_right (scalar_norm_add trace') (T_OMEGA11 t))
+ end.
+
+Theorem scalar_norm_add_stable : (t:nat) (term_stable (scalar_norm_add t)).
+
+Unfold term_stable scalar_norm_add; Intros trace; Elim trace; [
+ Exact reduce_stable
+| Intros n H e t; Elim apply_right_stable;
+ [ Exact (T_OMEGA11_stable e t) | Exact H ]].
+Save.
+
+(* \paragraph{Multiplication scalaire} *)
+Fixpoint scalar_norm [trace:nat] : term -> term := [t: term]
+ Cases trace of
+ O => (reduce t)
+ | (S trace') => (apply_right (scalar_norm trace') (T_OMEGA16 t))
+ end.
+
+Theorem scalar_norm_stable : (t:nat) (term_stable (scalar_norm t)).
+
+Unfold term_stable scalar_norm; Intros trace; Elim trace; [
+ Exact reduce_stable
+| Intros n H e t; Elim apply_right_stable;
+ [ Exact (T_OMEGA16_stable e t) | Exact H ]].
+Save.
+
+(* \paragraph{Somme d'une constante} *)
+Fixpoint add_norm [trace:nat] : term -> term := [t: term]
+ Cases trace of
+ O => (reduce t)
+ | (S trace') => (apply_right (add_norm trace') (Tplus_assoc_r t))
+ end.
+
+Theorem add_norm_stable : (t:nat) (term_stable (add_norm t)).
+
+Unfold term_stable add_norm; Intros trace; Elim trace; [
+ Exact reduce_stable
+| Intros n H e t; Elim apply_right_stable;
+ [ Exact (Tplus_assoc_r_stable e t) | Exact H ]].
+Save.
+
+(* \subsection{La fonction de normalisation des termes (moteur de réécriture)} *)
+
+
+Fixpoint rewrite [s: step] : term -> term :=
+ Cases s of
+ | (C_DO_BOTH s1 s2) => (apply_both (rewrite s1) (rewrite s2))
+ | (C_LEFT s) => (apply_left (rewrite s))
+ | (C_RIGHT s) => (apply_right (rewrite s))
+ | (C_SEQ s1 s2) => [t: term] (rewrite s2 (rewrite s1 t))
+ | C_NOP => [t:term] t
+ | C_OPP_PLUS => Topp_plus
+ | C_OPP_OPP => Topp_opp
+ | C_OPP_MULT_R => Topp_mult_r
+ | C_OPP_ONE => Topp_one
+ | C_REDUCE => reduce
+ | C_MULT_PLUS_DISTR => Tmult_plus_distr
+ | C_MULT_OPP_LEFT => Tmult_opp_left
+ | C_MULT_ASSOC_R => Tmult_assoc_r
+ | C_PLUS_ASSOC_R => Tplus_assoc_r
+ | C_PLUS_ASSOC_L => Tplus_assoc_l
+ | C_PLUS_PERMUTE => Tplus_permute
+ | C_PLUS_SYM => Tplus_sym
+ | C_RED0 => Tred_factor0
+ | C_RED1 => Tred_factor1
+ | C_RED2 => Tred_factor2
+ | C_RED3 => Tred_factor3
+ | C_RED4 => Tred_factor4
+ | C_RED5 => Tred_factor5
+ | C_RED6 => Tred_factor6
+ | C_MULT_ASSOC_REDUCED => Tmult_assoc_reduced
+ | C_MINUS => Tminus_def
+ | C_MULT_SYM => Tmult_sym
+ end.
+
+Theorem rewrite_stable : (s:step) (term_stable (rewrite s)).
+
+Induction s; Simpl; [
+ Intros; Apply apply_both_stable; Auto
+| Intros; Apply apply_left_stable; Auto
+| Intros; Apply apply_right_stable; Auto
+| Unfold term_stable; Intros; Elim H0; Apply H
+| Unfold term_stable; Auto
+| Exact Topp_plus_stable
+| Exact Topp_opp_stable
+| Exact Topp_mult_r_stable
+| Exact Topp_one_stable
+| Exact reduce_stable
+| Exact Tmult_plus_distr_stable
+| Exact Tmult_opp_left_stable
+| Exact Tmult_assoc_r_stable
+| Exact Tplus_assoc_r_stable
+| Exact Tplus_assoc_l_stable
+| Exact Tplus_permute_stable
+| Exact Tplus_sym_stable
+| Exact Tred_factor0_stable
+| Exact Tred_factor1_stable
+| Exact Tred_factor2_stable
+| Exact Tred_factor3_stable
+| Exact Tred_factor4_stable
+| Exact Tred_factor5_stable
+| Exact Tred_factor6_stable
+| Exact Tmult_assoc_reduced_stable
+| Exact Tminus_def_stable
+| Exact Tmult_sym_stable ].
+Save.
+
+(* \subsection{tactiques de résolution d'un but omega normalisé}
+ Trace de la procédure
+\subsubsection{Tactiques générant une contradiction}
+\paragraph{[O_CONSTANT_NOT_NUL]} *)
+
+Definition constant_not_nul [i:nat; h: hyps] :=
+ Cases (nth_hyps i h) of
+ (EqTerm (Tint ZERO) (Tint n)) =>
+ Case (eq_Z n ZERO) of
+ h
+ absurd
+ end
+ | _ => h
+ end.
+
+Theorem constant_not_nul_valid :
+ (i:nat) (valid_hyps (constant_not_nul i)).
+
+Unfold valid_hyps constant_not_nul; Intros;
+Generalize (nth_valid ep e i lp); Simplify; Simpl; (Elim_eq_Z z0 ZERO); Auto;
+Simpl; Intros H1 H2; Elim H1; Symmetry; Auto.
+Save.
+
+(* \paragraph{[O_CONSTANT_NEG]} *)
+
+Definition constant_neg [i:nat; h: hyps] :=
+ Cases (nth_hyps i h) of
+ (LeqTerm (Tint ZERO) (Tint (NEG n))) => absurd
+ | _ => h
+ end.
+
+Theorem constant_neg_valid : (i:nat) (valid_hyps (constant_neg i)).
+
+Unfold valid_hyps constant_neg; Intros;
+Generalize (nth_valid ep e i lp); Simplify; Simpl; Unfold Zle; Simpl;
+Intros H1; Elim H1; [ Assumption | Trivial ].
+Save.
+
+(* \paragraph{[NOT_EXACT_DIVIDE]} *)
+Definition not_exact_divide [k1,k2:Z; body:term; t:nat; i : nat; l:hyps] :=
+ Cases (nth_hyps i l) of
+ (EqTerm (Tint ZERO) b) =>
+ Case (eq_term
+ (scalar_norm_add t (Tplus (Tmult body (Tint k1)) (Tint k2))) b) of
+ Cases (Zcompare k2 ZERO) of
+ SUPERIEUR =>
+ Cases (Zcompare k1 k2) of
+ SUPERIEUR => absurd
+ | _ => l
+ end
+ | _ => l
+ end
+ l
+ end
+ | _ => l
+ end.
+
+Theorem not_exact_divide_valid : (k1,k2:Z; body:term; t:nat; i:nat)
+ (valid_hyps (not_exact_divide k1 k2 body t i)).
+
+Unfold valid_hyps not_exact_divide; Intros; Generalize (nth_valid ep e i lp);
+Simplify;
+(Elim_eq_term '(scalar_norm_add t (Tplus (Tmult body (Tint k1)) (Tint k2)))
+ 't1); Auto;
+Simplify;
+Intro H2; Elim H2; Simpl; Elim (scalar_norm_add_stable t e); Simpl;
+Intro H4; Absurd `(interp_term e body)*k1+k2 = 0`; [
+ Apply OMEGA4; Assumption | Symmetry; Auto ].
+
+Save.
+
+(* \paragraph{[O_CONTRADICTION]} *)
+
+Definition contradiction [t: nat; i,j:nat;l:hyps] :=
+ Cases (nth_hyps i l) of
+ (LeqTerm (Tint ZERO) b1) =>
+ Cases (nth_hyps j l) of
+ (LeqTerm (Tint ZERO) b2) =>
+ Cases (fusion_cancel t (Tplus b1 b2)) of
+ (Tint k) =>
+ Cases (Zcompare ZERO k) of
+ SUPERIEUR => absurd
+ | _ => l
+ end
+ | _ => l
+ end
+ | _ => l
+ end
+ | _ => l
+ end.
+
+Theorem contradiction_valid : (t,i,j: nat) (valid_hyps (contradiction t i j)).
+
+Unfold valid_hyps contradiction; Intros t i j ep e l H;
+Generalize (nth_valid ? ? i ? H); Generalize (nth_valid ? ? j ? H);
+Case (nth_hyps i l); Auto; Intros t1 t2; Case t1; Auto; Intros z; Case z; Auto;
+Case (nth_hyps j l); Auto; Intros t3 t4; Case t3; Auto; Intros z'; Case z';
+Auto; Simpl; Intros H1 H2;
+Generalize (refl_equal Z (interp_term e (fusion_cancel t (Tplus t2 t4))));
+Pattern 2 3 (fusion_cancel t (Tplus t2 t4));
+Case (fusion_cancel t (Tplus t2 t4));
+Simpl; Auto; Intro k; Elim (fusion_cancel_stable t);
+Simpl; Intro E; Generalize (OMEGA2 ? ? H2 H1); Rewrite E; Case k;
+Auto;Unfold Zle; Simpl; Intros p H3; Elim H3; Auto.
+
+Save.
+
+(* \paragraph{[O_NEGATE_CONTRADICT]} *)
+
+Definition negate_contradict [i1,i2:nat; h:hyps]:=
+ Cases (nth_hyps i1 h) of
+ (EqTerm (Tint ZERO) b1) =>
+ Cases (nth_hyps i2 h) of
+ (NeqTerm (Tint ZERO) b2) =>
+ Cases (eq_term b1 b2) of
+ true => absurd
+ | false => h
+ end
+ | _ => h
+ end
+ | (NeqTerm (Tint ZERO) b1) =>
+ Cases (nth_hyps i2 h) of
+ (EqTerm (Tint ZERO) b2) =>
+ Cases (eq_term b1 b2) of
+ true => absurd
+ | false => h
+ end
+ | _ => h
+ end
+ | _ => h
+ end.
+
+Definition negate_contradict_inv [t:nat; i1,i2:nat; h:hyps]:=
+ Cases (nth_hyps i1 h) of
+ (EqTerm (Tint ZERO) b1) =>
+ Cases (nth_hyps i2 h) of
+ (NeqTerm (Tint ZERO) b2) =>
+ Cases (eq_term b1 (scalar_norm t (Tmult b2 (Tint `-1`)))) of
+ true => absurd
+ | false => h
+ end
+ | _ => h
+ end
+ | (NeqTerm (Tint ZERO) b1) =>
+ Cases (nth_hyps i2 h) of
+ (EqTerm (Tint ZERO) b2) =>
+ Cases (eq_term b1 (scalar_norm t (Tmult b2 (Tint `-1`)))) of
+ true => absurd
+ | false => h
+ end
+ | _ => h
+ end
+ | _ => h
+ end.
+
+Theorem negate_contradict_valid :
+ (i,j:nat) (valid_hyps (negate_contradict i j)).
+
+Unfold valid_hyps negate_contradict; Intros i j ep e l H;
+Generalize (nth_valid ? ? i ? H); Generalize (nth_valid ? ? j ? H);
+Case (nth_hyps i l); Auto; Intros t1 t2; Case t1; Auto; Intros z; Case z; Auto;
+Case (nth_hyps j l); Auto; Intros t3 t4; Case t3; Auto; Intros z'; Case z';
+Auto; Simpl; Intros H1 H2; [
+ (Elim_eq_term t2 t4); Intro H3; [ Elim H1; Elim H3; Assumption | Assumption ]
+| (Elim_eq_term t2 t4); Intro H3;
+ [ Elim H2; Rewrite H3; Assumption | Assumption ]].
+
+Save.
+
+Theorem negate_contradict_inv_valid :
+ (t,i,j:nat) (valid_hyps (negate_contradict_inv t i j)).
+
+
+Unfold valid_hyps negate_contradict_inv; Intros t i j ep e l H;
+Generalize (nth_valid ? ? i ? H); Generalize (nth_valid ? ? j ? H);
+Case (nth_hyps i l); Auto; Intros t1 t2; Case t1; Auto; Intros z; Case z; Auto;
+Case (nth_hyps j l); Auto; Intros t3 t4; Case t3; Auto; Intros z'; Case z';
+Auto; Simpl; Intros H1 H2;
+(Pattern (eq_term t2 (scalar_norm t (Tmult t4 (Tint (NEG xH))))); Apply bool_ind2; Intro Aux; [
+ Generalize (eq_term_true t2 (scalar_norm t (Tmult t4 (Tint (NEG xH)))) Aux);
+ Clear Aux
+| Generalize (eq_term_false t2 (scalar_norm t (Tmult t4 (Tint (NEG xH)))) Aux);
+ Clear Aux ]); [
+ Intro H3; Elim H1; Generalize H2; Rewrite H3;
+ Rewrite <- (scalar_norm_stable t e); Simpl; Elim (interp_term e t4) ;
+ Simpl; Auto; Intros p H4; Discriminate H4
+ | Auto
+ | Intro H3; Elim H2; Rewrite H3; Elim (scalar_norm_stable t e); Simpl;
+ Elim H1; Simpl; Trivial
+ | Auto ].
+
+Save.
+
+(* \subsubsection{Tactiques générant une nouvelle équation} *)
+(* \paragraph{[O_SUM]}
+ C'est une oper2 valide mais elle traite plusieurs cas à la fois (suivant
+ les opérateurs de comparaison des deux arguments) d'où une
+ preuve un peu compliquée. On utilise quelques lemmes qui sont des
+ généralisations des théorèmes utilisés par OMEGA. *)
+
+Definition sum [k1,k2: Z; trace: (list t_fusion); prop1,prop2:proposition]:=
+ Cases prop1 of
+ (EqTerm (Tint ZERO) b1) =>
+ Cases prop2 of
+ (EqTerm (Tint ZERO) b2) =>
+ (EqTerm
+ (Tint ZERO)
+ (fusion trace
+ (Tplus (Tmult b1 (Tint k1)) (Tmult b2 (Tint k2)))))
+ | (LeqTerm (Tint ZERO) b2) =>
+ Cases (Zcompare k2 ZERO) of
+ SUPERIEUR =>
+ (LeqTerm
+ (Tint ZERO)
+ (fusion trace
+ (Tplus (Tmult b1 (Tint k1)) (Tmult b2 (Tint k2)))))
+ | _ => TrueTerm
+ end
+ | _ => TrueTerm
+ end
+ | (LeqTerm (Tint ZERO) b1) =>
+ Cases (Zcompare k1 ZERO) of
+ SUPERIEUR =>
+ Cases prop2 of
+ (EqTerm (Tint ZERO) b2) =>
+ (LeqTerm
+ (Tint ZERO)
+ (fusion trace
+ (Tplus (Tmult b1 (Tint k1)) (Tmult b2 (Tint k2)))))
+ | (LeqTerm (Tint ZERO) b2) =>
+ Cases (Zcompare k2 ZERO) of
+ SUPERIEUR =>
+ (LeqTerm
+ (Tint ZERO)
+ (fusion trace
+ (Tplus (Tmult b1 (Tint k1))
+ (Tmult b2 (Tint k2)))))
+ | _ => TrueTerm
+ end
+ | _ => TrueTerm
+ end
+ | _ => TrueTerm
+ end
+ | (NeqTerm (Tint ZERO) b1) =>
+ Cases prop2 of
+ (EqTerm (Tint ZERO) b2) =>
+ Case (eq_Z k1 ZERO) of
+ TrueTerm
+ (NeqTerm
+ (Tint ZERO)
+ (fusion trace
+ (Tplus (Tmult b1 (Tint k1)) (Tmult b2 (Tint k2)))))
+ end
+ | _ => TrueTerm
+ end
+ | _ => TrueTerm
+ end.
+
+Theorem sum1 :
+ (a,b,c,d:Z) (`0 = a`) -> (`0 = b`) -> (`0 = a*c + b*d`).
+
+Intros; Elim H; Elim H0; Simpl; Auto.
+Save.
+
+Theorem sum2 :
+ (a,b,c,d:Z) (`0 <= d`) -> (`0 = a`) -> (`0 <= b`) ->(`0 <= a*c + b*d`).
+
+Intros; Elim H0; Simpl; Generalize H H1; Case b; Case d;
+Unfold Zle; Simpl; Auto.
+Save.
+
+Theorem sum3 :
+ (a,b,c,d:Z) (`0 <= c`) -> (`0 <= d`) -> (`0 <= a`) -> (`0 <= b`) ->(`0 <= a*c + b*d`).
+
+Intros a b c d; Case a; Case b; Case c; Case d; Unfold Zle; Simpl; Auto.
+Save.
+
+Theorem sum4 : (k:Z) (Zcompare k `0`)=SUPERIEUR -> (`0 <= k`).
+
+Intro; Case k; Unfold Zle; Simpl; Auto; Intros; Discriminate.
+Save.
+
+Theorem sum5 :
+ (a,b,c,d:Z) (`c <> 0`) -> (`0 <> a`) -> (`0 = b`) -> (`0 <> a*c + b*d`).
+
+Intros a b c d H1 H2 H3; Elim H3; Simpl; Rewrite Zplus_sym;
+Simpl; Generalize H1 H2; Case a; Case c; Simpl; Intros; Try Discriminate;
+Assumption.
+Save.
+
+
+Theorem sum_valid : (k1,k2:Z; t:(list t_fusion)) (valid2 (sum k1 k2 t)).
+
+Unfold valid2; Intros k1 k2 t ep e p1 p2; Unfold sum; Simplify; Simpl; Auto;
+Try (Elim (fusion_stable t)); Simpl; Intros; [
+ Apply sum1; Assumption
+| Apply sum2; Try Assumption; Apply sum4; Assumption
+| Rewrite Zplus_sym; Apply sum2; Try Assumption; Apply sum4; Assumption
+| Apply sum3; Try Assumption; Apply sum4; Assumption
+| (Elim_eq_Z k1 ZERO); Simpl; Auto; Elim (fusion_stable t); Simpl; Intros;
+ Unfold Zne; Apply sum5; Assumption].
+Save.
+
+(* \paragraph{[O_EXACT_DIVIDE]}
+ c'est une oper1 valide mais on préfère une substitution a ce point la *)
+
+Definition exact_divide [k:Z; body:term; t: nat; prop:proposition] :=
+ Cases prop of
+ (EqTerm (Tint ZERO) b) =>
+ Case (eq_term (scalar_norm t (Tmult body (Tint k))) b) of
+ Case (eq_Z k ZERO) of
+ TrueTerm
+ (EqTerm (Tint ZERO) body)
+ end
+ TrueTerm
+ end
+ | _ => TrueTerm
+ end.
+
+Theorem exact_divide_valid :
+ (k:Z) (t:term) (n:nat) (valid1 (exact_divide k t n)).
+
+
+Unfold valid1 exact_divide; Intros k1 k2 t ep e p1; Simplify;Simpl; Auto;
+(Elim_eq_term '(scalar_norm t (Tmult k2 (Tint k1))) 't1); Simpl; Auto;
+(Elim_eq_Z 'k1 'ZERO); Simpl; Auto; Intros H1 H2; Elim H2;
+Elim scalar_norm_stable; Simpl; Generalize H1; Case (interp_term e k2);
+Try Trivial; (Case k1; Simpl; [
+ Intros; Absurd `0 = 0`; Assumption
+| Intros p2 p3 H3 H4; Discriminate H4
+| Intros p2 p3 H3 H4; Discriminate H4 ]).
+
+Save.
+
+
+
+(* \paragraph{[O_DIV_APPROX]}
+ La preuve reprend le schéma de la précédente mais on
+ est sur une opération de type valid1 et non sur une opération terminale. *)
+
+Definition divide_and_approx [k1,k2:Z; body:term; t:nat; prop:proposition] :=
+ Cases prop of
+ (LeqTerm (Tint ZERO) b) =>
+ Case (eq_term
+ (scalar_norm_add t (Tplus (Tmult body (Tint k1)) (Tint k2))) b) of
+ Cases (Zcompare k1 ZERO) of
+ SUPERIEUR =>
+ Cases (Zcompare k1 k2) of
+ SUPERIEUR =>(LeqTerm (Tint ZERO) body)
+ | _ => prop
+ end
+ | _ => prop
+ end
+ prop
+ end
+ | _ => prop
+ end.
+
+Theorem divide_and_approx_valid : (k1,k2:Z; body:term; t:nat)
+ (valid1 (divide_and_approx k1 k2 body t)).
+
+Unfold valid1 divide_and_approx; Intros k1 k2 body t ep e p1;Simplify;
+(Elim_eq_term '(scalar_norm_add t (Tplus (Tmult body (Tint k1)) (Tint k2))) 't1); Simplify; Auto; Intro E; Elim E; Simpl;
+Elim (scalar_norm_add_stable t e); Simpl; Intro H1;
+Apply Zmult_le_approx with 3 := H1; Assumption.
+Save.
+
+(* \paragraph{[MERGE_EQ]} *)
+
+Definition merge_eq [t: nat; prop1, prop2: proposition] :=
+ Cases prop1 of
+ (LeqTerm (Tint ZERO) b1) =>
+ Cases prop2 of
+ (LeqTerm (Tint ZERO) b2) =>
+ Case (eq_term b1 (scalar_norm t (Tmult b2 (Tint `-1`)))) of
+ (EqTerm (Tint ZERO) b1)
+ TrueTerm
+ end
+ | _ => TrueTerm
+ end
+ | _ => TrueTerm
+ end.
+
+Theorem merge_eq_valid : (n:nat) (valid2 (merge_eq n)).
+
+Unfold valid2 merge_eq; Intros n ep e p1 p2; Simplify; Simpl; Auto;
+Elim (scalar_norm_stable n e); Simpl; Intros; Symmetry;
+Apply OMEGA8 with 2 := H0; [ Assumption | Elim Zopp_one; Trivial ].
+Save.
+
+
+
+(* \paragraph{[O_CONSTANT_NUL]} *)
+
+Definition constant_nul [i:nat; h: hyps] :=
+ Cases (nth_hyps i h) of
+ (NeqTerm (Tint ZERO) (Tint ZERO)) => absurd
+ | _ => h
+ end.
+
+Theorem constant_nul_valid :
+ (i:nat) (valid_hyps (constant_nul i)).
+
+Unfold valid_hyps constant_nul; Intros; Generalize (nth_valid ep e i lp);
+Simplify; Simpl; Unfold Zne; Intro H1; Absurd `0=0`; Auto.
+Save.
+
+(* \paragraph{[O_STATE]} *)
+
+Definition state [m:Z;s:step; prop1,prop2:proposition] :=
+ Cases prop1 of
+ (EqTerm (Tint ZERO) b1) =>
+ Cases prop2 of
+ (EqTerm (Tint ZERO) (Tplus b2 (Topp b3))) =>
+ (EqTerm (Tint ZERO) (rewrite s (Tplus b1 (Tmult (Tplus (Topp b3) b2) (Tint m)))))
+ | _ => TrueTerm
+ end
+ | _ => TrueTerm
+ end.
+
+Theorem state_valid : (m:Z; s:step) (valid2 (state m s)).
+
+Unfold valid2; Intros m s ep e p1 p2; Unfold state; Simplify; Simpl;Auto;
+Elim (rewrite_stable s e); Simpl; Intros H1 H2; Elim H1;
+Rewrite (Zplus_sym `-(interp_term e t5)` `(interp_term e t3)`);
+Elim H2; Simpl; Reflexivity.
+
+Save.
+
+(* \subsubsection{Tactiques générant plusieurs but}
+ \paragraph{[O_SPLIT_INEQ]}
+ La seule pour le moment (tant que la normalisation n'est pas réfléchie). *)
+
+Definition split_ineq [i,t: nat; f1,f2:hyps -> lhyps; l:hyps] :=
+ Cases (nth_hyps i l) of
+ (NeqTerm (Tint ZERO) b1) =>
+ (app (f1 (cons (LeqTerm (Tint ZERO) (add_norm t (Tplus b1 (Tint `-1`)))) l))
+ (f2 (cons (LeqTerm (Tint ZERO)
+ (scalar_norm_add t
+ (Tplus (Tmult b1 (Tint `-1`)) (Tint `-1`))))
+ l)))
+ | _ => (cons l (nil ?))
+ end.
+
+Theorem split_ineq_valid :
+ (i,t: nat; f1,f2: hyps -> lhyps)
+ (valid_list_hyps f1) ->(valid_list_hyps f2) ->
+ (valid_list_hyps (split_ineq i t f1 f2)).
+
+Unfold valid_list_hyps split_ineq; Intros i t f1 f2 H1 H2 ep e lp H;
+Generalize (nth_valid ? ? i ? H);
+Case (nth_hyps i lp); Simpl; Auto; Intros t1 t2; Case t1; Simpl; Auto;
+Intros z; Case z; Simpl; Auto;
+Intro H3; Apply append_valid;Elim (OMEGA19 (interp_term e t2)) ;[
+ Intro H4; Left; Apply H1; Simpl; Elim (add_norm_stable t); Simpl; Auto
+| Intro H4; Right; Apply H2; Simpl; Elim (scalar_norm_add_stable t);
+ Simpl; Auto
+| Generalize H3; Unfold Zne not; Intros E1 E2; Apply E1; Symmetry; Trivial ].
+Save.
+
+
+(* \subsection{La fonction de rejeu de la trace} *)
+
+Fixpoint execute_omega [t: t_omega] : hyps -> lhyps :=
+ [l : hyps] Cases t of
+ | (O_CONSTANT_NOT_NUL n) => (singleton (constant_not_nul n l))
+ | (O_CONSTANT_NEG n) => (singleton (constant_neg n l))
+ | (O_DIV_APPROX k1 k2 body t cont n) =>
+ (execute_omega cont
+ (apply_oper_1 n (divide_and_approx k1 k2 body t) l))
+ | (O_NOT_EXACT_DIVIDE k1 k2 body t i) =>
+ (singleton (not_exact_divide k1 k2 body t i l))
+ | (O_EXACT_DIVIDE k body t cont n) =>
+ (execute_omega cont (apply_oper_1 n (exact_divide k body t) l))
+ | (O_SUM k1 i1 k2 i2 t cont) =>
+ (execute_omega cont (apply_oper_2 i1 i2 (sum k1 k2 t) l))
+ | (O_CONTRADICTION t i j) =>
+ (singleton (contradiction t i j l))
+ | (O_MERGE_EQ t i1 i2 cont) =>
+ (execute_omega cont (apply_oper_2 i1 i2 (merge_eq t) l))
+ | (O_SPLIT_INEQ t i cont1 cont2) =>
+ (split_ineq i t (execute_omega cont1) (execute_omega cont2) l)
+ | (O_CONSTANT_NUL i) => (singleton (constant_nul i l))
+ | (O_NEGATE_CONTRADICT i j) => (singleton (negate_contradict i j l))
+ | (O_NEGATE_CONTRADICT_INV t i j) => (singleton (negate_contradict_inv t i j l))
+ | (O_STATE m s i1 i2 cont) =>
+ (execute_omega cont (apply_oper_2 i1 i2 (state m s) l))
+ end.
+
+Theorem omega_valid : (t: t_omega) (valid_list_hyps (execute_omega t)).
+
+Induction t; Simpl; [
+ Unfold valid_list_hyps; Simpl; Intros; Left;
+ Apply (constant_not_nul_valid n ep e lp H)
+| Unfold valid_list_hyps; Simpl; Intros; Left;
+ Apply (constant_neg_valid n ep e lp H)
+| Unfold valid_list_hyps valid_hyps; Intros k1 k2 body n t' Ht' m ep e lp H;
+ Apply Ht';
+ Apply (apply_oper_1_valid m (divide_and_approx k1 k2 body n)
+ (divide_and_approx_valid k1 k2 body n) ep e lp H)
+| Unfold valid_list_hyps; Simpl; Intros; Left;
+ Apply (not_exact_divide_valid z z0 t0 n n0 ep e lp H)
+| Unfold valid_list_hyps valid_hyps; Intros k body n t' Ht' m ep e lp H;
+ Apply Ht';
+ Apply (apply_oper_1_valid m (exact_divide k body n)
+ (exact_divide_valid k body n) ep e lp H)
+| Unfold valid_list_hyps valid_hyps; Intros k1 i1 k2 i2 trace t' Ht' ep e lp H;
+ Apply Ht';
+ Apply (apply_oper_2_valid i1 i2 (sum k1 k2 trace)
+ (sum_valid k1 k2 trace) ep e lp H)
+| Unfold valid_list_hyps; Simpl; Intros; Left;
+ Apply (contradiction_valid n n0 n1 ep e lp H)
+| Unfold valid_list_hyps valid_hyps; Intros trace i1 i2 t' Ht' ep e lp H;
+ Apply Ht';
+ Apply (apply_oper_2_valid i1 i2 (merge_eq trace)
+ (merge_eq_valid trace) ep e lp H)
+| Intros t' i k1 H1 k2 H2; Unfold valid_list_hyps; Simpl; Intros ep e lp H;
+ Apply (split_ineq_valid i t' (execute_omega k1) (execute_omega k2)
+ H1 H2 ep e lp H)
+| Unfold valid_list_hyps; Simpl; Intros i ep e lp H; Left;
+ Apply (constant_nul_valid i ep e lp H)
+| Unfold valid_list_hyps; Simpl; Intros i j ep e lp H; Left;
+ Apply (negate_contradict_valid i j ep e lp H)
+| Unfold valid_list_hyps; Simpl; Intros n i j ep e lp H; Left;
+ Apply (negate_contradict_inv_valid n i j ep e lp H)
+| Unfold valid_list_hyps valid_hyps; Intros m s i1 i2 t' Ht' ep e lp H; Apply Ht';
+ Apply (apply_oper_2_valid i1 i2 (state m s) (state_valid m s) ep e lp H)
+].
+Save.
+
+
+(* \subsection{Les opérations globales sur le but}
+ \subsubsection{Normalisation} *)
+
+Definition move_right [s: step; p:proposition] :=
+ Cases p of
+ (EqTerm t1 t2) => (EqTerm (Tint ZERO) (rewrite s (Tplus t1 (Topp t2))))
+ | (LeqTerm t1 t2) => (LeqTerm (Tint ZERO) (rewrite s (Tplus t2 (Topp t1))))
+ | (GeqTerm t1 t2) => (LeqTerm (Tint ZERO) (rewrite s (Tplus t1 (Topp t2))))
+ | (LtTerm t1 t2) =>
+ (LeqTerm (Tint ZERO)
+ (rewrite s (Tplus (Tplus t2 (Tint `-1`)) (Topp t1))))
+ | (GtTerm t1 t2) =>
+ (LeqTerm (Tint ZERO)
+ (rewrite s (Tplus (Tplus t1 (Tint `-1`)) (Topp t2))))
+ | (NeqTerm t1 t2) => (NeqTerm (Tint ZERO) (rewrite s (Tplus t1 (Topp t2))))
+ | p => p
+ end.
+
+Theorem Zne_left_2 : (x,y:Z)(Zne x y)->(Zne `0` `x+(-y)`).
+Unfold Zne not; Intros x y H1 H2; Apply H1; Apply (Zsimpl_plus_l `-y`);
+Rewrite Zplus_sym; Elim H2; Rewrite Zplus_inverse_l; Trivial.
+Save.
+
+Theorem move_right_valid : (s: step) (valid1 (move_right s)).
+
+Unfold valid1 move_right; Intros s ep e p; Simplify; Simpl;
+Elim (rewrite_stable s e); Simpl; [
+ Symmetry; Apply Zegal_left; Assumption
+| Intro; Apply Zle_left; Assumption
+| Intro; Apply Zge_left; Assumption
+| Intro; Apply Zgt_left; Assumption
+| Intro; Apply Zlt_left; Assumption
+| Intro; Apply Zne_left_2; Assumption
+].
+Save.
+
+Definition do_normalize [i:nat; s: step] := (apply_oper_1 i (move_right s)).
+
+Theorem do_normalize_valid : (i:nat; s:step) (valid_hyps (do_normalize i s)).
+
+Intros; Unfold do_normalize; Apply apply_oper_1_valid; Apply move_right_valid.
+Save.
+
+Fixpoint do_normalize_list [l:(list step)] : nat -> hyps -> hyps :=
+ [i:nat; h:hyps] Cases l of
+ (cons s l') => (do_normalize_list l' (S i) (do_normalize i s h))
+ | nil => h
+ end.
+
+Theorem do_normalize_list_valid :
+ (l:(list step); i:nat) (valid_hyps (do_normalize_list l i)).
+
+Induction l; Simpl; Unfold valid_hyps; [
+ Auto
+| Intros a l' Hl' i ep e lp H; Unfold valid_hyps in Hl'; Apply Hl';
+ Apply (do_normalize_valid i a ep e lp); Assumption ].
+Save.
+
+Theorem normalize_goal :
+ (s: (list step); ep: PropList; env : (list Z); l: hyps)
+ (interp_goal ep env (do_normalize_list s O l)) ->
+ (interp_goal ep env l).
+
+Intros; Apply valid_goal with 2:=H; Apply do_normalize_list_valid.
+Save.
+
+(* \subsubsection{Exécution de la trace} *)
+
+Theorem execute_goal :
+ (t : t_omega; ep: PropList; env : (list Z); l: hyps)
+ (interp_list_goal ep env (execute_omega t l)) -> (interp_goal ep env l).
+
+Intros; Apply (goal_valid (execute_omega t) (omega_valid t) ep env l H).
+Save.
+
+
+Theorem append_goal :
+ (ep: PropList; e: (list Z)) (l1,l2:lhyps)
+ (interp_list_goal ep e l1) /\ (interp_list_goal ep e l2) ->
+ (interp_list_goal ep e (app l1 l2)).
+
+Intros ep e; Induction l1; [
+ Simpl; Intros l2 (H1, H2); Assumption
+| Simpl; Intros h1 t1 HR l2 ((H1 , H2), H3) ; Split; Auto].
+
+Save.
+
+Require Decidable.
+
+(* A simple decidability checker : if the proposition belongs to the
+ simple grammar describe below then it is decidable. Proof is by
+ induction and uses well known theorem about arithmetic and propositional
+ calculus *)
+
+Fixpoint decidability [p:proposition] : bool :=
+ Cases p of
+ (EqTerm _ _) => true
+ | (LeqTerm _ _) => true
+ | (GeqTerm _ _) => true
+ | (GtTerm _ _) => true
+ | (LtTerm _ _) => true
+ | (NeqTerm _ _) => true
+ | (FalseTerm) => true
+ | (TrueTerm) => true
+ | (Tnot t) => (decidability t)
+ | (Tand t1 t2) => (andb (decidability t1) (decidability t2))
+ | (Timp t1 t2) => (andb (decidability t1) (decidability t2))
+ | (Tor t1 t2) => (andb (decidability t1) (decidability t2))
+ | (Tprop _) => false
+ end
+.
+
+Theorem decidable_correct :
+ (ep: PropList) (e: (list Z)) (p:proposition)
+ (decidability p)=true -> (decidable (interp_proposition ep e p)).
+
+Induction p; Simpl; Intros; [
+ Apply dec_eq
+| Apply dec_Zle
+| Left;Auto
+| Right; Unfold not; Auto
+| Apply dec_not; Auto
+| Apply dec_Zge
+| Apply dec_Zgt
+| Apply dec_Zlt
+| Apply dec_Zne
+| Apply dec_or; Elim andb_prop with 1 := H1; Auto
+| Apply dec_and; Elim andb_prop with 1 := H1; Auto
+| Apply dec_imp; Elim andb_prop with 1 := H1; Auto
+| Discriminate H].
+
+Save.
+
+(* An interpretation function for a complete goal with an explicit
+ conclusion. We use an intermediate fixpoint. *)
+
+Fixpoint interp_full_goal
+ [envp: PropList;env : (list Z); c : proposition; l: hyps] : Prop :=
+ Cases l of
+ nil => (interp_proposition envp env c)
+ | (cons p' l') =>
+ (interp_proposition envp env p') -> (interp_full_goal envp env c l')
+ end.
+
+Definition interp_full
+ [ep: PropList;e : (list Z); lc : (hyps * proposition)] : Prop :=
+ Cases lc of (l,c) => (interp_full_goal ep e c l) end.
+
+(* Relates the interpretation of a complete goal with the interpretation
+ of its hypothesis and conclusion *)
+
+Theorem interp_full_false :
+ (ep: PropList; e : (list Z); l: hyps; c : proposition)
+ ((interp_hyps ep e l) -> (interp_proposition ep e c)) ->
+ (interp_full ep e (l,c)).
+
+Induction l; Unfold interp_full; Simpl; [
+ Auto
+| Intros a l1 H1 c H2 H3; Apply H1; Auto].
+
+Save.
+
+(* Push the conclusion in the list of hypothesis using a double negation
+ If the decidability cannot be "proven", then just forget about the
+ conclusion (equivalent of replacing it with false) *)
+
+Definition to_contradict [lc : hyps * proposition] :=
+ Cases lc of
+ (l,c) => (if (decidability c) then (cons (Tnot c) l) else l)
+ end.
+
+(* The previous operation is valid in the sense that the new list of
+ hypothesis implies the original goal *)
+
+Theorem to_contradict_valid :
+ (ep: PropList; e : (list Z); lc: hyps * proposition)
+ (interp_goal ep e (to_contradict lc)) -> (interp_full ep e lc).
+
+Intros ep e lc; Case lc; Intros l c; Simpl; (Pattern (decidability c));
+Apply bool_ind2; [
+ Simpl; Intros H H1; Apply interp_full_false; Intros H2; Apply not_not; [
+ Apply decidable_correct; Assumption
+ | Unfold 1 not; Intro H3; Apply hyps_to_goal with 2:=H2; Auto]
+| Intros H1 H2; Apply interp_full_false; Intro H3; Elim hyps_to_goal with 1:= H2; Assumption ].
+Save.
+
+(* [map_cons x l] adds [x] at the head of each list in [l] (which is a list
+ of lists *)
+
+Fixpoint map_cons [A:Set; x:A; l:(list (list A))] : (list (list A)) :=
+ Cases l of
+ nil => (nil ?)
+ | (cons l ll) => (cons (cons x l) (map_cons A x ll))
+ end.
+
+(* This function breaks up a list of hypothesis in a list of simpler
+ list of hypothesis that together implie the original one. The goal
+ of all this is to transform the goal in a list of solvable problems.
+ Note that :
+ - we need a way to drive the analysis as some hypotheis may not
+ require a split.
+ - this procedure must be perfectly mimicked by the ML part otherwise
+ hypothesis will get desynchronised and this will be a mess.
+ *)
+
+Fixpoint destructure_hyps [nn: nat] : hyps -> lhyps :=
+ [ll:hyps]Cases nn of
+ O => (cons ll (nil ?))
+ | (S n) =>
+ Cases ll of
+ nil => (cons (nil ?) (nil ?))
+ | (cons (Tor p1 p2) l) =>
+ (app (destructure_hyps n (cons p1 l))
+ (destructure_hyps n (cons p2 l)))
+ | (cons (Tand p1 p2) l) =>
+ (destructure_hyps n (cons p1 (cons p2 l)))
+ | (cons (Timp p1 p2) l) =>
+ (if (decidability p1) then
+ (app (destructure_hyps n (cons (Tnot p1) l))
+ (destructure_hyps n (cons p2 l)))
+ else (map_cons ? (Timp p1 p2) (destructure_hyps n l)))
+ | (cons (Tnot p) l) =>
+ Cases p of
+ (Tnot p1) =>
+ (if (decidability p1) then (destructure_hyps n (cons p1 l))
+ else (map_cons ? (Tnot (Tnot p1)) (destructure_hyps n l)))
+ | (Tor p1 p2) =>
+ (destructure_hyps n (cons (Tnot p1) (cons (Tnot p2) l)))
+ | (Tand p1 p2) =>
+ (if (decidability p1) then
+ (app (destructure_hyps n (cons (Tnot p1) l))
+ (destructure_hyps n (cons (Tnot p2) l)))
+ else (map_cons ? (Tnot p) (destructure_hyps n l)))
+ | _ => (map_cons ? (Tnot p) (destructure_hyps n l))
+ end
+ | (cons x l) => (map_cons ? x (destructure_hyps n l))
+ end
+ end.
+
+Theorem map_cons_val :
+ (ep: PropList; e : (list Z))
+ (p:proposition;l:lhyps)
+ (interp_proposition ep e p) ->
+ (interp_list_hyps ep e l) ->
+ (interp_list_hyps ep e (map_cons ? p l) ).
+
+Induction l; Simpl; [ Auto | Intros; Elim H1; Intro H2; Auto ].
+Save.
+
+Hints Resolve map_cons_val append_valid decidable_correct.
+
+Theorem destructure_hyps_valid :
+ (n:nat) (valid_list_hyps (destructure_hyps n)).
+
+Induction n; [
+ Unfold valid_list_hyps; Simpl; Auto
+| Unfold 2 valid_list_hyps; Intros n1 H ep e lp; Case lp; [
+ Simpl; Auto
+ | Intros p l; Case p;
+ Try (Simpl; Intros; Apply map_cons_val; Simpl; Elim H0; Auto); [
+ Intro p'; Case p';
+ Try (Simpl; Intros; Apply map_cons_val; Simpl; Elim H0; Auto); [
+ Simpl; Intros p1 (H1,H2); Pattern (decidability p1); Apply bool_ind2;
+ Intro H3; [
+ Apply H; Simpl; Split; [ Apply not_not; Auto | Assumption ]
+ | Auto]
+ | Simpl; Intros p1 p2 (H1,H2); Apply H; Simpl;
+ Elim not_or with 1 := H1; Auto
+ | Simpl; Intros p1 p2 (H1,H2);Pattern (decidability p1); Apply bool_ind2;
+ Intro H3; [
+ Apply append_valid; Elim not_and with 2 := H1; [
+ Intro; Left; Apply H; Simpl; Auto
+ | Intro; Right; Apply H; Simpl; Auto
+ | Auto ]
+ | Auto ]]
+ | Simpl; Intros p1 p2 (H1, H2); Apply append_valid;
+ (Elim H1; Intro H3; Simpl; [ Left | Right ]); Apply H; Simpl; Auto
+ | Simpl; Intros; Apply H; Simpl; Tauto
+ | Simpl; Intros p1 p2 (H1, H2); Pattern (decidability p1); Apply bool_ind2;
+ Intro H3; [
+ Apply append_valid; Elim imp_simp with 2:=H1; [
+ Intro H4; Left; Simpl; Apply H; Simpl; Auto
+ | Intro H4; Right; Simpl; Apply H; Simpl; Auto
+ | Auto ]
+ | Auto ]]]].
+
+Save.
+
+Definition prop_stable [f: proposition -> proposition] :=
+ (ep: PropList; e: (list Z); p:proposition)
+ (interp_proposition ep e p) <-> (interp_proposition ep e (f p)).
+
+Definition p_apply_left [f: proposition -> proposition; p : proposition]:=
+ Cases p of
+ (Timp x y) => (Timp (f x) y)
+ | (Tor x y) => (Tor (f x) y)
+ | (Tand x y) => (Tand (f x) y)
+ | (Tnot x) => (Tnot (f x))
+ | x => x
+ end.
+
+Theorem p_apply_left_stable :
+ (f : proposition -> proposition)
+ (prop_stable f) -> (prop_stable (p_apply_left f)).
+
+Unfold prop_stable; Intros f H ep e p; Split;
+(Case p; Simpl; Auto; Intros p1; Elim (H ep e p1); Tauto).
+Save.
+
+Definition p_apply_right [f: proposition -> proposition; p : proposition]:=
+ Cases p of
+ (Timp x y) => (Timp x (f y))
+ | (Tor x y) => (Tor x (f y))
+ | (Tand x y) => (Tand x (f y))
+ | (Tnot x) => (Tnot (f x))
+ | x => x
+ end.
+
+Theorem p_apply_right_stable :
+ (f : proposition -> proposition)
+ (prop_stable f) -> (prop_stable (p_apply_right f)).
+
+Unfold prop_stable; Intros f H ep e p; Split;
+(Case p; Simpl; Auto; [
+ Intros p1; Elim (H ep e p1); Tauto
+ | Intros p1 p2; Elim (H ep e p2); Tauto
+ | Intros p1 p2; Elim (H ep e p2); Tauto
+ | Intros p1 p2; Elim (H ep e p2); Tauto
+ ]).
+Save.
+
+Definition p_invert [f : proposition -> proposition; p : proposition] :=
+Cases p of
+ (EqTerm x y) => (Tnot (f (NeqTerm x y)))
+| (LeqTerm x y) => (Tnot (f (GtTerm x y)))
+| (GeqTerm x y) => (Tnot (f (LtTerm x y)))
+| (GtTerm x y) => (Tnot (f (LeqTerm x y)))
+| (LtTerm x y) => (Tnot (f (GeqTerm x y)))
+| (NeqTerm x y) => (Tnot (f (EqTerm x y)))
+| x => x
+end.
+
+Theorem p_invert_stable :
+ (f : proposition -> proposition)
+ (prop_stable f) -> (prop_stable (p_invert f)).
+
+Unfold prop_stable; Intros f H ep e p; Split;(Case p; Simpl; Auto; [
+ Intros t1 t2; Elim (H ep e (NeqTerm t1 t2)); Simpl; Unfold Zne;
+ Generalize (dec_eq (interp_term e t1) (interp_term e t2));
+ Unfold decidable; Tauto
+| Intros t1 t2; Elim (H ep e (GtTerm t1 t2)); Simpl; Unfold Zgt;
+ Generalize (dec_Zgt (interp_term e t1) (interp_term e t2));
+ Unfold decidable Zgt Zle; Tauto
+| Intros t1 t2; Elim (H ep e (LtTerm t1 t2)); Simpl; Unfold Zlt;
+ Generalize (dec_Zlt (interp_term e t1) (interp_term e t2));
+ Unfold decidable Zge; Tauto
+| Intros t1 t2; Elim (H ep e (LeqTerm t1 t2)); Simpl;
+ Generalize (dec_Zgt (interp_term e t1) (interp_term e t2)); Unfold Zle Zgt;
+ Unfold decidable; Tauto
+| Intros t1 t2; Elim (H ep e (GeqTerm t1 t2)); Simpl;
+ Generalize (dec_Zlt (interp_term e t1) (interp_term e t2)); Unfold Zge Zlt;
+ Unfold decidable; Tauto
+| Intros t1 t2; Elim (H ep e (EqTerm t1 t2)); Simpl;
+ Generalize (dec_eq (interp_term e t1) (interp_term e t2));
+ Unfold decidable Zne; Tauto ]).
+Save.
+
+Theorem Zlt_left_inv : (x,y:Z) `0 <= ((y + (-1)) + (-x))` -> `x<y`.
+
+Intros; Apply Zlt_S_n; Apply Zle_lt_n_Sm;
+Apply (Zsimpl_le_plus_r (Zplus `-1` (Zopp x))); Rewrite Zplus_assoc_l;
+Unfold Zs; Rewrite (Zplus_assoc_r x); Rewrite (Zplus_assoc_l y); Simpl;
+Rewrite Zero_right; Rewrite Zplus_inverse_r; Assumption.
+Save.
+
+Theorem move_right_stable : (s: step) (prop_stable (move_right s)).
+
+Unfold move_right prop_stable; Intros s ep e p; Split; [
+ Simplify; Simpl; Elim (rewrite_stable s e); Simpl; [
+ Symmetry; Apply Zegal_left; Assumption
+ | Intro; Apply Zle_left; Assumption
+ | Intro; Apply Zge_left; Assumption
+ | Intro; Apply Zgt_left; Assumption
+ | Intro; Apply Zlt_left; Assumption
+ | Intro; Apply Zne_left_2; Assumption ]
+| Case p; Simpl; Intros; Auto; Generalize H; Elim (rewrite_stable s); Simpl;
+ Intro H1; [
+ Rewrite (Zplus_n_O (interp_term e t0)); Rewrite H1; Rewrite Zplus_permute;
+ Rewrite Zplus_inverse_r; Rewrite Zero_right; Trivial
+ | Apply (Zsimpl_le_plus_r (Zopp (interp_term e t))); Rewrite Zplus_inverse_r;
+ Assumption
+ | Apply Zle_ge; Apply (Zsimpl_le_plus_r (Zopp (interp_term e t0)));
+ Rewrite Zplus_inverse_r; Assumption
+ | Apply Zlt_gt; Apply Zlt_left_inv; Assumption
+ | Apply Zlt_left_inv; Assumption
+ | Unfold Zne not; Unfold Zne in H1; Intro H2; Apply H1; Rewrite H2;
+ Rewrite Zplus_inverse_r; Trivial ]].
+Save.
+
+
+Fixpoint p_rewrite [s: p_step] : proposition -> proposition :=
+ Cases s of
+ | (P_LEFT s) => (p_apply_left (p_rewrite s))
+ | (P_RIGHT s) => (p_apply_right (p_rewrite s))
+ | (P_STEP s) => (move_right s)
+ | (P_INVERT s) => (p_invert (move_right s))
+ | P_NOP => [p:proposition]p
+ end.
+
+Theorem p_rewrite_stable : (s : p_step) (prop_stable (p_rewrite s)).
+
+
+Induction s; Simpl; [
+ Intros; Apply p_apply_left_stable; Trivial
+| Intros; Apply p_apply_right_stable; Trivial
+| Intros; Apply p_invert_stable; Apply move_right_stable
+| Apply move_right_stable
+| Unfold prop_stable; Simpl; Intros; Split; Auto ].
+Save.
+
+Fixpoint normalize_hyps [l: (list h_step)] : hyps -> hyps :=
+ [lh:hyps] Cases l of
+ nil => lh
+ | (cons (pair_step i s) r) =>
+ (normalize_hyps r (apply_oper_1 i (p_rewrite s) lh))
+ end.
+
+Theorem normalize_hyps_valid :
+ (l: (list h_step)) (valid_hyps (normalize_hyps l)).
+
+Induction l; Unfold valid_hyps; Simpl; [
+ Auto
+| Intros n_s r; Case n_s; Intros n s H ep e lp H1; Apply H;
+ Apply apply_oper_1_valid; [
+ Unfold valid1; Intros ep1 e1 p1 H2; Elim (p_rewrite_stable s ep1 e1 p1);
+ Auto
+ | Assumption ]].
+Save.
+
+Theorem normalize_hyps_goal :
+ (s: (list h_step); ep: PropList; env : (list Z); l: hyps)
+ (interp_goal ep env (normalize_hyps s l)) ->
+ (interp_goal ep env l).
+
+Intros; Apply valid_goal with 2:=H; Apply normalize_hyps_valid.
+Save.
+
+Fixpoint extract_hyp_pos [s: (list direction)] : proposition -> proposition :=
+ [p: proposition]
+ Cases s of
+ | (cons D_left l) =>
+ Cases p of
+ (Tand x y) => (extract_hyp_pos l x)
+ | _ => p
+ end
+ | (cons D_right l) =>
+ Cases p of
+ (Tand x y) => (extract_hyp_pos l y)
+ | _ => p
+ end
+ | (cons D_mono l) =>
+ Cases p of
+ (Tnot x ) => (extract_hyp_neg l x)
+ | _ => p
+ end
+ | _ => p
+ end
+with extract_hyp_neg [s: (list direction)] : proposition -> proposition :=
+ [p: proposition]
+ Cases s of
+ | (cons D_left l) =>
+ Cases p of
+ (Tor x y) => (extract_hyp_neg l x)
+ | (Timp x y) =>
+ (if (decidability x) then (extract_hyp_pos l x) else (Tnot p))
+ | _ => (Tnot p)
+ end
+ | (cons D_right l) =>
+ Cases p of
+ (Tor x y) => (extract_hyp_neg l y)
+ | (Timp x y) => (extract_hyp_neg l y)
+ | _ => (Tnot p)
+ end
+ | (cons D_mono l) =>
+ Cases p of
+ (Tnot x) =>
+ (if (decidability x) then (extract_hyp_pos l x) else (Tnot p))
+ | _ => (Tnot p)
+ end
+ | _ =>
+ Cases p of
+ (Tnot x) => (if (decidability x) then x else (Tnot p))
+ | _ => (Tnot p)
+ end
+ end.
+
+Definition co_valid1 [f: proposition -> proposition] :=
+ (ep : PropList; e: (list Z)) (p1: proposition)
+ (interp_proposition ep e (Tnot p1)) -> (interp_proposition ep e (f p1)).
+
+Theorem extract_valid :
+ (s: (list direction))
+ ((valid1 (extract_hyp_pos s)) /\ (co_valid1 (extract_hyp_neg s))).
+
+Unfold valid1 co_valid1; Induction s; [
+ Split; [
+ Simpl; Auto
+ | Intros ep e p1; Case p1; Simpl; Auto; Intro p; Pattern (decidability p);
+ Apply bool_ind2; [
+ Intro H; Generalize (decidable_correct ep e p H); Unfold decidable; Tauto
+ | Simpl; Auto]]
+| Intros a s' (H1,H2); Simpl in H2; Split; Intros ep e p; Case a; Auto;
+ Case p; Auto; Simpl; Intros;
+ (Apply H1; Tauto) Orelse (Apply H2; Tauto) Orelse
+ (Pattern (decidability p0); Apply bool_ind2; [
+ Intro H3; Generalize (decidable_correct ep e p0 H3);Unfold decidable;
+ Intro H4; Apply H1; Tauto
+ | Intro; Tauto ])].
+
+Save.
+
+Fixpoint decompose_solve [s: e_step] : hyps -> lhyps :=
+ [h:hyps]
+ Cases s of
+ (E_SPLIT i dl s1 s2) =>
+ (Cases (extract_hyp_pos dl (nth_hyps i h)) of
+ (Tor x y) =>
+ (app (decompose_solve s1 (cons x h))
+ (decompose_solve s2 (cons y h)))
+ | (Tnot (Tand x y)) =>
+ (if (decidability x) then
+ (app (decompose_solve s1 (cons (Tnot x) h))
+ (decompose_solve s2 (cons (Tnot y) h)))
+ else (cons h (nil hyps)))
+ | _ => (cons h (nil hyps))
+ end)
+ | (E_EXTRACT i dl s1) =>
+ (decompose_solve s1 (cons (extract_hyp_pos dl (nth_hyps i h)) h))
+ | (E_SOLVE t) => (execute_omega t h)
+ end.
+
+Theorem decompose_solve_valid :
+ (s:e_step)(valid_list_goal (decompose_solve s)).
+
+Intro s; Apply goal_valid; Unfold valid_list_hyps; Elim s; Simpl; Intros; [
+ Cut (interp_proposition ep e1 (extract_hyp_pos l (nth_hyps n lp))); [
+ Case (extract_hyp_pos l (nth_hyps n lp)); Simpl; Auto; [
+ Intro p; Case p; Simpl;Auto; Intros p1 p2 H2;
+ Pattern (decidability p1); Apply bool_ind2; [
+ Intro H3; Generalize (decidable_correct ep e1 p1 H3);
+ Intro H4; Apply append_valid; Elim H4; Intro H5; [
+ Right; Apply H0; Simpl; Tauto
+ | Left; Apply H; Simpl; Tauto ]
+ | Simpl; Auto]
+ | Intros p1 p2 H2; Apply append_valid; Simpl; Elim H2; [
+ Intros H3; Left; Apply H; Simpl; Auto
+ | Intros H3; Right; Apply H0; Simpl; Auto ]]
+ | Elim (extract_valid l); Intros H2 H3; Apply H2; Apply nth_valid; Auto]
+| Intros; Apply H; Simpl; Split; [
+ Elim (extract_valid l); Intros H2 H3; Apply H2; Apply nth_valid; Auto
+ | Auto ]
+| Apply omega_valid with 1:= H].
+
+Save.
+
+(* \subsection{La dernière étape qui élimine tous les séquents inutiles} *)
+
+Definition valid_lhyps [f: lhyps -> lhyps] :=
+ (ep : PropList; e : (list Z)) (lp: lhyps)
+ (interp_list_hyps ep e lp) -> (interp_list_hyps ep e (f lp)).
+
+Fixpoint reduce_lhyps [lp:lhyps] : lhyps :=
+ Cases lp of
+ (cons (cons FalseTerm nil) lp') => (reduce_lhyps lp')
+ | (cons x lp') => (cons x (reduce_lhyps lp'))
+ | nil => (nil hyps)
+ end.
+
+Theorem reduce_lhyps_valid : (valid_lhyps reduce_lhyps).
+
+Unfold valid_lhyps; Intros ep e lp; Elim lp; [
+ Simpl; Auto
+| Intros a l HR; Elim a; [
+ Simpl; Tauto
+ | Intros a1 l1; Case l1; Case a1; Simpl; Try Tauto]].
+Save.
+
+Theorem do_reduce_lhyps :
+ (envp: PropList; env: (list Z); l: lhyps)
+ (interp_list_goal envp env (reduce_lhyps l)) ->
+ (interp_list_goal envp env l).
+
+Intros envp env l H; Apply list_goal_to_hyps; Intro H1;
+Apply list_hyps_to_goal with 1 := H; Apply reduce_lhyps_valid; Assumption.
+Save.
+
+Definition concl_to_hyp := [p:proposition]
+ (if (decidability p) then (Tnot p) else TrueTerm).
+
+Definition do_concl_to_hyp :
+ (envp: PropList; env: (list Z); c : proposition; l:hyps)
+ (interp_goal envp env (cons (concl_to_hyp c) l)) ->
+ (interp_goal_concl envp env c l).
+
+Simpl; Intros envp env c l; Induction l; [
+ Simpl; Unfold concl_to_hyp; Pattern (decidability c); Apply bool_ind2; [
+ Intro H; Generalize (decidable_correct envp env c H); Unfold decidable;
+ Simpl; Tauto
+ | Simpl; Intros H1 H2; Elim H2; Trivial]
+| Simpl; Tauto ].
+Save.
+
+Definition omega_tactic :=
+ [t1:e_step ; t2:(list h_step) ; c:proposition; l:hyps]
+ (reduce_lhyps
+ (decompose_solve t1 (normalize_hyps t2 (cons (concl_to_hyp c) l)))).
+
+Theorem do_omega:
+ (t1: e_step ; t2: (list h_step);
+ envp: PropList; env: (list Z); c: proposition; l:hyps)
+ (interp_list_goal envp env (omega_tactic t1 t2 c l)) ->
+ (interp_goal_concl envp env c l).
+
+Unfold omega_tactic; Intros; Apply do_concl_to_hyp;
+Apply (normalize_hyps_goal t2); Apply (decompose_solve_valid t1);
+Apply do_reduce_lhyps; Assumption.
+Save.