aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Pierre Letouzey <pierre.letouzey@inria.fr>2017-03-22 11:24:27 +0100
committerGravatar Pierre Letouzey <pierre.letouzey@inria.fr>2017-06-13 10:30:29 +0200
commit295107103aaa86db8a31abb0e410123212648d45 (patch)
tree15928f2d0e3752e70938401555faddb48661f34d
parent423d3202fa0f244db36a0b1b45edfa61829201e6 (diff)
BigNums: remove files about BigN,BigZ,BigQ (now in an separate git repo)
See now https://github.com/coq/bignums Int31 is still in the stdlib. Some proofs there has be adapted to avoid the need for BigNumPrelude.
-rw-r--r--.gitignore1
-rw-r--r--CHANGES4
-rw-r--r--Makefile3
-rw-r--r--Makefile.build7
-rw-r--r--Makefile.common12
-rw-r--r--doc/stdlib/index-list.html.template31
-rw-r--r--plugins/syntax/int31_syntax.ml100
-rw-r--r--plugins/syntax/int31_syntax_plugin.mlpack1
-rw-r--r--plugins/syntax/numbers_syntax.ml313
-rw-r--r--plugins/syntax/numbers_syntax_plugin.mlpack1
-rw-r--r--test-suite/failure/int31.v2
-rw-r--r--test-suite/output/Int31Syntax.out14
-rw-r--r--test-suite/output/Int31Syntax.v13
-rw-r--r--test-suite/output/NumbersSyntax.out67
-rw-r--r--test-suite/output/NumbersSyntax.v50
-rw-r--r--test-suite/success/NumberScopes.v21
-rw-r--r--test-suite/success/bigQ.v66
-rw-r--r--theories/Numbers/BigNumPrelude.v411
-rw-r--r--theories/Numbers/Cyclic/Abstract/CyclicAxioms.v2
-rw-r--r--theories/Numbers/Cyclic/Abstract/DoubleType.v (renamed from theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v)1
-rw-r--r--theories/Numbers/Cyclic/Abstract/NZCyclic.v23
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v317
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v437
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v966
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v1494
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v519
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v475
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v621
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v1369
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v356
-rw-r--r--theories/Numbers/Cyclic/Int31/Cyclic31.v255
-rw-r--r--theories/Numbers/Cyclic/ZModulo/ZModulo.v7
-rw-r--r--theories/Numbers/Integer/BigZ/BigZ.v208
-rw-r--r--theories/Numbers/Integer/BigZ/ZMake.v759
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSig.v135
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v527
-rw-r--r--theories/Numbers/Natural/BigN/BigN.v198
-rw-r--r--theories/Numbers/Natural/BigN/NMake.v1706
-rw-r--r--theories/Numbers/Natural/BigN/NMake_gen.ml1017
-rw-r--r--theories/Numbers/Natural/BigN/Nbasic.v569
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSig.v124
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v487
-rw-r--r--theories/Numbers/Rational/BigQ/BigQ.v162
-rw-r--r--theories/Numbers/Rational/BigQ/QMake.v1283
-rw-r--r--theories/Numbers/Rational/SpecViaQ/QSig.v229
45 files changed, 310 insertions, 15053 deletions
diff --git a/.gitignore b/.gitignore
index e52091ee2..db12a9de2 100644
--- a/.gitignore
+++ b/.gitignore
@@ -150,7 +150,6 @@ plugins/ssr/ssrvernac.ml
kernel/byterun/coq_jumptbl.h
kernel/copcodes.ml
tools/tolink.ml
-theories/Numbers/Natural/BigN/NMake_gen.v
ide/index_urls.txt
.lia.cache
checker/names.ml
diff --git a/CHANGES b/CHANGES
index fc95b5ec2..b5aaad725 100644
--- a/CHANGES
+++ b/CHANGES
@@ -50,6 +50,10 @@ Standard Library
and, consequently, choice of representatives in equivalence classes.
Various proof-theoretic characterizations of choice over setoids in
file ChoiceFacts.v.
+- The BigN, BigZ, BigZ libraries are not part anymore of Coq standard
+ library, they are now provided by a separate repository
+ https://github.com/coq/bignums
+ The split has been done just after the Int31 library.
- IZR (Reals) has been changed to produce a compact representation of
integers. As a consequence, IZR is no longer convertible to INR and
diff --git a/Makefile b/Makefile
index 91b024913..66721a3ad 100644
--- a/Makefile
+++ b/Makefile
@@ -89,8 +89,7 @@ EXISTINGMLI := $(call find, '*.mli')
GENML4FILES:= $(ML4FILES:.ml4=.ml)
export GENMLFILES:=$(LEXFILES:.mll=.ml) tools/tolink.ml kernel/copcodes.ml
export GENHFILES:=kernel/byterun/coq_jumptbl.h
-export GENVFILES:=theories/Numbers/Natural/BigN/NMake_gen.v
-export GENFILES:=$(GENMLFILES) $(GENMLIFILES) $(GENHFILES) $(GENVFILES)
+export GENFILES:=$(GENMLFILES) $(GENMLIFILES) $(GENHFILES)
# NB: all files in $(GENFILES) can be created initially, while
# .ml files in $(GENML4FILES) might need some intermediate building.
diff --git a/Makefile.build b/Makefile.build
index 6e048ce94..bae1e7ec7 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -610,11 +610,6 @@ plugins: $(PLUGINSVO) $(PLUGINSCMO)
.PHONY: coqlib theories plugins
-# One of the .v files is macro-generated
-
-theories/Numbers/Natural/BigN/NMake_gen.v: theories/Numbers/Natural/BigN/NMake_gen.ml
- $(OCAML) $< $(TOTARGET)
-
# The .vo files in Init are built with the -noinit option
theories/Init/%.vo theories/Init/%.glob: theories/Init/%.v $(VO_TOOLS_DEP)
@@ -636,7 +631,7 @@ endif
# Dependencies of .v files
-%.v.d: $(D_DEPEND_BEFORE_SRC) %.v $(D_DEPEND_AFTER_SRC) $(COQDEPBOOT) $(GENVFILES)
+%.v.d: $(D_DEPEND_BEFORE_SRC) %.v $(D_DEPEND_AFTER_SRC) $(COQDEPBOOT)
$(SHOW)'COQDEP $<'
$(HIDE)$(COQDEPBOOT) -boot $(DYNDEP) "$<" $(TOTARGET)
diff --git a/Makefile.common b/Makefile.common
index b2e1d47df..ec5e6ac85 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -129,8 +129,8 @@ RTAUTOCMO:=plugins/rtauto/rtauto_plugin.cmo
NATSYNTAXCMO:=plugins/syntax/nat_syntax_plugin.cmo
OTHERSYNTAXCMO:=$(addprefix plugins/syntax/, \
z_syntax_plugin.cmo \
- numbers_syntax_plugin.cmo \
r_syntax_plugin.cmo \
+ int31_syntax_plugin.cmo \
ascii_syntax_plugin.cmo \
string_syntax_plugin.cmo )
DERIVECMO:=plugins/derive/derive_plugin.cmo
@@ -161,14 +161,8 @@ LINKCMX:=$(CORECMA:.cma=.cmxa) $(STATICPLUGINS:.cmo=.cmx)
# vo files
###########################################################################
-GENVOFILES := $(GENVFILES:.v=.vo)
-
-THEORIESVO := $(patsubst %.v,%.vo,$(shell find theories -type f -name "*.v")) \
- $(filter theories/%, $(GENVOFILES))
-
-PLUGINSVO := $(patsubst %.v,%.vo,$(shell find plugins -type f -name "*.v")) \
- $(filter plugins/%, $(GENVOFILES))
-
+THEORIESVO := $(patsubst %.v,%.vo,$(shell find theories -type f -name "*.v"))
+PLUGINSVO := $(patsubst %.v,%.vo,$(shell find plugins -type f -name "*.v"))
ALLVO := $(THEORIESVO) $(PLUGINSVO)
VFILES := $(ALLVO:.vo=.v)
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index aeb0de48a..1b847414f 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -224,7 +224,6 @@ through the <tt>Require Import</tt> command.</p>
<dd>
theories/Numbers/BinNums.v
theories/Numbers/NumPrelude.v
- theories/Numbers/BigNumPrelude.v
theories/Numbers/NaryFunctions.v
</dd>
@@ -256,16 +255,7 @@ through the <tt>Require Import</tt> command.</p>
<dd>
theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
theories/Numbers/Cyclic/Abstract/NZCyclic.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
- theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
+ theories/Numbers/Cyclic/Abstract/DoubleType.v
theories/Numbers/Cyclic/Int31/Cyclic31.v
theories/Numbers/Cyclic/Int31/Ring31.v
theories/Numbers/Cyclic/Int31/Int31.v
@@ -298,12 +288,6 @@ through the <tt>Require Import</tt> command.</p>
theories/Numbers/Natural/Abstract/NProperties.v
theories/Numbers/Natural/Binary/NBinary.v
theories/Numbers/Natural/Peano/NPeano.v
- theories/Numbers/Natural/SpecViaZ/NSig.v
- theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
- theories/Numbers/Natural/BigN/BigN.v
- theories/Numbers/Natural/BigN/Nbasic.v
- theories/Numbers/Natural/BigN/NMake.v
- theories/Numbers/Natural/BigN/NMake_gen.v
</dd>
<dt> <b>&nbsp;&nbsp;Integer</b>:
@@ -331,19 +315,6 @@ through the <tt>Require Import</tt> command.</p>
theories/Numbers/Integer/Abstract/ZDivTrunc.v
theories/Numbers/Integer/Binary/ZBinary.v
theories/Numbers/Integer/NatPairs/ZNatPairs.v
- theories/Numbers/Integer/SpecViaZ/ZSig.v
- theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
- theories/Numbers/Integer/BigZ/BigZ.v
- theories/Numbers/Integer/BigZ/ZMake.v
- </dd>
-
- <dt><b>&nbsp;&nbsp;Rational</b>:
- Abstract and 31-bits-words-based rational arithmetic
- </dt>
- <dd>
- theories/Numbers/Rational/SpecViaQ/QSig.v
- theories/Numbers/Rational/BigQ/BigQ.v
- theories/Numbers/Rational/BigQ/QMake.v
</dd>
</dl>
</dd>
diff --git a/plugins/syntax/int31_syntax.ml b/plugins/syntax/int31_syntax.ml
new file mode 100644
index 000000000..5d1412ba7
--- /dev/null
+++ b/plugins/syntax/int31_syntax.ml
@@ -0,0 +1,100 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open API
+
+(* Poor's man DECLARE PLUGIN *)
+let __coq_plugin_name = "int31_syntax_plugin"
+let () = Mltop.add_known_module __coq_plugin_name
+
+(* digit-based syntax for int31 *)
+
+open Bigint
+open Names
+open Globnames
+open Glob_term
+
+(*** Constants for locating int31 constructors ***)
+
+let make_dir l = DirPath.make (List.rev_map Id.of_string l)
+let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)
+
+let make_mind mp id = Names.MutInd.make2 mp (Label.make id)
+let make_mind_mpfile dir id = make_mind (ModPath.MPfile (make_dir dir)) id
+let make_mind_mpdot dir modname id =
+ let mp = ModPath.MPdot (ModPath.MPfile (make_dir dir), Label.make modname)
+ in make_mind mp id
+
+
+(* int31 stuff *)
+let int31_module = ["Coq"; "Numbers"; "Cyclic"; "Int31"; "Int31"]
+let int31_path = make_path int31_module "int31"
+let int31_id = make_mind_mpfile int31_module
+let int31_scope = "int31_scope"
+
+let int31_construct = ConstructRef ((int31_id "int31",0),1)
+
+let int31_0 = ConstructRef ((int31_id "digits",0),1)
+let int31_1 = ConstructRef ((int31_id "digits",0),2)
+
+(*** Definition of the Non_closed exception, used in the pretty printing ***)
+exception Non_closed
+
+(*** Parsing for int31 in digital notation ***)
+
+(* parses a *non-negative* integer (from bigint.ml) into an int31
+ wraps modulo 2^31 *)
+let int31_of_pos_bigint ?loc n =
+ let ref_construct = CAst.make ?loc (GRef (int31_construct, None)) in
+ let ref_0 = CAst.make ?loc (GRef (int31_0, None)) in
+ let ref_1 = CAst.make ?loc (GRef (int31_1, None)) in
+ let rec args counter n =
+ if counter <= 0 then
+ []
+ else
+ let (q,r) = div2_with_rest n in
+ (if r then ref_1 else ref_0)::(args (counter-1) q)
+ in
+ CAst.make ?loc (GApp (ref_construct, List.rev (args 31 n)))
+
+let error_negative ?loc =
+ CErrors.user_err ?loc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.")
+
+let interp_int31 ?loc n =
+ if is_pos_or_zero n then
+ int31_of_pos_bigint ?loc n
+ else
+ error_negative ?loc
+
+(* Pretty prints an int31 *)
+
+let bigint_of_int31 =
+ let rec args_parsing args cur =
+ match args with
+ | [] -> cur
+ | { CAst.v = GRef (b,_) }::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur)
+ | { CAst.v = GRef (b,_) }::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur))
+ | _ -> raise Non_closed
+ in
+ function
+ | { CAst.v = GApp ({ CAst.v = GRef (c, _) }, args) } when eq_gr c int31_construct -> args_parsing args zero
+ | _ -> raise Non_closed
+
+let uninterp_int31 i =
+ try
+ Some (bigint_of_int31 i)
+ with Non_closed ->
+ None
+
+(* Actually declares the interpreter for int31 *)
+let _ = Notation.declare_numeral_interpreter int31_scope
+ (int31_path, int31_module)
+ interp_int31
+ ([CAst.make (GRef (int31_construct, None))],
+ uninterp_int31,
+ true)
diff --git a/plugins/syntax/int31_syntax_plugin.mlpack b/plugins/syntax/int31_syntax_plugin.mlpack
new file mode 100644
index 000000000..54a5bc0cd
--- /dev/null
+++ b/plugins/syntax/int31_syntax_plugin.mlpack
@@ -0,0 +1 @@
+Int31_syntax
diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml
deleted file mode 100644
index fb657c47c..000000000
--- a/plugins/syntax/numbers_syntax.ml
+++ /dev/null
@@ -1,313 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open API
-
-(* Poor's man DECLARE PLUGIN *)
-let __coq_plugin_name = "numbers_syntax_plugin"
-let () = Mltop.add_known_module __coq_plugin_name
-
-(* digit-based syntax for int31, bigN bigZ and bigQ *)
-
-open Bigint
-open Names
-open Globnames
-open Glob_term
-
-(*** Constants for locating int31 / bigN / bigZ / bigQ constructors ***)
-
-let make_dir l = DirPath.make (List.rev_map Id.of_string l)
-let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)
-
-let make_mind mp id = Names.MutInd.make2 mp (Label.make id)
-let make_mind_mpfile dir id = make_mind (ModPath.MPfile (make_dir dir)) id
-let make_mind_mpdot dir modname id =
- let mp = ModPath.MPdot (ModPath.MPfile (make_dir dir), Label.make modname)
- in make_mind mp id
-
-
-(* int31 stuff *)
-let int31_module = ["Coq"; "Numbers"; "Cyclic"; "Int31"; "Int31"]
-let int31_path = make_path int31_module "int31"
-let int31_id = make_mind_mpfile int31_module
-let int31_scope = "int31_scope"
-
-let int31_construct = ConstructRef ((int31_id "int31",0),1)
-
-let int31_0 = ConstructRef ((int31_id "digits",0),1)
-let int31_1 = ConstructRef ((int31_id "digits",0),2)
-
-
-(* bigN stuff*)
-let zn2z_module = ["Coq"; "Numbers"; "Cyclic"; "DoubleCyclic"; "DoubleType"]
-let zn2z_path = make_path zn2z_module "zn2z"
-let zn2z_id = make_mind_mpfile zn2z_module
-
-let zn2z_W0 = ConstructRef ((zn2z_id "zn2z",0),1)
-let zn2z_WW = ConstructRef ((zn2z_id "zn2z",0),2)
-
-let bigN_module = ["Coq"; "Numbers"; "Natural"; "BigN"; "BigN" ]
-let bigN_path = make_path (bigN_module@["BigN"]) "t"
-let bigN_t = make_mind_mpdot bigN_module "BigN" "t'"
-let bigN_scope = "bigN_scope"
-
-(* number of inlined level of bigN (actually the level 0 to n_inlined-1 are inlined) *)
-let n_inlined = 7
-
-let bigN_constructor i =
- ConstructRef ((bigN_t,0),(min i n_inlined)+1)
-
-(*bigZ stuff*)
-let bigZ_module = ["Coq"; "Numbers"; "Integer"; "BigZ"; "BigZ" ]
-let bigZ_path = make_path (bigZ_module@["BigZ"]) "t"
-let bigZ_t = make_mind_mpdot bigZ_module "BigZ" "t_"
-let bigZ_scope = "bigZ_scope"
-
-let bigZ_pos = ConstructRef ((bigZ_t,0),1)
-let bigZ_neg = ConstructRef ((bigZ_t,0),2)
-
-
-(*bigQ stuff*)
-let bigQ_module = ["Coq"; "Numbers"; "Rational"; "BigQ"; "BigQ"]
-let bigQ_path = make_path (bigQ_module@["BigQ"]) "t"
-let bigQ_t = make_mind_mpdot bigQ_module "BigQ" "t_"
-let bigQ_scope = "bigQ_scope"
-
-let bigQ_z = ConstructRef ((bigQ_t,0),1)
-
-
-(*** Definition of the Non_closed exception, used in the pretty printing ***)
-exception Non_closed
-
-(*** Parsing for int31 in digital notation ***)
-
-(* parses a *non-negative* integer (from bigint.ml) into an int31
- wraps modulo 2^31 *)
-let int31_of_pos_bigint ?loc n =
- let ref_construct = CAst.make ?loc @@ GRef (int31_construct, None) in
- let ref_0 = CAst.make ?loc @@ GRef (int31_0, None) in
- let ref_1 = CAst.make ?loc @@ GRef (int31_1, None) in
- let rec args counter n =
- if counter <= 0 then
- []
- else
- let (q,r) = div2_with_rest n in
- (if r then ref_1 else ref_0)::(args (counter-1) q)
- in
- CAst.make ?loc @@ GApp (ref_construct, List.rev (args 31 n))
-
-let error_negative ?loc =
- CErrors.user_err ?loc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.")
-
-let interp_int31 ?loc n =
- if is_pos_or_zero n then
- int31_of_pos_bigint ?loc n
- else
- error_negative ?loc
-
-(* Pretty prints an int31 *)
-
-let bigint_of_int31 =
- let rec args_parsing args cur =
- match args with
- | [] -> cur
- | { CAst.v = GRef (b,_) }::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur)
- | { CAst.v = GRef (b,_) }::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur))
- | _ -> raise Non_closed
- in
- function
- | { CAst.v = GApp ({ CAst.v = GRef (c, _)}, args) } when eq_gr c int31_construct -> args_parsing args zero
- | _ -> raise Non_closed
-
-let uninterp_int31 i =
- try
- Some (bigint_of_int31 i)
- with Non_closed ->
- None
-
-(* Actually declares the interpreter for int31 *)
-let _ = Notation.declare_numeral_interpreter int31_scope
- (int31_path, int31_module)
- interp_int31
- ([CAst.make @@ GRef (int31_construct, None)],
- uninterp_int31,
- true)
-
-
-(*** Parsing for bigN in digital notation ***)
-(* the base for bigN (in Coq) that is 2^31 in our case *)
-let base = pow two 31
-
-(* base of the bigN of height N : (2^31)^(2^n) *)
-let rank n =
- let rec rk n pow2 =
- if n <= 0 then pow2
- else rk (n-1) (mult pow2 pow2)
- in rk n base
-
-(* splits a number bi at height n, that is the rest needs 2^n int31 to be stored
- it is expected to be used only when the quotient would also need 2^n int31 to be
- stored *)
-let split_at n bi =
- euclid bi (rank (n-1))
-
-(* search the height of the Coq bigint needed to represent the integer bi *)
-let height bi =
- let rec hght n pow2 =
- if less_than bi pow2 then n
- else hght (n+1) (mult pow2 pow2)
- in hght 0 base
-
-(* n must be a non-negative integer (from bigint.ml) *)
-let word_of_pos_bigint ?loc hght n =
- let ref_W0 = CAst.make ?loc @@ GRef (zn2z_W0, None) in
- let ref_WW = CAst.make ?loc @@ GRef (zn2z_WW, None) in
- let rec decomp hgt n =
- if hgt <= 0 then
- int31_of_pos_bigint ?loc n
- else if equal n zero then
- CAst.make ?loc @@ GApp (ref_W0, [CAst.make ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)])
- else
- let (h,l) = split_at hgt n in
- CAst.make ?loc @@ GApp (ref_WW, [CAst.make ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None);
- decomp (hgt-1) h;
- decomp (hgt-1) l])
- in
- decomp hght n
-
-let bigN_of_pos_bigint ?loc n =
- let h = height n in
- let ref_constructor = CAst.make ?loc @@ GRef (bigN_constructor h, None) in
- let word = word_of_pos_bigint ?loc h n in
- let args =
- if h < n_inlined then [word]
- else [Nat_syntax_plugin.Nat_syntax.nat_of_int ?loc (of_int (h-n_inlined));word]
- in
- CAst.make ?loc @@ GApp (ref_constructor, args)
-
-let bigN_error_negative ?loc =
- CErrors.user_err ?loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.")
-
-let interp_bigN ?loc n =
- if is_pos_or_zero n then
- bigN_of_pos_bigint ?loc n
- else
- bigN_error_negative ?loc
-
-
-(* Pretty prints a bigN *)
-
-let bigint_of_word =
- let rec get_height rc =
- match rc with
- | { CAst.v = GApp ({ CAst.v = GRef(c,_)}, [_;lft;rght]) } when eq_gr c zn2z_WW ->
- 1+max (get_height lft) (get_height rght)
- | _ -> 0
- in
- let rec transform hght rc =
- match rc with
- | { CAst.v = GApp ({ CAst.v = GRef(c,_)},_)} when eq_gr c zn2z_W0-> zero
- | { CAst.v = GApp ({ CAst.v = GRef(c,_)}, [_;lft;rght]) } when eq_gr c zn2z_WW->
- let new_hght = hght-1 in
- add (mult (rank new_hght)
- (transform new_hght lft))
- (transform new_hght rght)
- | _ -> bigint_of_int31 rc
- in
- fun rc ->
- let hght = get_height rc in
- transform hght rc
-
-let bigint_of_bigN rc =
- match rc with
- | { CAst.v = GApp (_,[one_arg]) } -> bigint_of_word one_arg
- | { CAst.v = GApp (_,[_;second_arg]) } -> bigint_of_word second_arg
- | _ -> raise Non_closed
-
-let uninterp_bigN rc =
- try
- Some (bigint_of_bigN rc)
- with Non_closed ->
- None
-
-
-(* declare the list of constructors of bigN used in the declaration of the
- numeral interpreter *)
-
-let bigN_list_of_constructors =
- let rec build i =
- if i < n_inlined+1 then
- (CAst.make @@ GRef (bigN_constructor i,None))::(build (i+1))
- else
- []
- in
- build 0
-
-(* Actually declares the interpreter for bigN *)
-let _ = Notation.declare_numeral_interpreter bigN_scope
- (bigN_path, bigN_module)
- interp_bigN
- (bigN_list_of_constructors,
- uninterp_bigN,
- true)
-
-
-(*** Parsing for bigZ in digital notation ***)
-let interp_bigZ ?loc n =
- let ref_pos = CAst.make ?loc @@ GRef (bigZ_pos, None) in
- let ref_neg = CAst.make ?loc @@ GRef (bigZ_neg, None) in
- if is_pos_or_zero n then
- CAst.make ?loc @@ GApp (ref_pos, [bigN_of_pos_bigint ?loc n])
- else
- CAst.make ?loc @@ GApp (ref_neg, [bigN_of_pos_bigint ?loc (neg n)])
-
-(* pretty printing functions for bigZ *)
-let bigint_of_bigZ = function
- | { CAst.v = GApp ({ CAst.v = GRef(c,_) }, [one_arg])} when eq_gr c bigZ_pos -> bigint_of_bigN one_arg
- | { CAst.v = GApp ({ CAst.v = GRef(c,_) }, [one_arg])} when eq_gr c bigZ_neg ->
- let opp_val = bigint_of_bigN one_arg in
- if equal opp_val zero then
- raise Non_closed
- else
- neg opp_val
- | _ -> raise Non_closed
-
-
-let uninterp_bigZ rc =
- try
- Some (bigint_of_bigZ rc)
- with Non_closed ->
- None
-
-(* Actually declares the interpreter for bigZ *)
-let _ = Notation.declare_numeral_interpreter bigZ_scope
- (bigZ_path, bigZ_module)
- interp_bigZ
- ([CAst.make @@ GRef (bigZ_pos, None);
- CAst.make @@ GRef (bigZ_neg, None)],
- uninterp_bigZ,
- true)
-
-(*** Parsing for bigQ in digital notation ***)
-let interp_bigQ ?loc n =
- let ref_z = CAst.make ?loc @@ GRef (bigQ_z, None) in
- CAst.make ?loc @@ GApp (ref_z, [interp_bigZ ?loc n])
-
-let uninterp_bigQ rc =
- try match rc with
- | { CAst.v = GApp ({ CAst.v = GRef(c,_)}, [one_arg]) } when eq_gr c bigQ_z ->
- Some (bigint_of_bigZ one_arg)
- | _ -> None (* we don't pretty-print yet fractions *)
- with Non_closed -> None
-
-(* Actually declares the interpreter for bigQ *)
-let _ = Notation.declare_numeral_interpreter bigQ_scope
- (bigQ_path, bigQ_module)
- interp_bigQ
- ([CAst.make @@ GRef (bigQ_z, None)], uninterp_bigQ,
- true)
diff --git a/plugins/syntax/numbers_syntax_plugin.mlpack b/plugins/syntax/numbers_syntax_plugin.mlpack
deleted file mode 100644
index e48c00a0d..000000000
--- a/plugins/syntax/numbers_syntax_plugin.mlpack
+++ /dev/null
@@ -1 +0,0 @@
-Numbers_syntax
diff --git a/test-suite/failure/int31.v b/test-suite/failure/int31.v
index b1d112247..ed4c9f9c7 100644
--- a/test-suite/failure/int31.v
+++ b/test-suite/failure/int31.v
@@ -1,4 +1,4 @@
-Require Import Int31 BigN.
+Require Import Int31 Cyclic31.
Open Scope int31_scope.
diff --git a/test-suite/output/Int31Syntax.out b/test-suite/output/Int31Syntax.out
new file mode 100644
index 000000000..4e8796c14
--- /dev/null
+++ b/test-suite/output/Int31Syntax.out
@@ -0,0 +1,14 @@
+I31
+ : digits31 int31
+2
+ : int31
+660865024
+ : int31
+2 + 2
+ : int31
+2 + 2
+ : int31
+ = 4
+ : int31
+ = 710436486
+ : int31
diff --git a/test-suite/output/Int31Syntax.v b/test-suite/output/Int31Syntax.v
new file mode 100644
index 000000000..83be3b976
--- /dev/null
+++ b/test-suite/output/Int31Syntax.v
@@ -0,0 +1,13 @@
+Require Import Int31 Cyclic31.
+
+Open Scope int31_scope.
+Check I31. (* Would be nice to have I31 : digits->digits->...->int31
+ For the moment, I31 : digits31 int31, which is better
+ than (fix nfun .....) size int31 *)
+Check 2.
+Check 1000000000000000000. (* = 660865024, after modulo 2^31 *)
+Check (add31 2 2).
+Check (2+2).
+Eval vm_compute in 2+2.
+Eval vm_compute in 65675757 * 565675998.
+Close Scope int31_scope.
diff --git a/test-suite/output/NumbersSyntax.out b/test-suite/output/NumbersSyntax.out
deleted file mode 100644
index b2677b6ad..000000000
--- a/test-suite/output/NumbersSyntax.out
+++ /dev/null
@@ -1,67 +0,0 @@
-I31
- : digits31 int31
-2
- : int31
-660865024
- : int31
-2 + 2
- : int31
-2 + 2
- : int31
- = 4
- : int31
- = 710436486
- : int31
-2
- : BigN.t'
-1000000000000000000
- : BigN.t'
-2 + 2
- : bigN
-2 + 2
- : bigN
- = 4
- : bigN
- = 37151199385380486
- : bigN
- = 1267650600228229401496703205376
- : bigN
-2
- : BigZ.t_
--1000000000000000000
- : BigZ.t_
-2 + 2
- : BigZ.t_
-2 + 2
- : BigZ.t_
- = 4
- : BigZ.t_
- = 37151199385380486
- : BigZ.t_
- = 1267650600228229401496703205376
- : BigZ.t_
-2
- : BigQ.t_
--1000000000000000000
- : BigQ.t_
-2 + 2
- : bigQ
-2 + 2
- : bigQ
- = 4
- : bigQ
- = 37151199385380486
- : bigQ
-6562 # 456
- : BigQ.t_
- = 3281 # 228
- : bigQ
- = -1 # 10000
- : bigQ
- = 100
- : bigQ
- = 515377520732011331036461129765621272702107522001
- # 1267650600228229401496703205376
- : bigQ
- = 1
- : bigQ
diff --git a/test-suite/output/NumbersSyntax.v b/test-suite/output/NumbersSyntax.v
deleted file mode 100644
index 4fbf56ab1..000000000
--- a/test-suite/output/NumbersSyntax.v
+++ /dev/null
@@ -1,50 +0,0 @@
-
-Require Import BigQ.
-
-Open Scope int31_scope.
-Check I31. (* Would be nice to have I31 : digits->digits->...->int31
- For the moment, I31 : digits31 int31, which is better
- than (fix nfun .....) size int31 *)
-Check 2.
-Check 1000000000000000000. (* = 660865024, after modulo 2^31 *)
-Check (add31 2 2).
-Check (2+2).
-Eval vm_compute in 2+2.
-Eval vm_compute in 65675757 * 565675998.
-Close Scope int31_scope.
-
-Open Scope bigN_scope.
-Check 2.
-Check 1000000000000000000.
-Check (BigN.add 2 2).
-Check (2+2).
-Eval vm_compute in 2+2.
-Eval vm_compute in 65675757 * 565675998.
-Eval vm_compute in 2^100.
-Close Scope bigN_scope.
-
-Open Scope bigZ_scope.
-Check 2.
-Check -1000000000000000000.
-Check (BigZ.add 2 2).
-Check (2+2).
-Eval vm_compute in 2+2.
-Eval vm_compute in 65675757 * 565675998.
-Eval vm_compute in (-2)^100.
-Close Scope bigZ_scope.
-
-Open Scope bigQ_scope.
-Check 2.
-Check -1000000000000000000.
-Check (BigQ.add 2 2).
-Check (2+2).
-Eval vm_compute in 2+2.
-Eval vm_compute in 65675757 * 565675998.
-(* fractions *)
-Check (6562 # 456). (* Nota: # is BigQ.Qq i.e. base fractions *)
-Eval vm_compute in (BigQ.red (6562 # 456)).
-Eval vm_compute in (1/-10000).
-Eval vm_compute in (BigQ.red (1/(1/100))). (* back to integers... *)
-Eval vm_compute in ((2/3)^(-100)).
-Eval vm_compute in BigQ.red ((2/3)^(-1000) * (2/3)^(1000)).
-Close Scope bigQ_scope.
diff --git a/test-suite/success/NumberScopes.v b/test-suite/success/NumberScopes.v
index 6d7872107..155863747 100644
--- a/test-suite/success/NumberScopes.v
+++ b/test-suite/success/NumberScopes.v
@@ -39,24 +39,3 @@ Definition f_nat (x:nat) := x.
Definition f_nat' (x:Nat.t) := x.
Check (f_nat 1).
Check (f_nat' 1).
-
-Require Import BigN.
-Check (BigN.add 1 2).
-Check (BigN.add_comm 1 2).
-Check (BigN.min_comm 1 2).
-Definition f_bigN (x:bigN) := x.
-Check (f_bigN 1).
-
-Require Import BigZ.
-Check (BigZ.add 1 2).
-Check (BigZ.add_comm 1 2).
-Check (BigZ.min_comm 1 2).
-Definition f_bigZ (x:bigZ) := x.
-Check (f_bigZ 1).
-
-Require Import BigQ.
-Check (BigQ.add 1 2).
-Check (BigQ.add_comm 1 2).
-Check (BigQ.min_comm 1 2).
-Definition f_bigQ (x:bigQ) := x.
-Check (f_bigQ 1). \ No newline at end of file
diff --git a/test-suite/success/bigQ.v b/test-suite/success/bigQ.v
deleted file mode 100644
index 7fd0cf669..000000000
--- a/test-suite/success/bigQ.v
+++ /dev/null
@@ -1,66 +0,0 @@
-Require Import BigQ.
-Import List.
-
-Definition pi_4_approx_low' :=
-(5066193963420194617885108698600649932059391557720145469382602092416947640628637390992675949693715109726079394291478795603894419483819297806310615866892414925850691415582239745615128821983865262221858109336884967754871321668348027076234335167119885298878199925731495390387858629042311908406056230882123787019283378509712244687397013657159455607193734144010901984756727174636853404278421831024545476850410085042498464474261035780891759930905778986584183710930670670301831474144997069400304290351567959717683444430666444319233768399342338059169002790777424962570605618705584660815518973602995097110557181643034682308210782171804373210646804613922337450953858508244032293753591878060539465788294318856859293281629951093130167801471787011911886414492513677892193100809508943832528344473873460853362957387889412799458784754514139679847887887544849825173792522272708046699681079289358082661375778523609867456540595586031625044964543428047238934233579184772793670436643502740076366994465457847106782560289782615794595755672643440040123002018908935362541166831619056664637901929131328502017686713274283777724453661234225382109584471950444925886358166551424008707439387934109226545596919797083495958300914344992836193126080289565652575543234385558967555959267746932292860747199382633363026440008828134867747920263181610216905129926037611247017868033961426567047355301676870662406173724238530061264149506666345040372864118731705584795947926329181826992456072045382170981478151356381437136818835196834068650217794381425547036331194595892801393225038235274901050364737353586927051766717037643833477566087835266968086513005761986678747515870298138062157791066648217784877968385924845017637219384732843791052551854695220023477365706464590594542001161575677402761543188277502092362285265847964496740584911576627239093631932307473445797386335961743298553548881544486940399236133577915988716682746485564575640818803540680574730591500432326858763829791848612343662539095316357052823005419355719381626599487868023399182174939253393897549026675976384326749445831606130546375395770778462506203752920470130305293966478109733954117063941901686840180727195741528561335809865193566993349413786715403053579411364371500063193205131503024022217701373077790337150298315820556080596579100618643147698304927957576213733526923182742441048553793831725592624850721293495085399785588171300815789795594858916409701139277050529011775828846362873246196866089783324522718656445008090114701320562608474099248873638488023114015981013142490827777895317580810590743940417298263300561876701828404744082864248409230009391001735746615476377303707782123483770118391136826609366946585715225248587168403619476143657107412319421501162805102723455593551478028055839072686207007765300258935153546418515706362733656094770289090398825190320430416955807878686642673124733998295439657633866090085982598765253268688814792672416195730086607425842181518560588819896560847103627615434844684536463752986969865794019299978956052589825441828842338163389851892617560591840546654410705167593310272272965900821031821380595084783691324416454359888103920904935692840264474003367023256964191100139001239923263691779167792867186165635514824889759796850863175082506408142175595463676408992027105356481220754473245821534527625758942093801142305560662681150069082553674495761075895588095760081401141419460482860852822686860785424514171214889677926763812031823537071721974799922995763666175738785000806081164280471363125324839717808977470218218571800106898347366938927189989988149888641129263448064762730769285877330997355234347773807099829665997515649429224335217107760728789764718885665291038706425454675746218345291274054088843647602239258308472486102933167465443294268551209015027897159307743987020521392788721231001835675584104894174434637260464035122611721657641428625505184886116917149318963070896162119215386541876236027342810162765609201440423207771441367926085768438143507025739041041240810056881304230519058117534418374553879198061289605354335880794397478047346975609179199801003098836622253165101961484972165230151495472006888128587168049198312469715081555662345452800468933420359802645393289853553618279788400476187713990872203669487294118461245455333004125835663010526985716431187034663870796866708678078952110615910196519835267441831874676895301527286826106517027821074816850326548617513767142627360001181210946100011774672126943957522004190414960909074050454565964857276407084991922274068961845339154089866785707764290964299529444616711194034827611771558783466230353209661849406004241580029437779784290315347968833708422223285859451369907260780956405036020581705441364379616715041818815829810906212826084485200785283123265202151252852134381195424724503189247411069117189489985791487434549080447866370484866697404176437230771558469231403088139693477706784802801265075586678597768511791952562627345622499328
- # 100788726492580594349650258277496659410917619472657560321971265983799894639441017438166498752997098978003489632843381325240982516059309714013145358125224597827602157516585886911710102182473475545864474089191789296685473601331678556438310133356793199956062857423397512495293688453655805536015029176541424005214818033707522950635262669828538132795615008381824067071229426026518897202246241637377064076189277685257166926338187911595052586669184297526234794666364657344206795357967279911782849686515024121916258300642000317525374433525235296287037535618423661645124459323811792936193272341688261801253469089129439519903538495370298752436267926761998785090092411372633429302950606054074205533246665546979112178855223925266166034953000200646676762301817000435641690517142795144469005596172113586738287118865058604922865654348297975054956781513943444060257230946224520058527667925776273088622386666860662470481606622952298649177217986593047495967209669116410592230626047083795555559776477430548946990993890380787911273437967786556742804566652408275798339221179283430482118140020742719695900657696142739101628984271513292954605191778803974738871043737934546460016184719168074062912083778327025499841998124431899131874519812228674255796948879306477894924710085384116453080236862135706628989104070747737689294987000148388110561753028594988959655591699155508380909698460304884908709246116411180876105681720036833487450945730831039969246996849503525429840196651386469599438064049723005123629385485140945945416764414133189625489032807860400751723995946290581976152580477047961138617997133510128194027510895265424780627975864980749945631413855375897945293107842908479797077570371447220506451229526132919408351287454305932886749170523056147842439813407002950370505941417426433452282518739345666494683448699945734453214481915512562995906034771246088038719298959180199052759295868161570318718927430655393250250811804905393113074074574608255523847592006804881016594060188745212933427473833239777228852952217878690668413947367586040297784502192683200664398064682201012931468052982448022330449955215606614483165425935154496289535573901139223034819824408001205784146243892228030383941863746839845526558421740316887532141893650230936137269356278754487130882868595412163277284772124736531380334814212708066069618080153747333573454834500999083737284449542481264971030785043701582134343596645346132964567391370300568578875509971483039720438955919863275044932311289587494336123538202079503922025306586828117649623642521324286648529829664567232756108169459356549144779085080036654897525078792273443307070502103724611233768453196294899770515940520895908289018412144327894912660060761908970811602375085884115384049610753387776858733798341463052471017393165656926510611173543365663267563198760597092606598728110197523695339144204179424646442294307593146446562536865057987897899655645968129515654148044008249646703504419478535298270862753806142083172190778193001810574370442181909146645889199829207284871551220439225371051511970054965951914399901815408791418836185742573331879114400013259342896515702942707292473805188905427717363630137869116872433627556880809120353079342030725196065815470427569172350436988386579444534375353968759750750178342190349607711313840613843718547859929387259163285524671855725511880656411741012446023392964655239624520090988149679656514996202498334816938716757663800773997302639681907686195671083505910700098597156238624351157219093280177066146218516478636356056420098245995113668018177690728654922707281126889313941750547830163078886329630807850633273613622550216189245162735650139455042125252043274668279981753287687674520319519360593091620297805736177366738063651905396783336064579717230286821545930579779462534206093794040878198825916141099864730374109311705285661056855668930671948265232862757146615431791375559792290479316263924560826544387396762768331402198937951439504767950821089741987629257538953417586416459087855138539304027013800937360598578194413362672871055543854633921502486683911956250444582746421552178164852341035733290405311280719066037175324627429434912416361334254696649419037348733709488576582107382055914938194078813926926742828297826939120316120573453588052056773875836843924877773978390546387248009519202370375478981843515393806263037580338009594022254079586380520797699651840576286033587273591899639699077044271208886940540056794360292760863657703246410020854088849880453524038877935317875884698324859548991680533307680053872403383516589028793015681082435908524045497475001609824047204954932626536311826911363867426654549346914317405110707189532251727848751560224936842128628673253616256326013555922159336370177663785738170802777550686079119049748734352584409583136667752555307842739679930698964098088960000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000)%bigQ
-.
-
-Definition pi_4_approx_high' :=
-(5066193963420194617885108698600649932059391557720145469382602092416947640628637390992675949693715109726079394291478795603894419483819297806310615866892414925850691415582239745615128821983865262221858109336884967754871321668348027076234335167119885298878199925731495390387858629042311908406056230882123787019283378509712244687397013657159455607193734144010901984756727174636853404278421831024545476850410085042498464474261035780891759930905778986584183710930670670301831474144997069400304290351567959717683444430666444319233768399342338059169002790777424962570605618705584660815518973602995097110557181643034682308210788409308322071457087096445676662503017187903223859814905546579050729173916234740628466315449085686468204847296426235788544874405450791749423436215032927889914519102361378633666267941326393265376660400091389373564825046526561381561278586121772300141564909333667988204680492088607706214346458601842899721615765319505314310192693665547163360402786722105590252780194994950097926184146718893770363322073641336811404180286358079915338791029818581497746089864894356686643882883410392601500048021013346713450539807687779704798018559373507951388092945938366448668853081682176581336156031434604604833692503597621519809826880683536141897075567053733515342478008373282599947520770191238802249392773327261328133194484586433840861730959791563023761306622956165536481335792721379318928171897265310054788931201902441066997927781894934061720760080768154565282051604447333036111267534150649674590201404453202347064545359869105856798745664471694795576801148562495225166002814304124970965817043547048503388910163287916513427409193998045119986267987892522931703487420953769290650229176116308194977201080691718825944370436642709192983358059711255925052564016519597530235976618244111239816418652282585432539731271068892992142956810775762851238126881225206289553948196520384709574383566733478326330112084307565420647201107231840508040019131253750047046446929758911912155202166566751947087545292626353331520202690130850009389387290465497377022080531269511355734944672010542204118978272180881335465227900174033380001851066811103401787656367819132934758616060307366679580043123632565656840669377840733018248707250548277181001911990237151790533341326223932843775840498222236867608395855700891719880219904948672458645420169533565809609056209006342663841718949396996175294237942265325043426430990062217643279654512512640557763489491751115437780462208361129433667449740743123546232162409802316714286708788831227582498585478334315076725145986771341647015244092760289407649044493584479944044779273447198382196766547779885914425854375158084417582279211000449529495605127376707776277159376010648950025135061284601443461110447113346277147728593420397807946636800365109579479211273476195727270004743568492888900356505584731622538401071221591141889158461271000051210318027818802379539544396973228585821742794928813630781709195703717312953337431290682263448669168179857644544116657440168099166467471736180072984407514757289757495435699300593165669101965987430482600019222913485092771346963058673132443387835726110205958057187517487684058179749952286341120230051432903482992282688283815697442898155194928723360957436110770317998431272108100149791425689283090777721270428030993332057319821685391144252815655146410678839177846108260765981523812232294638190350688210999605869296307711846463311346627138400477211801219366400312514793356564308532012682051019030257269068628100171220662165246389309014292764479226570049772046255291379151017129899157296574099437276707879597755725339406865738613810979022640265737120949077721294633786520294559343155148383011293584240192753971366644780434237846862975993387453786681995831719537733846579480995517357440575781962659282856696638992709756358478461648462532279323701121386551383509193782388241965285971965887701816406255233933761008649762854363984142178331798953040874526844255758512982810004271235810681505829473926495256537353108899526434200682024946218302499640511518360332022463196599199779172637638655415918976955930735312156870786600023896830267884391447789311101069654521354446521135407720085038662159974712373018912537116964809382149581004863115431780452188813210275393919111435118030412595133958954313836191108258769640843644195537185904547405641078708492098917460393911427237155683288565433183738513871595286090814836422982384810033331519971102974091067660369548406192526284519976668985518575216481570167748402860759832933071281814538397923687510782620605409323050353840034866296214149657376249634795555007199540807313397329050410326609108411299737760271566308288500400587417017113933243099961248847368789383209110747378488312550109911605079801570534271874115018095746872468910162721975463388518648962869080447866370484866697404176437230771558469231403088139693477706784802801265075586678597768511791952562627345622499328
- # 100788726492580594349650258277496659410917619472657560321971265983799894639441017438166498752997098978003489632843381325240982516059309714013145358125224597827602157516585886911710102182473475545864474089191789296685473601331678556438310133356793199956062857423397512495293688453655805536015029176541424005214818033707522950635262669828538132795615008381824067071229426026518897202246241637377064076189277685257166926338187911595052586669184297526234794666364657344206795357967279911782849686515024121916258300642000317525374433525235296287037535618423661645124459323811792936193272341688261801253469089129439519903538495370298752436267926761998785090092411372633429302950606054074205533246665546979112178855223925266166034953000200646676762301817000435641690517142795144469005596172113586738287118865058604922865654348297975054956781513943444060257230946224520058527667925776273088622386666860662470481606622952298649177217986593047495967209669116410592230626047083795555559776477430548946990993890380787911273437967786556742804566652408275798339221179283430482118140020742719695900657696142739101628984271513292954605191778803974738871043737934546460016184719168074062912083778327025499841998124431899131874519812228674255796948879306477894924710085384116453080236862135706628989104070747737689294987000148388110561753028594988959655591699155508380909698460304884908709246116411180876105681720036833487450945730831039969246996849503525429840196651386469599438064049723005123629385485140945945416764414133189625489032807860400751723995946290581976152580477047961138617997133510128194027510895265424780627975864980749945631413855375897945293107842908479797077570371447220506451229526132919408351287454305932886749170523056147842439813407002950370505941417426433452282518739345666494683448699945734453214481915512562995906034771246088038719298959180199052759295868161570318718927430655393250250811804905393113074074574608255523847592006804881016594060188745212933427473833239777228852952217878690668413947367586040297784502192683200664398064682201012931468052982448022330449955215606614483165425935154496289535573901139223034819824408001205784146243892228030383941863746839845526558421740316887532141893650230936137269356278754487130882868595412163277284772124736531380334814212708066069618080153747333573454834500999083737284449542481264971030785043701582134343596645346132964567391370300568578875509971483039720438955919863275044932311289587494336123538202079503922025306586828117649623642521324286648529829664567232756108169459356549144779085080036654897525078792273443307070502103724611233768453196294899770515940520895908289018412144327894912660060761908970811602375085884115384049610753387776858733798341463052471017393165656926510611173543365663267563198760597092606598728110197523695339144204179424646442294307593146446562536865057987897899655645968129515654148044008249646703504419478535298270862753806142083172190778193001810574370442181909146645889199829207284871551220439225371051511970054965951914399901815408791418836185742573331879114400013259342896515702942707292473805188905427717363630137869116872433627556880809120353079342030725196065815470427569172350436988386579444534375353968759750750178342190349607711313840613843718547859929387259163285524671855725511880656411741012446023392964655239624520090988149679656514996202498334816938716757663800773997302639681907686195671083505910700098597156238624351157219093280177066146218516478636356056420098245995113668018177690728654922707281126889313941750547830163078886329630807850633273613622550216189245162735650139455042125252043274668279981753287687674520319519360593091620297805736177366738063651905396783336064579717230286821545930579779462534206093794040878198825916141099864730374109311705285661056855668930671948265232862757146615431791375559792290479316263924560826544387396762768331402198937951439504767950821089741987629257538953417586416459087855138539304027013800937360598578194413362672871055543854633921502486683911956250444582746421552178164852341035733290405311280719066037175324627429434912416361334254696649419037348733709488576582107382055914938194078813926926742828297826939120316120573453588052056773875836843924877773978390546387248009519202370375478981843515393806263037580338009594022254079586380520797699651840576286033587273591899639699077044271208886940540056794360292760863657703246410020854088849880453524038877935317875884698324859548991680533307680053872403383516589028793015681082435908524045497475001609824047204954932626536311826911363867426654549346914317405110707189532251727848751560224936842128628673253616256326013555922159336370177663785738170802777550686079119049748734352584409583136667752555307842739679930698964098088960000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000000)%bigQ
-.
-
-Fixpoint numden_Rcontfrac_tailrecB (accu: list bigZ) (n1 d1: bigZ) (n2 d2: bigZ) (fuel: nat) {struct fuel}:
- (list bigZ * bigQ * bigQ) :=
- let default := (rev_append accu nil, BigQ.div (BigQ.Qz n1) (BigQ.Qz d1), BigQ.div (BigQ.Qz n2) (BigQ.Qz d2)) in
- match fuel with
- | O => default
- | S fuel' =>
- let '(q1, r1) := BigZ.div_eucl n1 d1 in
- let '(q2, r2) := BigZ.div_eucl n2 d2 in
- match BigZ.eqb q1 q2 with
- | false => default
- | true =>
- let r1_is_zero := BigZ.eqb r1 0 in
- let r2_is_zero := BigZ.eqb r2 0 in
- match Bool.eqb r1_is_zero r2_is_zero with
- | false => default
- | true =>
- match r1_is_zero with
- | true =>
- match BigZ.eqb q1 1 with
- | true => (rev_append accu nil, 1%bigQ, 1%bigQ)
- | false => (rev_append ((q1 - 1)%bigZ :: accu) nil, 1%bigQ, 1%bigQ)
- end
- | false => numden_Rcontfrac_tailrecB (q1 :: accu) d1 r1 d2 r2 fuel'
- end
- end
- end
- end.
-
-Definition Bnum b :=
- match b with
- | BigQ.Qz t => t
- | BigQ.Qq n d =>
- if (d =? BigN.zero)%bigN then 0%bigZ else n
- end.
-
-Definition Bden b :=
- match b with
- | BigQ.Qz _ => 1%bigN
- | BigQ.Qq _ d => if (d =? BigN.zero)%bigN then 1%bigN else d
- end.
-
-Definition rat_Rcontfrac_tailrecB q1 q2 :=
- numden_Rcontfrac_tailrecB nil (Bnum q1) (BigZ.Pos (Bden q1)) (Bnum q2) (BigZ.Pos (Bden q2)).
-
-Definition pi_4_contfrac :=
- rat_Rcontfrac_tailrecB pi_4_approx_low' pi_4_approx_high' 3000.
-
-(* The following used to fail because of a non canonical representation of 0 in
-the bytecode interpreter. Bug reported privately by Tahina Ramananandro. *)
-Goal pi_4_contfrac = pi_4_contfrac.
-vm_compute.
-reflexivity.
-Qed.
diff --git a/theories/Numbers/BigNumPrelude.v b/theories/Numbers/BigNumPrelude.v
deleted file mode 100644
index bd8930872..000000000
--- a/theories/Numbers/BigNumPrelude.v
+++ /dev/null
@@ -1,411 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-(** * BigNumPrelude *)
-
-(** Auxiliary functions & theorems used for arbitrary precision efficient
- numbers. *)
-
-
-Require Import ArithRing.
-Require Export ZArith.
-Require Export Znumtheory.
-Require Export Zpow_facts.
-
-Declare ML Module "numbers_syntax_plugin".
-
-(* *** Nota Bene ***
- All results that were general enough have been moved in ZArith.
- Only remain here specialized lemmas and compatibility elements.
- (P.L. 5/11/2007).
-*)
-
-
-Local Open Scope Z_scope.
-
-(* For compatibility of scripts, weaker version of some lemmas of Z.div *)
-
-Lemma Zlt0_not_eq : forall n, 0<n -> n<>0.
-Proof.
- auto with zarith.
-Qed.
-
-Definition Zdiv_mult_cancel_r a b c H := Zdiv.Zdiv_mult_cancel_r a b c (Zlt0_not_eq _ H).
-Definition Zdiv_mult_cancel_l a b c H := Zdiv.Zdiv_mult_cancel_r a b c (Zlt0_not_eq _ H).
-Definition Z_div_plus_l a b c H := Zdiv.Z_div_plus_full_l a b c (Zlt0_not_eq _ H).
-
-(* Automation *)
-
-Hint Extern 2 (Z.le _ _) =>
- (match goal with
- |- Zpos _ <= Zpos _ => exact (eq_refl _)
-| H: _ <= ?p |- _ <= ?p => apply Z.le_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Z.lt_le_incl; apply Z.le_lt_trans with (2 := H)
- end).
-
-Hint Extern 2 (Z.lt _ _) =>
- (match goal with
- |- Zpos _ < Zpos _ => exact (eq_refl _)
-| H: _ <= ?p |- _ <= ?p => apply Z.lt_le_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Z.le_lt_trans with (2 := H)
- end).
-
-
-Hint Resolve Z.lt_gt Z.le_ge Z_div_pos: zarith.
-
-(**************************************
- Properties of order and product
- **************************************)
-
- Theorem beta_lex: forall a b c d beta,
- a * beta + b <= c * beta + d ->
- 0 <= b < beta -> 0 <= d < beta ->
- a <= c.
- Proof.
- intros a b c d beta H1 (H3, H4) (H5, H6).
- assert (a - c < 1); auto with zarith.
- apply Z.mul_lt_mono_pos_r with beta; auto with zarith.
- apply Z.le_lt_trans with (d - b); auto with zarith.
- rewrite Z.mul_sub_distr_r; auto with zarith.
- Qed.
-
- Theorem beta_lex_inv: forall a b c d beta,
- a < c -> 0 <= b < beta ->
- 0 <= d < beta ->
- a * beta + b < c * beta + d.
- Proof.
- intros a b c d beta H1 (H3, H4) (H5, H6).
- case (Z.le_gt_cases (c * beta + d) (a * beta + b)); auto with zarith.
- intros H7. contradict H1. apply Z.le_ngt. apply beta_lex with (1 := H7); auto.
- Qed.
-
- Lemma beta_mult : forall h l beta,
- 0 <= h < beta -> 0 <= l < beta -> 0 <= h*beta+l < beta^2.
- Proof.
- intros h l beta H1 H2;split. auto with zarith.
- rewrite <- (Z.add_0_r (beta^2)); rewrite Z.pow_2_r;
- apply beta_lex_inv;auto with zarith.
- Qed.
-
- Lemma Zmult_lt_b :
- forall b x y, 0 <= x < b -> 0 <= y < b -> 0 <= x * y <= b^2 - 2*b + 1.
- Proof.
- intros b x y (Hx1,Hx2) (Hy1,Hy2);split;auto with zarith.
- apply Z.le_trans with ((b-1)*(b-1)).
- apply Z.mul_le_mono_nonneg;auto with zarith.
- apply Z.eq_le_incl; ring.
- Qed.
-
- Lemma sum_mul_carry : forall xh xl yh yl wc cc beta,
- 1 < beta ->
- 0 <= wc < beta ->
- 0 <= xh < beta ->
- 0 <= xl < beta ->
- 0 <= yh < beta ->
- 0 <= yl < beta ->
- 0 <= cc < beta^2 ->
- wc*beta^2 + cc = xh*yl + xl*yh ->
- 0 <= wc <= 1.
- Proof.
- intros xh xl yh yl wc cc beta U H1 H2 H3 H4 H5 H6 H7.
- assert (H8 := Zmult_lt_b beta xh yl H2 H5).
- assert (H9 := Zmult_lt_b beta xl yh H3 H4).
- split;auto with zarith.
- apply beta_lex with (cc) (beta^2 - 2) (beta^2); auto with zarith.
- Qed.
-
- Theorem mult_add_ineq: forall x y cross beta,
- 0 <= x < beta ->
- 0 <= y < beta ->
- 0 <= cross < beta ->
- 0 <= x * y + cross < beta^2.
- Proof.
- intros x y cross beta HH HH1 HH2.
- split; auto with zarith.
- apply Z.le_lt_trans with ((beta-1)*(beta-1)+(beta-1)); auto with zarith.
- apply Z.add_le_mono; auto with zarith.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- rewrite ?Z.mul_sub_distr_l, ?Z.mul_sub_distr_r, Z.pow_2_r; auto with zarith.
- Qed.
-
- Theorem mult_add_ineq2: forall x y c cross beta,
- 0 <= x < beta ->
- 0 <= y < beta ->
- 0 <= c*beta + cross <= 2*beta - 2 ->
- 0 <= x * y + (c*beta + cross) < beta^2.
- Proof.
- intros x y c cross beta HH HH1 HH2.
- split; auto with zarith.
- apply Z.le_lt_trans with ((beta-1)*(beta-1)+(2*beta-2));auto with zarith.
- apply Z.add_le_mono; auto with zarith.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- rewrite ?Z.mul_sub_distr_l, ?Z.mul_sub_distr_r, Z.pow_2_r; auto with zarith.
- Qed.
-
-Theorem mult_add_ineq3: forall x y c cross beta,
- 0 <= x < beta ->
- 0 <= y < beta ->
- 0 <= cross <= beta - 2 ->
- 0 <= c <= 1 ->
- 0 <= x * y + (c*beta + cross) < beta^2.
- Proof.
- intros x y c cross beta HH HH1 HH2 HH3.
- apply mult_add_ineq2;auto with zarith.
- split;auto with zarith.
- apply Z.le_trans with (1*beta+cross);auto with zarith.
- Qed.
-
-Hint Rewrite Z.mul_1_r Z.mul_0_r Z.mul_1_l Z.mul_0_l Z.add_0_l Z.add_0_r Z.sub_0_r: rm10.
-
-
-(**************************************
- Properties of Z.div and Z.modulo
-**************************************)
-
-Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
- Proof.
- intros a b H H1;case (Z_mod_lt a b);auto with zarith;intros H2 H3;split;auto.
- case (Z.le_gt_cases b a); intros H4; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- Qed.
-
-
- Theorem Zmod_distr: forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
- (2 ^a * r + t) mod (2 ^ b) = (2 ^a * r) mod (2 ^ b) + t.
- Proof.
- intros a b r t (H1, H2) H3 (H4, H5).
- assert (t < 2 ^ b).
- apply Z.lt_le_trans with (1:= H5); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- rewrite Zplus_mod; auto with zarith.
- rewrite Zmod_small with (a := t); auto with zarith.
- apply Zmod_small; auto with zarith.
- split; auto with zarith.
- assert (0 <= 2 ^a * r); auto with zarith.
- apply Z.add_nonneg_nonneg; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- pattern (2 ^ b) at 2; replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a);
- try ring.
- apply Z.add_le_lt_mono; auto with zarith.
- replace b with ((b - a) + a); try ring.
- rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Z.mul_1_l (2 ^a));
- try rewrite <- Z.mul_sub_distr_r.
- rewrite (Z.mul_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l;
- auto with zarith.
- rewrite (Z.mul_comm (2 ^a)); apply Z.mul_le_mono_nonneg_r; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- Qed.
-
- Theorem Zmod_shift_r:
- forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
- (r * 2 ^a + t) mod (2 ^ b) = (r * 2 ^a) mod (2 ^ b) + t.
- Proof.
- intros a b r t (H1, H2) H3 (H4, H5).
- assert (t < 2 ^ b).
- apply Z.lt_le_trans with (1:= H5); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- rewrite Zplus_mod; auto with zarith.
- rewrite Zmod_small with (a := t); auto with zarith.
- apply Zmod_small; auto with zarith.
- split; auto with zarith.
- assert (0 <= 2 ^a * r); auto with zarith.
- apply Z.add_nonneg_nonneg; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- pattern (2 ^ b) at 2;replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); try ring.
- apply Z.add_le_lt_mono; auto with zarith.
- replace b with ((b - a) + a); try ring.
- rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Z.mul_1_l (2 ^a));
- try rewrite <- Z.mul_sub_distr_r.
- repeat rewrite (fun x => Z.mul_comm x (2 ^ a)); rewrite Zmult_mod_distr_l;
- auto with zarith.
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- Qed.
-
- Theorem Zdiv_shift_r:
- forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
- (r * 2 ^a + t) / (2 ^ b) = (r * 2 ^a) / (2 ^ b).
- Proof.
- intros a b r t (H1, H2) H3 (H4, H5).
- assert (Eq: t < 2 ^ b); auto with zarith.
- apply Z.lt_le_trans with (1 := H5); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- pattern (r * 2 ^ a) at 1; rewrite Z_div_mod_eq with (b := 2 ^ b);
- auto with zarith.
- rewrite <- Z.add_assoc.
- rewrite <- Zmod_shift_r; auto with zarith.
- rewrite (Z.mul_comm (2 ^ b)); rewrite Z_div_plus_full_l; auto with zarith.
- rewrite (fun x y => @Zdiv_small (x mod y)); auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- Qed.
-
-
- Lemma shift_unshift_mod : forall n p a,
- 0 <= a < 2^n ->
- 0 <= p <= n ->
- a * 2^p = a / 2^(n - p) * 2^n + (a*2^p) mod 2^n.
- Proof.
- intros n p a H1 H2.
- pattern (a*2^p) at 1;replace (a*2^p) with
- (a*2^p/2^n * 2^n + a*2^p mod 2^n).
- 2:symmetry;rewrite (Z.mul_comm (a*2^p/2^n));apply Z_div_mod_eq.
- replace (a * 2 ^ p / 2 ^ n) with (a / 2 ^ (n - p));trivial.
- replace (2^n) with (2^(n-p)*2^p).
- symmetry;apply Zdiv_mult_cancel_r.
- destruct H1;trivial.
- cut (0 < 2^p); auto with zarith.
- rewrite <- Zpower_exp.
- replace (n-p+p) with n;trivial. ring.
- omega. omega.
- apply Z.lt_gt. apply Z.pow_pos_nonneg;auto with zarith.
- Qed.
-
-
- Lemma shift_unshift_mod_2 : forall n p a, 0 <= p <= n ->
- ((a * 2 ^ (n - p)) mod (2^n) / 2 ^ (n - p)) mod (2^n) =
- a mod 2 ^ p.
- Proof.
- intros.
- rewrite Zmod_small.
- rewrite Zmod_eq by (auto with zarith).
- unfold Z.sub at 1.
- rewrite Z_div_plus_l by (auto with zarith).
- assert (2^n = 2^(n-p)*2^p).
- rewrite <- Zpower_exp by (auto with zarith).
- replace (n-p+p) with n; auto with zarith.
- rewrite H0.
- rewrite <- Zdiv_Zdiv, Z_div_mult by (auto with zarith).
- rewrite (Z.mul_comm (2^(n-p))), Z.mul_assoc.
- rewrite <- Z.mul_opp_l.
- rewrite Z_div_mult by (auto with zarith).
- symmetry; apply Zmod_eq; auto with zarith.
-
- remember (a * 2 ^ (n - p)) as b.
- destruct (Z_mod_lt b (2^n)); auto with zarith.
- split.
- apply Z_div_pos; auto with zarith.
- apply Zdiv_lt_upper_bound; auto with zarith.
- apply Z.lt_le_trans with (2^n); auto with zarith.
- rewrite <- (Z.mul_1_r (2^n)) at 1.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- cut (0 < 2 ^ (n-p)); auto with zarith.
- Qed.
-
- Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p.
- Proof.
- intros p x Hle;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_le_lower_bound;auto with zarith.
- replace (2^p) with 0.
- destruct x;compute;intro;discriminate.
- destruct p;trivial;discriminate.
- Qed.
-
- Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y.
- Proof.
- intros p x y H;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_lt_upper_bound;auto with zarith.
- apply Z.lt_le_trans with y;auto with zarith.
- rewrite <- (Z.mul_1_r y);apply Z.mul_le_mono_nonneg;auto with zarith.
- assert (0 < 2^p);auto with zarith.
- replace (2^p) with 0.
- destruct x;change (0<y);auto with zarith.
- destruct p;trivial;discriminate.
- Qed.
-
- Theorem Zgcd_div_pos a b:
- 0 < b -> 0 < Z.gcd a b -> 0 < b / Z.gcd a b.
- Proof.
- intros Hb Hg.
- assert (H : 0 <= b / Z.gcd a b) by (apply Z.div_pos; auto with zarith).
- Z.le_elim H; trivial.
- rewrite (Zdivide_Zdiv_eq (Z.gcd a b) b), <- H, Z.mul_0_r in Hb;
- auto using Z.gcd_divide_r with zarith.
- Qed.
-
- Theorem Zdiv_neg a b:
- a < 0 -> 0 < b -> a / b < 0.
- Proof.
- intros Ha Hb.
- assert (b > 0) by omega.
- generalize (Z_mult_div_ge a _ H); intros.
- assert (b * (a / b) < 0)%Z.
- apply Z.le_lt_trans with a; auto with zarith.
- destruct b; try (compute in Hb; discriminate).
- destruct (a/Zpos p)%Z.
- compute in H1; discriminate.
- compute in H1; discriminate.
- compute; auto.
- Qed.
-
- Lemma Zdiv_gcd_zero : forall a b, b / Z.gcd a b = 0 -> b <> 0 ->
- Z.gcd a b = 0.
- Proof.
- intros.
- generalize (Zgcd_is_gcd a b); destruct 1.
- destruct H2 as (k,Hk).
- generalize H; rewrite Hk at 1.
- destruct (Z.eq_dec (Z.gcd a b) 0) as [H'|H']; auto.
- rewrite Z_div_mult_full; auto.
- intros; subst k; simpl in *; subst b; elim H0; auto.
- Qed.
-
- Lemma Zgcd_mult_rel_prime : forall a b c,
- Z.gcd a c = 1 -> Z.gcd b c = 1 -> Z.gcd (a*b) c = 1.
- Proof.
- intros.
- rewrite Zgcd_1_rel_prime in *.
- apply rel_prime_sym; apply rel_prime_mult; apply rel_prime_sym; auto.
- Qed.
-
- Lemma Zcompare_gt : forall (A:Type)(a a':A)(p q:Z),
- match (p?=q)%Z with Gt => a | _ => a' end =
- if Z_le_gt_dec p q then a' else a.
- Proof.
- intros.
- destruct Z_le_gt_dec as [H|H].
- red in H.
- destruct (p?=q)%Z; auto; elim H; auto.
- rewrite H; auto.
- Qed.
-
-Theorem Zbounded_induction :
- (forall Q : Z -> Prop, forall b : Z,
- Q 0 ->
- (forall n, 0 <= n -> n < b - 1 -> Q n -> Q (n + 1)) ->
- forall n, 0 <= n -> n < b -> Q n)%Z.
-Proof.
-intros Q b Q0 QS.
-set (Q' := fun n => (n < b /\ Q n) \/ (b <= n)).
-assert (H : forall n, 0 <= n -> Q' n).
-apply natlike_rec2; unfold Q'.
-destruct (Z.le_gt_cases b 0) as [H | H]. now right. left; now split.
-intros n H IH. destruct IH as [[IH1 IH2] | IH].
-destruct (Z.le_gt_cases (b - 1) n) as [H1 | H1].
-right; auto with zarith.
-left. split; [auto with zarith | now apply (QS n)].
-right; auto with zarith.
-unfold Q' in *; intros n H1 H2. destruct (H n H1) as [[H3 H4] | H3].
-assumption. now apply Z.le_ngt in H3.
-Qed.
-
-Lemma Zsquare_le x : x <= x*x.
-Proof.
-destruct (Z.lt_ge_cases 0 x).
-- rewrite <- Z.mul_1_l at 1.
- rewrite <- Z.mul_le_mono_pos_r; auto with zarith.
-- pose proof (Z.square_nonneg x); auto with zarith.
-Qed.
diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
index 3312161ae..857580198 100644
--- a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
+++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
@@ -17,7 +17,7 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import Znumtheory.
-Require Import BigNumPrelude.
+Require Import Zpow_facts.
Require Import DoubleType.
Local Open Scope Z_scope.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/Abstract/DoubleType.v
index abd567a85..d60c19ea5 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
+++ b/theories/Numbers/Cyclic/Abstract/DoubleType.v
@@ -67,4 +67,3 @@ Fixpoint word (w:Type) (n:nat) : Type :=
| O => w
| S n => zn2z (word w n)
end.
-
diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
index df9b83392..3f9b7b297 100644
--- a/theories/Numbers/Cyclic/Abstract/NZCyclic.v
+++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
@@ -9,7 +9,8 @@
(************************************************************************)
Require Export NZAxioms.
-Require Import BigNumPrelude.
+Require Import ZArith.
+Require Import Zpow_facts.
Require Import DoubleType.
Require Import CyclicAxioms.
@@ -139,6 +140,26 @@ rewrite 2 ZnZ.of_Z_correct; auto with zarith.
symmetry; apply Zmod_small; auto with zarith.
Qed.
+Theorem Zbounded_induction :
+ (forall Q : Z -> Prop, forall b : Z,
+ Q 0 ->
+ (forall n, 0 <= n -> n < b - 1 -> Q n -> Q (n + 1)) ->
+ forall n, 0 <= n -> n < b -> Q n)%Z.
+Proof.
+intros Q b Q0 QS.
+set (Q' := fun n => (n < b /\ Q n) \/ (b <= n)).
+assert (H : forall n, 0 <= n -> Q' n).
+apply natlike_rec2; unfold Q'.
+destruct (Z.le_gt_cases b 0) as [H | H]. now right. left; now split.
+intros n H IH. destruct IH as [[IH1 IH2] | IH].
+destruct (Z.le_gt_cases (b - 1) n) as [H1 | H1].
+right; auto with zarith.
+left. split; [auto with zarith | now apply (QS n)].
+right; auto with zarith.
+unfold Q' in *; intros n H1 H2. destruct (H n H1) as [[H3 H4] | H3].
+assumption. now apply Z.le_ngt in H3.
+Qed.
+
Lemma B_holds : forall n : Z, 0 <= n < wB -> B n.
Proof.
intros n [H1 H2].
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
deleted file mode 100644
index 407bcca4b..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
+++ /dev/null
@@ -1,317 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-
-Local Open Scope Z_scope.
-
-Section DoubleAdd.
- Variable w : Type.
- Variable w_0 : w.
- Variable w_1 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_W0 : w -> zn2z w.
- Variable ww_1 : zn2z w.
- Variable w_succ_c : w -> carry w.
- Variable w_add_c : w -> w -> carry w.
- Variable w_add_carry_c : w -> w -> carry w.
- Variable w_succ : w -> w.
- Variable w_add : w -> w -> w.
- Variable w_add_carry : w -> w -> w.
-
- Definition ww_succ_c x :=
- match x with
- | W0 => C0 ww_1
- | WW xh xl =>
- match w_succ_c xl with
- | C0 l => C0 (WW xh l)
- | C1 l =>
- match w_succ_c xh with
- | C0 h => C0 (WW h w_0)
- | C1 h => C1 W0
- end
- end
- end.
-
- Definition ww_succ x :=
- match x with
- | W0 => ww_1
- | WW xh xl =>
- match w_succ_c xl with
- | C0 l => WW xh l
- | C1 l => w_W0 (w_succ xh)
- end
- end.
-
- Definition ww_add_c x y :=
- match x, y with
- | W0, _ => C0 y
- | _, W0 => C0 x
- | WW xh xl, WW yh yl =>
- match w_add_c xl yl with
- | C0 l =>
- match w_add_c xh yh with
- | C0 h => C0 (WW h l)
- | C1 h => C1 (w_WW h l)
- end
- | C1 l =>
- match w_add_carry_c xh yh with
- | C0 h => C0 (WW h l)
- | C1 h => C1 (w_WW h l)
- end
- end
- end.
-
- Variable R : Type.
- Variable f0 f1 : zn2z w -> R.
-
- Definition ww_add_c_cont x y :=
- match x, y with
- | W0, _ => f0 y
- | _, W0 => f0 x
- | WW xh xl, WW yh yl =>
- match w_add_c xl yl with
- | C0 l =>
- match w_add_c xh yh with
- | C0 h => f0 (WW h l)
- | C1 h => f1 (w_WW h l)
- end
- | C1 l =>
- match w_add_carry_c xh yh with
- | C0 h => f0 (WW h l)
- | C1 h => f1 (w_WW h l)
- end
- end
- end.
-
- (* ww_add et ww_add_carry conserve la forme normale s'il n'y a pas
- de debordement *)
- Definition ww_add x y :=
- match x, y with
- | W0, _ => y
- | _, W0 => x
- | WW xh xl, WW yh yl =>
- match w_add_c xl yl with
- | C0 l => WW (w_add xh yh) l
- | C1 l => WW (w_add_carry xh yh) l
- end
- end.
-
- Definition ww_add_carry_c x y :=
- match x, y with
- | W0, W0 => C0 ww_1
- | W0, WW yh yl => ww_succ_c (WW yh yl)
- | WW xh xl, W0 => ww_succ_c (WW xh xl)
- | WW xh xl, WW yh yl =>
- match w_add_carry_c xl yl with
- | C0 l =>
- match w_add_c xh yh with
- | C0 h => C0 (WW h l)
- | C1 h => C1 (WW h l)
- end
- | C1 l =>
- match w_add_carry_c xh yh with
- | C0 h => C0 (WW h l)
- | C1 h => C1 (w_WW h l)
- end
- end
- end.
-
- Definition ww_add_carry x y :=
- match x, y with
- | W0, W0 => ww_1
- | W0, WW yh yl => ww_succ (WW yh yl)
- | WW xh xl, W0 => ww_succ (WW xh xl)
- | WW xh xl, WW yh yl =>
- match w_add_carry_c xl yl with
- | C0 l => WW (w_add xh yh) l
- | C1 l => WW (w_add_carry xh yh) l
- end
- end.
-
- (*Section DoubleProof.*)
- Variable w_digits : positive.
- Variable w_to_Z : w -> Z.
-
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
- Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_1 : [|w_1|] = 1.
- Variable spec_ww_1 : [[ww_1]] = 1.
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
- Variable spec_w_succ_c : forall x, [+|w_succ_c x|] = [|x|] + 1.
- Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
- Variable spec_w_add_carry_c :
- forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1.
- Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB.
- Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB.
- Variable spec_w_add_carry :
- forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB.
-
- Lemma spec_ww_succ_c : forall x, [+[ww_succ_c x]] = [[x]] + 1.
- Proof.
- destruct x as [ |xh xl];simpl. apply spec_ww_1.
- generalize (spec_w_succ_c xl);destruct (w_succ_c xl) as [l|l];
- intro H;unfold interp_carry in H. simpl;rewrite H;ring.
- rewrite <- Z.add_assoc;rewrite <- H;rewrite Z.mul_1_l.
- assert ([|l|] = 0). generalize (spec_to_Z xl)(spec_to_Z l);omega.
- rewrite H0;generalize (spec_w_succ_c xh);destruct (w_succ_c xh) as [h|h];
- intro H1;unfold interp_carry in H1.
- simpl;rewrite H1;rewrite spec_w_0;ring.
- unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB.
- assert ([|xh|] = wB - 1). generalize (spec_to_Z xh)(spec_to_Z h);omega.
- rewrite H2;ring.
- Qed.
-
- Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
- Proof.
- destruct x as [ |xh xl];trivial.
- destruct y as [ |yh yl]. rewrite Z.add_0_r;trivial.
- simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
- with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring.
- generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l];
- intros H;unfold interp_carry in H;rewrite <- H.
- generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h];
- intros H1;unfold interp_carry in *;rewrite <- H1. trivial.
- repeat rewrite Z.mul_1_l;rewrite spec_w_WW;rewrite wwB_wBwB; ring.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh)
- as [h|h]; intros H1;unfold interp_carry in *;rewrite <- H1.
- simpl;ring.
- repeat rewrite Z.mul_1_l;rewrite wwB_wBwB;rewrite spec_w_WW;ring.
- Qed.
-
- Section Cont.
- Variable P : zn2z w -> zn2z w -> R -> Prop.
- Variable x y : zn2z w.
- Variable spec_f0 : forall r, [[r]] = [[x]] + [[y]] -> P x y (f0 r).
- Variable spec_f1 : forall r, wwB + [[r]] = [[x]] + [[y]] -> P x y (f1 r).
-
- Lemma spec_ww_add_c_cont : P x y (ww_add_c_cont x y).
- Proof.
- destruct x as [ |xh xl];trivial.
- apply spec_f0;trivial.
- destruct y as [ |yh yl].
- apply spec_f0;rewrite Z.add_0_r;trivial.
- simpl.
- generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l];
- intros H;unfold interp_carry in H.
- generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h];
- intros H1;unfold interp_carry in *.
- apply spec_f0. simpl;rewrite H;rewrite H1;ring.
- apply spec_f1. simpl;rewrite spec_w_WW;rewrite H.
- rewrite Z.add_assoc;rewrite wwB_wBwB. rewrite Z.pow_2_r; rewrite <- Z.mul_add_distr_r.
- rewrite Z.mul_1_l in H1;rewrite H1;ring.
- generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh)
- as [h|h]; intros H1;unfold interp_carry in *.
- apply spec_f0;simpl;rewrite H1. rewrite Z.mul_add_distr_r.
- rewrite <- Z.add_assoc;rewrite H;ring.
- apply spec_f1. rewrite spec_w_WW;rewrite wwB_wBwB.
- rewrite Z.add_assoc; rewrite Z.pow_2_r; rewrite <- Z.mul_add_distr_r.
- rewrite Z.mul_1_l in H1;rewrite H1. rewrite Z.mul_add_distr_r.
- rewrite <- Z.add_assoc;rewrite H; simpl; ring.
- Qed.
-
- End Cont.
-
- Lemma spec_ww_add_carry_c :
- forall x y, [+[ww_add_carry_c x y]] = [[x]] + [[y]] + 1.
- Proof.
- destruct x as [ |xh xl];intro y.
- exact (spec_ww_succ_c y).
- destruct y as [ |yh yl].
- rewrite Z.add_0_r;exact (spec_ww_succ_c (WW xh xl)).
- simpl; replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1)
- with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring.
- generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl)
- as [l|l];intros H;unfold interp_carry in H;rewrite <- H.
- generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h];
- intros H1;unfold interp_carry in H1;rewrite <- H1. trivial.
- unfold interp_carry;repeat rewrite Z.mul_1_l;simpl;rewrite wwB_wBwB;ring.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh)
- as [h|h];intros H1;unfold interp_carry in H1;rewrite <- H1. trivial.
- unfold interp_carry;rewrite spec_w_WW;
- repeat rewrite Z.mul_1_l;simpl;rewrite wwB_wBwB;ring.
- Qed.
-
- Lemma spec_ww_succ : forall x, [[ww_succ x]] = ([[x]] + 1) mod wwB.
- Proof.
- destruct x as [ |xh xl];simpl.
- rewrite spec_ww_1;rewrite Zmod_small;trivial.
- split;[intro;discriminate|apply wwB_pos].
- rewrite <- Z.add_assoc;generalize (spec_w_succ_c xl);
- destruct (w_succ_c xl) as[l|l];intro H;unfold interp_carry in H;rewrite <-H.
- rewrite Zmod_small;trivial.
- rewrite wwB_wBwB;apply beta_mult;apply spec_to_Z.
- assert ([|l|] = 0). clear spec_ww_1 spec_w_1 spec_w_0.
- assert (H1:= spec_to_Z l); assert (H2:= spec_to_Z xl); omega.
- rewrite H0;rewrite Z.add_0_r;rewrite <- Z.mul_add_distr_r;rewrite wwB_wBwB.
- rewrite Z.pow_2_r; rewrite Zmult_mod_distr_r;try apply lt_0_wB.
- rewrite spec_w_W0;rewrite spec_w_succ;trivial.
- Qed.
-
- Lemma spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB.
- Proof.
- destruct x as [ |xh xl];intros y.
- rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
- destruct y as [ |yh yl].
- change [[W0]] with 0;rewrite Z.add_0_r.
- rewrite Zmod_small;trivial.
- exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)).
- simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
- with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring.
- generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l];
- unfold interp_carry;intros H;simpl;rewrite <- H.
- rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial.
- Qed.
-
- Lemma spec_ww_add_carry :
- forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB.
- Proof.
- destruct x as [ |xh xl];intros y.
- exact (spec_ww_succ y).
- destruct y as [ |yh yl].
- change [[W0]] with 0;rewrite Z.add_0_r. exact (spec_ww_succ (WW xh xl)).
- simpl;replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1)
- with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring.
- generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl)
- as [l|l];unfold interp_carry;intros H;rewrite <- H;simpl ww_to_Z.
- rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial.
- Qed.
-
-(* End DoubleProof. *)
-End DoubleAdd.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
deleted file mode 100644
index e94a891dd..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
+++ /dev/null
@@ -1,437 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith Ndigits.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-
-Local Open Scope Z_scope.
-
-Local Infix "<<" := Pos.shiftl_nat (at level 30).
-
-Section DoubleBase.
- Variable w : Type.
- Variable w_0 : w.
- Variable w_1 : w.
- Variable w_Bm1 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_0W : w -> zn2z w.
- Variable w_digits : positive.
- Variable w_zdigits: w.
- Variable w_add: w -> w -> zn2z w.
- Variable w_to_Z : w -> Z.
- Variable w_compare : w -> w -> comparison.
-
- Definition ww_digits := xO w_digits.
-
- Definition ww_zdigits := w_add w_zdigits w_zdigits.
-
- Definition ww_to_Z := zn2z_to_Z (base w_digits) w_to_Z.
-
- Definition ww_1 := WW w_0 w_1.
-
- Definition ww_Bm1 := WW w_Bm1 w_Bm1.
-
- Definition ww_WW xh xl : zn2z (zn2z w) :=
- match xh, xl with
- | W0, W0 => W0
- | _, _ => WW xh xl
- end.
-
- Definition ww_W0 h : zn2z (zn2z w) :=
- match h with
- | W0 => W0
- | _ => WW h W0
- end.
-
- Definition ww_0W l : zn2z (zn2z w) :=
- match l with
- | W0 => W0
- | _ => WW W0 l
- end.
-
- Definition double_WW (n:nat) :=
- match n return word w n -> word w n -> word w (S n) with
- | O => w_WW
- | S n =>
- fun (h l : zn2z (word w n)) =>
- match h, l with
- | W0, W0 => W0
- | _, _ => WW h l
- end
- end.
-
- Definition double_wB n := base (w_digits << n).
-
- Fixpoint double_to_Z (n:nat) : word w n -> Z :=
- match n return word w n -> Z with
- | O => w_to_Z
- | S n => zn2z_to_Z (double_wB n) (double_to_Z n)
- end.
-
- Fixpoint extend_aux (n:nat) (x:zn2z w) {struct n}: word w (S n) :=
- match n return word w (S n) with
- | O => x
- | S n1 => WW W0 (extend_aux n1 x)
- end.
-
- Definition extend (n:nat) (x:w) : word w (S n) :=
- let r := w_0W x in
- match r with
- | W0 => W0
- | _ => extend_aux n r
- end.
-
- Definition double_0 n : word w n :=
- match n return word w n with
- | O => w_0
- | S _ => W0
- end.
-
- Definition double_split (n:nat) (x:zn2z (word w n)) :=
- match x with
- | W0 =>
- match n return word w n * word w n with
- | O => (w_0,w_0)
- | S _ => (W0, W0)
- end
- | WW h l => (h,l)
- end.
-
- Definition ww_compare x y :=
- match x, y with
- | W0, W0 => Eq
- | W0, WW yh yl =>
- match w_compare w_0 yh with
- | Eq => w_compare w_0 yl
- | _ => Lt
- end
- | WW xh xl, W0 =>
- match w_compare xh w_0 with
- | Eq => w_compare xl w_0
- | _ => Gt
- end
- | WW xh xl, WW yh yl =>
- match w_compare xh yh with
- | Eq => w_compare xl yl
- | Lt => Lt
- | Gt => Gt
- end
- end.
-
-
- (* Return the low part of the composed word*)
- Fixpoint get_low (n : nat) {struct n}:
- word w n -> w :=
- match n return (word w n -> w) with
- | 0%nat => fun x => x
- | S n1 =>
- fun x =>
- match x with
- | W0 => w_0
- | WW _ x1 => get_low n1 x1
- end
- end.
-
-
- Section DoubleProof.
- Notation wB := (base w_digits).
- Notation wwB := (base ww_digits).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[[ x ]]" := (ww_to_Z x) (at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB ww_to_Z c) (at level 0, c at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB ww_to_Z c) (at level 0, c at level 99).
- Notation "[! n | x !]" := (double_to_Z n x) (at level 0, x at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_1 : [|w_1|] = 1.
- Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1.
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_w_compare : forall x y,
- w_compare x y = Z.compare [|x|] [|y|].
-
- Lemma wwB_wBwB : wwB = wB^2.
- Proof.
- unfold base, ww_digits;rewrite Z.pow_2_r; rewrite (Pos2Z.inj_xO w_digits).
- replace (2 * Zpos w_digits) with (Zpos w_digits + Zpos w_digits).
- apply Zpower_exp; unfold Z.ge;simpl;intros;discriminate.
- ring.
- Qed.
-
- Lemma spec_ww_1 : [[ww_1]] = 1.
- Proof. simpl;rewrite spec_w_0;rewrite spec_w_1;ring. Qed.
-
- Lemma spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1.
- Proof. simpl;rewrite spec_w_Bm1;rewrite wwB_wBwB;ring. Qed.
-
- Lemma lt_0_wB : 0 < wB.
- Proof.
- unfold base;apply Z.pow_pos_nonneg. unfold Z.lt;reflexivity.
- unfold Z.le;intros H;discriminate H.
- Qed.
-
- Lemma lt_0_wwB : 0 < wwB.
- Proof. rewrite wwB_wBwB; rewrite Z.pow_2_r; apply Z.mul_pos_pos;apply lt_0_wB. Qed.
-
- Lemma wB_pos: 1 < wB.
- Proof.
- unfold base;apply Z.lt_le_trans with (2^1). unfold Z.lt;reflexivity.
- apply Zpower_le_monotone. unfold Z.lt;reflexivity.
- split;unfold Z.le;intros H. discriminate H.
- clear spec_w_0W w_0W spec_w_Bm1 spec_to_Z spec_w_WW w_WW.
- destruct w_digits; discriminate H.
- Qed.
-
- Lemma wwB_pos: 1 < wwB.
- Proof.
- assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Z.mul_1_r 1).
- rewrite Z.pow_2_r.
- apply Zmult_lt_compat2;(split;[unfold Z.lt;reflexivity|trivial]).
- apply Z.lt_le_incl;trivial.
- Qed.
-
- Theorem wB_div_2: 2 * (wB / 2) = wB.
- Proof.
- clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W
- spec_to_Z;unfold base.
- assert (2 ^ Zpos w_digits = 2 * (2 ^ (Zpos w_digits - 1))).
- pattern 2 at 2; rewrite <- Z.pow_1_r.
- rewrite <- Zpower_exp; auto with zarith.
- f_equal; auto with zarith.
- case w_digits; compute; intros; discriminate.
- rewrite H; f_equal; auto with zarith.
- rewrite Z.mul_comm; apply Z_div_mult; auto with zarith.
- Qed.
-
- Theorem wwB_div_2 : wwB / 2 = wB / 2 * wB.
- Proof.
- clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W
- spec_to_Z.
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- pattern wB at 1; rewrite <- wB_div_2; auto.
- rewrite <- Z.mul_assoc.
- repeat (rewrite (Z.mul_comm 2); rewrite Z_div_mult); auto with zarith.
- Qed.
-
- Lemma mod_wwB : forall z x,
- (z*wB + [|x|]) mod wwB = (z mod wB)*wB + [|x|].
- Proof.
- intros z x.
- rewrite Zplus_mod.
- pattern wwB at 1;rewrite wwB_wBwB; rewrite Z.pow_2_r.
- rewrite Zmult_mod_distr_r;try apply lt_0_wB.
- rewrite (Zmod_small [|x|]).
- apply Zmod_small;rewrite wwB_wBwB;apply beta_mult;try apply spec_to_Z.
- apply Z_mod_lt;apply Z.lt_gt;apply lt_0_wB.
- destruct (spec_to_Z x);split;trivial.
- change [|x|] with (0*wB+[|x|]). rewrite wwB_wBwB.
- rewrite Z.pow_2_r;rewrite <- (Z.add_0_r (wB*wB));apply beta_lex_inv.
- apply lt_0_wB. apply spec_to_Z. split;[apply Z.le_refl | apply lt_0_wB].
- Qed.
-
- Lemma wB_div : forall x y, ([|x|] * wB + [|y|]) / wB = [|x|].
- Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- intros x y;unfold base;rewrite Zdiv_shift_r;auto with zarith.
- rewrite Z_div_mult;auto with zarith.
- destruct (spec_to_Z x);trivial.
- Qed.
-
- Lemma wB_div_plus : forall x y p,
- 0 <= p ->
- ([|x|]*wB + [|y|]) / 2^(Zpos w_digits + p) = [|x|] / 2^p.
- Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- intros x y p Hp;rewrite Zpower_exp;auto with zarith.
- rewrite <- Zdiv_Zdiv;auto with zarith.
- rewrite wB_div;trivial.
- Qed.
-
- Lemma lt_wB_wwB : wB < wwB.
- Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- unfold base;apply Zpower_lt_monotone;auto with zarith.
- assert (0 < Zpos w_digits). compute;reflexivity.
- unfold ww_digits;rewrite Pos2Z.inj_xO;auto with zarith.
- Qed.
-
- Lemma w_to_Z_wwB : forall x, x < wB -> x < wwB.
- Proof.
- intros x H;apply Z.lt_trans with wB;trivial;apply lt_wB_wwB.
- Qed.
-
- Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB.
- Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- destruct x as [ |h l];simpl.
- split;[apply Z.le_refl|apply lt_0_wwB].
- assert (H:=spec_to_Z h);assert (L:=spec_to_Z l);split.
- apply Z.add_nonneg_nonneg;auto with zarith.
- rewrite <- (Z.add_0_r wwB);rewrite wwB_wBwB; rewrite Z.pow_2_r;
- apply beta_lex_inv;auto with zarith.
- Qed.
-
- Lemma double_wB_wwB : forall n, double_wB n * double_wB n = double_wB (S n).
- Proof.
- intros n;unfold double_wB;simpl.
- unfold base. rewrite (Pos2Z.inj_xO (_ << _)).
- replace (2 * Zpos (w_digits << n)) with
- (Zpos (w_digits << n) + Zpos (w_digits << n)) by ring.
- symmetry; apply Zpower_exp;intro;discriminate.
- Qed.
-
- Lemma double_wB_pos:
- forall n, 0 <= double_wB n.
- Proof.
- intros n; unfold double_wB, base; auto with zarith.
- Qed.
-
- Lemma double_wB_more_digits:
- forall n, wB <= double_wB n.
- Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- intros n; elim n; clear n; auto.
- unfold double_wB, "<<"; auto with zarith.
- intros n H1; rewrite <- double_wB_wwB.
- apply Z.le_trans with (wB * 1).
- rewrite Z.mul_1_r; apply Z.le_refl.
- unfold base; auto with zarith.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- apply Z.le_trans with wB; auto with zarith.
- unfold base.
- rewrite <- (Z.pow_0_r 2).
- apply Z.pow_le_mono_r; auto with zarith.
- Qed.
-
- Lemma spec_double_to_Z :
- forall n (x:word w n), 0 <= [!n | x!] < double_wB n.
- Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- induction n;intros. exact (spec_to_Z x).
- unfold double_to_Z;fold double_to_Z.
- destruct x;unfold zn2z_to_Z.
- unfold double_wB,base;split;auto with zarith.
- assert (U0:= IHn w0);assert (U1:= IHn w1).
- split;auto with zarith.
- apply Z.lt_le_trans with ((double_wB n - 1) * double_wB n + double_wB n).
- assert (double_to_Z n w0*double_wB n <= (double_wB n - 1)*double_wB n).
- apply Z.mul_le_mono_nonneg_r;auto with zarith.
- auto with zarith.
- rewrite <- double_wB_wwB.
- replace ((double_wB n - 1) * double_wB n + double_wB n) with (double_wB n * double_wB n);
- [auto with zarith | ring].
- Qed.
-
- Lemma spec_get_low:
- forall n x,
- [!n | x!] < wB -> [|get_low n x|] = [!n | x!].
- Proof.
- clear spec_w_1 spec_w_Bm1.
- intros n; elim n; auto; clear n.
- intros n Hrec x; case x; clear x; auto.
- intros xx yy; simpl.
- destruct (spec_double_to_Z n xx) as [F1 _]. Z.le_elim F1.
- - (* 0 < [!n | xx!] *)
- intros; exfalso.
- assert (F3 := double_wB_more_digits n).
- destruct (spec_double_to_Z n yy) as [F4 _].
- assert (F5: 1 * wB <= [!n | xx!] * double_wB n);
- auto with zarith.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- unfold base; auto with zarith.
- - (* 0 = [!n | xx!] *)
- rewrite <- F1; rewrite Z.mul_0_l, Z.add_0_l.
- intros; apply Hrec; auto.
- Qed.
-
- Lemma spec_double_WW : forall n (h l : word w n),
- [!S n|double_WW n h l!] = [!n|h!] * double_wB n + [!n|l!].
- Proof.
- induction n;simpl;intros;trivial.
- destruct h;auto.
- destruct l;auto.
- Qed.
-
- Lemma spec_extend_aux : forall n x, [!S n|extend_aux n x!] = [[x]].
- Proof. induction n;simpl;trivial. Qed.
-
- Lemma spec_extend : forall n x, [!S n|extend n x!] = [|x|].
- Proof.
- intros n x;assert (H:= spec_w_0W x);unfold extend.
- destruct (w_0W x);simpl;trivial.
- rewrite <- H;exact (spec_extend_aux n (WW w0 w1)).
- Qed.
-
- Lemma spec_double_0 : forall n, [!n|double_0 n!] = 0.
- Proof. destruct n;trivial. Qed.
-
- Lemma spec_double_split : forall n x,
- let (h,l) := double_split n x in
- [!S n|x!] = [!n|h!] * double_wB n + [!n|l!].
- Proof.
- destruct x;simpl;auto.
- destruct n;simpl;trivial.
- rewrite spec_w_0;trivial.
- Qed.
-
- Lemma wB_lex_inv: forall a b c d,
- a < c ->
- a * wB + [|b|] < c * wB + [|d|].
- Proof.
- intros a b c d H1; apply beta_lex_inv with (1 := H1); auto.
- Qed.
-
- Ltac comp2ord := match goal with
- | |- Lt = (?x ?= ?y) => symmetry; change (x < y)
- | |- Gt = (?x ?= ?y) => symmetry; change (x > y); apply Z.lt_gt
- end.
-
- Lemma spec_ww_compare : forall x y,
- ww_compare x y = Z.compare [[x]] [[y]].
- Proof.
- destruct x as [ |xh xl];destruct y as [ |yh yl];simpl;trivial.
- (* 1st case *)
- rewrite 2 spec_w_compare, spec_w_0.
- destruct (Z.compare_spec 0 [|yh|]) as [H|H|H].
- rewrite <- H;simpl. reflexivity.
- symmetry. change (0 < [|yh|]*wB+[|yl|]).
- change 0 with (0*wB+0). rewrite <- spec_w_0 at 2.
- apply wB_lex_inv;trivial.
- absurd (0 <= [|yh|]). apply Z.lt_nge; trivial.
- destruct (spec_to_Z yh);trivial.
- (* 2nd case *)
- rewrite 2 spec_w_compare, spec_w_0.
- destruct (Z.compare_spec [|xh|] 0) as [H|H|H].
- rewrite H;simpl;reflexivity.
- absurd (0 <= [|xh|]). apply Z.lt_nge; trivial.
- destruct (spec_to_Z xh);trivial.
- comp2ord.
- change 0 with (0*wB+0). rewrite <- spec_w_0 at 2.
- apply wB_lex_inv;trivial.
- (* 3rd case *)
- rewrite 2 spec_w_compare.
- destruct (Z.compare_spec [|xh|] [|yh|]) as [H|H|H].
- rewrite H.
- symmetry. apply Z.add_compare_mono_l.
- comp2ord. apply wB_lex_inv;trivial.
- comp2ord. apply wB_lex_inv;trivial.
- Qed.
-
-
- End DoubleProof.
-
-End DoubleBase.
-
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
deleted file mode 100644
index 4ebe8fac1..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
+++ /dev/null
@@ -1,966 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-Require Import DoubleAdd.
-Require Import DoubleSub.
-Require Import DoubleMul.
-Require Import DoubleSqrt.
-Require Import DoubleLift.
-Require Import DoubleDivn1.
-Require Import DoubleDiv.
-Require Import CyclicAxioms.
-
-Local Open Scope Z_scope.
-
-
-Section Z_2nZ.
-
- Context {t : Type}{ops : ZnZ.Ops t}.
-
- Let w_digits := ZnZ.digits.
- Let w_zdigits := ZnZ.zdigits.
-
- Let w_to_Z := ZnZ.to_Z.
- Let w_of_pos := ZnZ.of_pos.
- Let w_head0 := ZnZ.head0.
- Let w_tail0 := ZnZ.tail0.
-
- Let w_0 := ZnZ.zero.
- Let w_1 := ZnZ.one.
- Let w_Bm1 := ZnZ.minus_one.
-
- Let w_compare := ZnZ.compare.
- Let w_eq0 := ZnZ.eq0.
-
- Let w_opp_c := ZnZ.opp_c.
- Let w_opp := ZnZ.opp.
- Let w_opp_carry := ZnZ.opp_carry.
-
- Let w_succ_c := ZnZ.succ_c.
- Let w_add_c := ZnZ.add_c.
- Let w_add_carry_c := ZnZ.add_carry_c.
- Let w_succ := ZnZ.succ.
- Let w_add := ZnZ.add.
- Let w_add_carry := ZnZ.add_carry.
-
- Let w_pred_c := ZnZ.pred_c.
- Let w_sub_c := ZnZ.sub_c.
- Let w_sub_carry_c := ZnZ.sub_carry_c.
- Let w_pred := ZnZ.pred.
- Let w_sub := ZnZ.sub.
- Let w_sub_carry := ZnZ.sub_carry.
-
-
- Let w_mul_c := ZnZ.mul_c.
- Let w_mul := ZnZ.mul.
- Let w_square_c := ZnZ.square_c.
-
- Let w_div21 := ZnZ.div21.
- Let w_div_gt := ZnZ.div_gt.
- Let w_div := ZnZ.div.
-
- Let w_mod_gt := ZnZ.modulo_gt.
- Let w_mod := ZnZ.modulo.
-
- Let w_gcd_gt := ZnZ.gcd_gt.
- Let w_gcd := ZnZ.gcd.
-
- Let w_add_mul_div := ZnZ.add_mul_div.
-
- Let w_pos_mod := ZnZ.pos_mod.
-
- Let w_is_even := ZnZ.is_even.
- Let w_sqrt2 := ZnZ.sqrt2.
- Let w_sqrt := ZnZ.sqrt.
-
- Let _zn2z := zn2z t.
-
- Let wB := base w_digits.
-
- Let w_Bm2 := w_pred w_Bm1.
-
- Let ww_1 := ww_1 w_0 w_1.
- Let ww_Bm1 := ww_Bm1 w_Bm1.
-
- Let w_add2 a b := match w_add_c a b with C0 p => WW w_0 p | C1 p => WW w_1 p end.
-
- Let _ww_digits := xO w_digits.
-
- Let _ww_zdigits := w_add2 w_zdigits w_zdigits.
-
- Let to_Z := zn2z_to_Z wB w_to_Z.
-
- Let w_W0 := ZnZ.WO.
- Let w_0W := ZnZ.OW.
- Let w_WW := ZnZ.WW.
-
- Let ww_of_pos p :=
- match w_of_pos p with
- | (N0, l) => (N0, WW w_0 l)
- | (Npos ph,l) =>
- let (n,h) := w_of_pos ph in (n, w_WW h l)
- end.
-
- Let head0 :=
- Eval lazy beta delta [ww_head0] in
- ww_head0 w_0 w_0W w_compare w_head0 w_add2 w_zdigits _ww_zdigits.
-
- Let tail0 :=
- Eval lazy beta delta [ww_tail0] in
- ww_tail0 w_0 w_0W w_compare w_tail0 w_add2 w_zdigits _ww_zdigits.
-
- Let ww_WW := Eval lazy beta delta [ww_WW] in (@ww_WW t).
- Let ww_0W := Eval lazy beta delta [ww_0W] in (@ww_0W t).
- Let ww_W0 := Eval lazy beta delta [ww_W0] in (@ww_W0 t).
-
- (* ** Comparison ** *)
- Let compare :=
- Eval lazy beta delta[ww_compare] in ww_compare w_0 w_compare.
-
- Let eq0 (x:zn2z t) :=
- match x with
- | W0 => true
- | _ => false
- end.
-
- (* ** Opposites ** *)
- Let opp_c :=
- Eval lazy beta delta [ww_opp_c] in ww_opp_c w_0 w_opp_c w_opp_carry.
-
- Let opp :=
- Eval lazy beta delta [ww_opp] in ww_opp w_0 w_opp_c w_opp_carry w_opp.
-
- Let opp_carry :=
- Eval lazy beta delta [ww_opp_carry] in ww_opp_carry w_WW ww_Bm1 w_opp_carry.
-
- (* ** Additions ** *)
-
- Let succ_c :=
- Eval lazy beta delta [ww_succ_c] in ww_succ_c w_0 ww_1 w_succ_c.
-
- Let add_c :=
- Eval lazy beta delta [ww_add_c] in ww_add_c w_WW w_add_c w_add_carry_c.
-
- Let add_carry_c :=
- Eval lazy beta iota delta [ww_add_carry_c ww_succ_c] in
- ww_add_carry_c w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c.
-
- Let succ :=
- Eval lazy beta delta [ww_succ] in ww_succ w_W0 ww_1 w_succ_c w_succ.
-
- Let add :=
- Eval lazy beta delta [ww_add] in ww_add w_add_c w_add w_add_carry.
-
- Let add_carry :=
- Eval lazy beta iota delta [ww_add_carry ww_succ] in
- ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ w_add w_add_carry.
-
- (* ** Subtractions ** *)
-
- Let pred_c :=
- Eval lazy beta delta [ww_pred_c] in ww_pred_c w_Bm1 w_WW ww_Bm1 w_pred_c.
-
- Let sub_c :=
- Eval lazy beta iota delta [ww_sub_c ww_opp_c] in
- ww_sub_c w_0 w_WW w_opp_c w_opp_carry w_sub_c w_sub_carry_c.
-
- Let sub_carry_c :=
- Eval lazy beta iota delta [ww_sub_carry_c ww_pred_c ww_opp_carry] in
- ww_sub_carry_c w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_c w_sub_carry_c.
-
- Let pred :=
- Eval lazy beta delta [ww_pred] in ww_pred w_Bm1 w_WW ww_Bm1 w_pred_c w_pred.
-
- Let sub :=
- Eval lazy beta iota delta [ww_sub ww_opp] in
- ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry.
-
- Let sub_carry :=
- Eval lazy beta iota delta [ww_sub_carry ww_pred ww_opp_carry] in
- ww_sub_carry w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_carry_c w_pred
- w_sub w_sub_carry.
-
-
- (* ** Multiplication ** *)
-
- Let mul_c :=
- Eval lazy beta iota delta [ww_mul_c double_mul_c] in
- ww_mul_c w_0 w_1 w_WW w_W0 w_mul_c add_c add add_carry.
-
- Let karatsuba_c :=
- Eval lazy beta iota delta [ww_karatsuba_c double_mul_c kara_prod] in
- ww_karatsuba_c w_0 w_1 w_WW w_W0 w_compare w_add w_sub w_mul_c
- add_c add add_carry sub_c sub.
-
- Let mul :=
- Eval lazy beta delta [ww_mul] in
- ww_mul w_W0 w_add w_mul_c w_mul add.
-
- Let square_c :=
- Eval lazy beta delta [ww_square_c] in
- ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add add_carry.
-
- (* Division operation *)
-
- Let div32 :=
- Eval lazy beta iota delta [w_div32] in
- w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c
- w_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c.
-
- Let div21 :=
- Eval lazy beta iota delta [ww_div21] in
- ww_div21 w_0 w_0W div32 ww_1 compare sub.
-
- Let low (p: zn2z t) := match p with WW _ p1 => p1 | _ => w_0 end.
-
- Let add_mul_div :=
- Eval lazy beta delta [ww_add_mul_div] in
- ww_add_mul_div w_0 w_WW w_W0 w_0W compare w_add_mul_div sub w_zdigits low.
-
- Let div_gt :=
- Eval lazy beta delta [ww_div_gt] in
- ww_div_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp
- w_opp_carry w_sub_c w_sub w_sub_carry
- w_div_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits.
-
- Let div :=
- Eval lazy beta delta [ww_div] in ww_div ww_1 compare div_gt.
-
- Let mod_gt :=
- Eval lazy beta delta [ww_mod_gt] in
- ww_mod_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry
- w_mod_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div w_zdigits.
-
- Let mod_ :=
- Eval lazy beta delta [ww_mod] in ww_mod compare mod_gt.
-
- Let pos_mod :=
- Eval lazy beta delta [ww_pos_mod] in
- ww_pos_mod w_0 w_zdigits w_WW w_pos_mod compare w_0W low sub _ww_zdigits.
-
- Let is_even :=
- Eval lazy beta delta [ww_is_even] in ww_is_even w_is_even.
-
- Let sqrt2 :=
- Eval lazy beta delta [ww_sqrt2] in
- ww_sqrt2 w_is_even w_compare w_0 w_1 w_Bm1 w_0W w_sub w_square_c
- w_div21 w_add_mul_div w_zdigits w_add_c w_sqrt2 w_pred pred_c
- pred add_c add sub_c add_mul_div.
-
- Let sqrt :=
- Eval lazy beta delta [ww_sqrt] in
- ww_sqrt w_is_even w_0 w_sub w_add_mul_div w_zdigits
- _ww_zdigits w_sqrt2 pred add_mul_div head0 compare low.
-
- Let gcd_gt_fix :=
- Eval cbv beta delta [ww_gcd_gt_aux ww_gcd_gt_body] in
- ww_gcd_gt_aux w_0 w_WW w_0W w_compare w_opp_c w_opp w_opp_carry
- w_sub_c w_sub w_sub_carry w_gcd_gt
- w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div
- w_zdigits.
-
- Let gcd_cont :=
- Eval lazy beta delta [gcd_cont] in gcd_cont ww_1 w_1 w_compare.
-
- Let gcd_gt :=
- Eval lazy beta delta [ww_gcd_gt] in
- ww_gcd_gt w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont.
-
- Let gcd :=
- Eval lazy beta delta [ww_gcd] in
- ww_gcd compare w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont.
-
- Definition lor (x y : zn2z t) :=
- match x, y with
- | W0, _ => y
- | _, W0 => x
- | WW hx lx, WW hy ly => WW (ZnZ.lor hx hy) (ZnZ.lor lx ly)
- end.
-
- Definition land (x y : zn2z t) :=
- match x, y with
- | W0, _ => W0
- | _, W0 => W0
- | WW hx lx, WW hy ly => WW (ZnZ.land hx hy) (ZnZ.land lx ly)
- end.
-
- Definition lxor (x y : zn2z t) :=
- match x, y with
- | W0, _ => y
- | _, W0 => x
- | WW hx lx, WW hy ly => WW (ZnZ.lxor hx hy) (ZnZ.lxor lx ly)
- end.
-
- (* ** Record of operators on 2 words *)
-
- Global Instance mk_zn2z_ops : ZnZ.Ops (zn2z t) | 1 :=
- ZnZ.MkOps _ww_digits _ww_zdigits
- to_Z ww_of_pos head0 tail0
- W0 ww_1 ww_Bm1
- compare eq0
- opp_c opp opp_carry
- succ_c add_c add_carry_c
- succ add add_carry
- pred_c sub_c sub_carry_c
- pred sub sub_carry
- mul_c mul square_c
- div21 div_gt div
- mod_gt mod_
- gcd_gt gcd
- add_mul_div
- pos_mod
- is_even
- sqrt2
- sqrt
- lor
- land
- lxor.
-
- Global Instance mk_zn2z_ops_karatsuba : ZnZ.Ops (zn2z t) | 2 :=
- ZnZ.MkOps _ww_digits _ww_zdigits
- to_Z ww_of_pos head0 tail0
- W0 ww_1 ww_Bm1
- compare eq0
- opp_c opp opp_carry
- succ_c add_c add_carry_c
- succ add add_carry
- pred_c sub_c sub_carry_c
- pred sub sub_carry
- karatsuba_c mul square_c
- div21 div_gt div
- mod_gt mod_
- gcd_gt gcd
- add_mul_div
- pos_mod
- is_even
- sqrt2
- sqrt
- lor
- land
- lxor.
-
- (* Proof *)
- Context {specs : ZnZ.Specs ops}.
-
- Create HintDb ZnZ.
-
- Hint Resolve
- ZnZ.spec_to_Z
- ZnZ.spec_of_pos
- ZnZ.spec_0
- ZnZ.spec_1
- ZnZ.spec_m1
- ZnZ.spec_compare
- ZnZ.spec_eq0
- ZnZ.spec_opp_c
- ZnZ.spec_opp
- ZnZ.spec_opp_carry
- ZnZ.spec_succ_c
- ZnZ.spec_add_c
- ZnZ.spec_add_carry_c
- ZnZ.spec_succ
- ZnZ.spec_add
- ZnZ.spec_add_carry
- ZnZ.spec_pred_c
- ZnZ.spec_sub_c
- ZnZ.spec_sub_carry_c
- ZnZ.spec_pred
- ZnZ.spec_sub
- ZnZ.spec_sub_carry
- ZnZ.spec_mul_c
- ZnZ.spec_mul
- ZnZ.spec_square_c
- ZnZ.spec_div21
- ZnZ.spec_div_gt
- ZnZ.spec_div
- ZnZ.spec_modulo_gt
- ZnZ.spec_modulo
- ZnZ.spec_gcd_gt
- ZnZ.spec_gcd
- ZnZ.spec_head0
- ZnZ.spec_tail0
- ZnZ.spec_add_mul_div
- ZnZ.spec_pos_mod
- ZnZ.spec_is_even
- ZnZ.spec_sqrt2
- ZnZ.spec_sqrt
- ZnZ.spec_WO
- ZnZ.spec_OW
- ZnZ.spec_WW : ZnZ.
-
- Ltac wwauto := unfold ww_to_Z; eauto with ZnZ.
-
- Let wwB := base _ww_digits.
-
- Notation "[| x |]" := (to_Z x) (at level 0, x at level 99).
-
- Notation "[+| c |]" :=
- (interp_carry 1 wwB to_Z c) (at level 0, c at level 99).
-
- Notation "[-| c |]" :=
- (interp_carry (-1) wwB to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (zn2z_to_Z wwB to_Z x) (at level 0, x at level 99).
-
- Let spec_ww_to_Z : forall x, 0 <= [| x |] < wwB.
- Proof. refine (spec_ww_to_Z w_digits w_to_Z _); wwauto. Qed.
-
- Let spec_ww_of_pos : forall p,
- Zpos p = (Z.of_N (fst (ww_of_pos p)))*wwB + [|(snd (ww_of_pos p))|].
- Proof.
- unfold ww_of_pos;intros.
- rewrite (ZnZ.spec_of_pos p). unfold w_of_pos.
- case (ZnZ.of_pos p); intros. simpl.
- destruct n; simpl ZnZ.to_Z.
- simpl;unfold w_to_Z,w_0; rewrite ZnZ.spec_0;trivial.
- unfold Z.of_N.
- rewrite (ZnZ.spec_of_pos p0).
- case (ZnZ.of_pos p0); intros. simpl.
- unfold fst, snd,Z.of_N, to_Z, wB, w_digits, w_to_Z, w_WW.
- rewrite ZnZ.spec_WW.
- replace wwB with (wB*wB).
- unfold wB,w_to_Z,w_digits;destruct n;ring.
- symmetry. rewrite <- Z.pow_2_r; exact (wwB_wBwB w_digits).
- Qed.
-
- Let spec_ww_0 : [|W0|] = 0.
- Proof. reflexivity. Qed.
-
- Let spec_ww_1 : [|ww_1|] = 1.
- Proof. refine (spec_ww_1 w_0 w_1 w_digits w_to_Z _ _);wwauto. Qed.
-
- Let spec_ww_Bm1 : [|ww_Bm1|] = wwB - 1.
- Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);wwauto. Qed.
-
- Let spec_ww_compare :
- forall x y, compare x y = Z.compare [|x|] [|y|].
- Proof.
- refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);wwauto.
- Qed.
-
- Let spec_ww_eq0 : forall x, eq0 x = true -> [|x|] = 0.
- Proof. destruct x;simpl;intros;trivial;discriminate. Qed.
-
- Let spec_ww_opp_c : forall x, [-|opp_c x|] = -[|x|].
- Proof.
- refine(spec_ww_opp_c w_0 w_0 W0 w_opp_c w_opp_carry w_digits w_to_Z _ _ _ _);
- wwauto.
- Qed.
-
- Let spec_ww_opp : forall x, [|opp x|] = (-[|x|]) mod wwB.
- Proof.
- refine(spec_ww_opp w_0 w_0 W0 w_opp_c w_opp_carry w_opp
- w_digits w_to_Z _ _ _ _ _);
- wwauto.
- Qed.
-
- Let spec_ww_opp_carry : forall x, [|opp_carry x|] = wwB - [|x|] - 1.
- Proof.
- refine (spec_ww_opp_carry w_WW ww_Bm1 w_opp_carry w_digits w_to_Z _ _ _);
- wwauto.
- Qed.
-
- Let spec_ww_succ_c : forall x, [+|succ_c x|] = [|x|] + 1.
- Proof.
- refine (spec_ww_succ_c w_0 w_0 ww_1 w_succ_c w_digits w_to_Z _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_add_c : forall x y, [+|add_c x y|] = [|x|] + [|y|].
- Proof.
- refine (spec_ww_add_c w_WW w_add_c w_add_carry_c w_digits w_to_Z _ _ _);wwauto.
- Qed.
-
- Let spec_ww_add_carry_c : forall x y, [+|add_carry_c x y|] = [|x|]+[|y|]+1.
- Proof.
- refine (spec_ww_add_carry_c w_0 w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c
- w_digits w_to_Z _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_succ : forall x, [|succ x|] = ([|x|] + 1) mod wwB.
- Proof.
- refine (spec_ww_succ w_W0 ww_1 w_succ_c w_succ w_digits w_to_Z _ _ _ _ _);
- wwauto.
- Qed.
-
- Let spec_ww_add : forall x y, [|add x y|] = ([|x|] + [|y|]) mod wwB.
- Proof.
- refine (spec_ww_add w_add_c w_add w_add_carry w_digits w_to_Z _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_add_carry : forall x y, [|add_carry x y|]=([|x|]+[|y|]+1)mod wwB.
- Proof.
- refine (spec_ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ
- w_add w_add_carry w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_pred_c : forall x, [-|pred_c x|] = [|x|] - 1.
- Proof.
- refine (spec_ww_pred_c w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_digits w_to_Z
- _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_sub_c : forall x y, [-|sub_c x y|] = [|x|] - [|y|].
- Proof.
- refine (spec_ww_sub_c w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c
- w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_sub_carry_c : forall x y, [-|sub_carry_c x y|] = [|x|]-[|y|]-1.
- Proof.
- refine (spec_ww_sub_carry_c w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c
- w_sub_c w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_pred : forall x, [|pred x|] = ([|x|] - 1) mod wwB.
- Proof.
- refine (spec_ww_pred w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_pred w_digits w_to_Z
- _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_sub : forall x y, [|sub x y|] = ([|x|] - [|y|]) mod wwB.
- Proof.
- refine (spec_ww_sub w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c w_opp
- w_sub w_sub_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_sub_carry : forall x y, [|sub_carry x y|]=([|x|]-[|y|]-1) mod wwB.
- Proof.
- refine (spec_ww_sub_carry w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c
- w_sub_carry_c w_pred w_sub w_sub_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _);
- wwauto.
- Qed.
-
- Let spec_ww_mul_c : forall x y, [[mul_c x y ]] = [|x|] * [|y|].
- Proof.
- refine (spec_ww_mul_c w_0 w_1 w_WW w_W0 w_mul_c add_c add add_carry w_digits
- w_to_Z _ _ _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_karatsuba_c : forall x y, [[karatsuba_c x y ]] = [|x|] * [|y|].
- Proof.
- refine (spec_ww_karatsuba_c _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
- _ _ _ _ _ _ _ _ _ _ _ _); wwauto.
- unfold w_digits; apply ZnZ.spec_more_than_1_digit; auto.
- Qed.
-
- Let spec_ww_mul : forall x y, [|mul x y|] = ([|x|] * [|y|]) mod wwB.
- Proof.
- refine (spec_ww_mul w_W0 w_add w_mul_c w_mul add w_digits w_to_Z _ _ _ _ _);
- wwauto.
- Qed.
-
- Let spec_ww_square_c : forall x, [[square_c x]] = [|x|] * [|x|].
- Proof.
- refine (spec_ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add
- add_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_w_div32 : forall a1 a2 a3 b1 b2,
- wB / 2 <= (w_to_Z b1) ->
- [|WW a1 a2|] < [|WW b1 b2|] ->
- let (q, r) := div32 a1 a2 a3 b1 b2 in
- (w_to_Z a1) * wwB + (w_to_Z a2) * wB + (w_to_Z a3) =
- (w_to_Z q) * ((w_to_Z b1)*wB + (w_to_Z b2)) + [|r|] /\
- 0 <= [|r|] < (w_to_Z b1)*wB + w_to_Z b2.
- Proof.
- refine (spec_w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c
- w_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c w_digits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
- unfold w_Bm2, w_to_Z, w_pred, w_Bm1.
- rewrite ZnZ.spec_pred, ZnZ.spec_m1.
- unfold w_digits;rewrite Zmod_small. ring.
- assert (H:= wB_pos(ZnZ.digits)). omega.
- exact ZnZ.spec_div21.
- Qed.
-
- Let spec_ww_div21 : forall a1 a2 b,
- wwB/2 <= [|b|] ->
- [|a1|] < [|b|] ->
- let (q,r) := div21 a1 a2 b in
- [|a1|] *wwB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
- Proof.
- refine (spec_ww_div21 w_0 w_0W div32 ww_1 compare sub w_digits w_to_Z
- _ _ _ _ _ _ _);wwauto.
- Qed.
-
- Let spec_add2: forall x y,
- [|w_add2 x y|] = w_to_Z x + w_to_Z y.
- unfold w_add2.
- intros xh xl; generalize (ZnZ.spec_add_c xh xl).
- unfold w_add_c; case ZnZ.add_c; unfold interp_carry; simpl ww_to_Z.
- intros w0 Hw0; simpl; unfold w_to_Z; rewrite Hw0.
- unfold w_0; rewrite ZnZ.spec_0; simpl; auto with zarith.
- intros w0; rewrite Z.mul_1_l; simpl.
- unfold w_to_Z, w_1; rewrite ZnZ.spec_1; auto with zarith.
- rewrite Z.mul_1_l; auto.
- Qed.
-
- Let spec_low: forall x,
- w_to_Z (low x) = [|x|] mod wB.
- intros x; case x; simpl low.
- unfold ww_to_Z, w_to_Z, w_0; rewrite ZnZ.spec_0; simpl; wwauto.
- intros xh xl; simpl.
- rewrite Z.add_comm; rewrite Z_mod_plus; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- unfold wB, base; eauto with ZnZ zarith.
- unfold wB, base; eauto with ZnZ zarith.
- Qed.
-
- Let spec_ww_digits:
- [|_ww_zdigits|] = Zpos (xO w_digits).
- Proof.
- unfold w_to_Z, _ww_zdigits.
- rewrite spec_add2.
- unfold w_to_Z, w_zdigits, w_digits.
- rewrite ZnZ.spec_zdigits; auto.
- rewrite Pos2Z.inj_xO; auto with zarith.
- Qed.
-
-
- Let spec_ww_head00 : forall x, [|x|] = 0 -> [|head0 x|] = Zpos _ww_digits.
- Proof.
- refine (spec_ww_head00 w_0 w_0W
- w_compare w_head0 w_add2 w_zdigits _ww_zdigits
- w_to_Z _ _ _ (eq_refl _ww_digits) _ _ _ _); wwauto.
- exact ZnZ.spec_head00.
- exact ZnZ.spec_zdigits.
- Qed.
-
- Let spec_ww_head0 : forall x, 0 < [|x|] ->
- wwB/ 2 <= 2 ^ [|head0 x|] * [|x|] < wwB.
- Proof.
- refine (spec_ww_head0 w_0 w_0W w_compare w_head0
- w_add2 w_zdigits _ww_zdigits
- w_to_Z _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_zdigits.
- Qed.
-
- Let spec_ww_tail00 : forall x, [|x|] = 0 -> [|tail0 x|] = Zpos _ww_digits.
- Proof.
- refine (spec_ww_tail00 w_0 w_0W
- w_compare w_tail0 w_add2 w_zdigits _ww_zdigits
- w_to_Z _ _ _ (eq_refl _ww_digits) _ _ _ _); wwauto.
- exact ZnZ.spec_tail00.
- exact ZnZ.spec_zdigits.
- Qed.
-
-
- Let spec_ww_tail0 : forall x, 0 < [|x|] ->
- exists y, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail0 x|].
- Proof.
- refine (spec_ww_tail0 (w_digits := w_digits) w_0 w_0W w_compare w_tail0
- w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_zdigits.
- Qed.
-
- Lemma spec_ww_add_mul_div : forall x y p,
- [|p|] <= Zpos _ww_digits ->
- [| add_mul_div p x y |] =
- ([|x|] * (2 ^ [|p|]) +
- [|y|] / (2 ^ ((Zpos _ww_digits) - [|p|]))) mod wwB.
- Proof.
- refine (@spec_ww_add_mul_div t w_0 w_WW w_W0 w_0W compare w_add_mul_div
- sub w_digits w_zdigits low w_to_Z
- _ _ _ _ _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_zdigits.
- Qed.
-
- Let spec_ww_div_gt : forall a b,
- [|a|] > [|b|] -> 0 < [|b|] ->
- let (q,r) := div_gt a b in
- [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|].
- Proof.
-refine
-(@spec_ww_div_gt t w_digits w_0 w_WW w_0W w_compare w_eq0
- w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt
- w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
-).
- exact ZnZ.spec_0.
- exact ZnZ.spec_to_Z.
- wwauto.
- wwauto.
- exact ZnZ.spec_compare.
- exact ZnZ.spec_eq0.
- exact ZnZ.spec_opp_c.
- exact ZnZ.spec_opp.
- exact ZnZ.spec_opp_carry.
- exact ZnZ.spec_sub_c.
- exact ZnZ.spec_sub.
- exact ZnZ.spec_sub_carry.
- exact ZnZ.spec_div_gt.
- exact ZnZ.spec_add_mul_div.
- exact ZnZ.spec_head0.
- exact ZnZ.spec_div21.
- exact spec_w_div32.
- exact ZnZ.spec_zdigits.
- exact spec_ww_digits.
- exact spec_ww_1.
- exact spec_ww_add_mul_div.
- Qed.
-
- Let spec_ww_div : forall a b, 0 < [|b|] ->
- let (q,r) := div a b in
- [|a|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
- Proof.
- refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);wwauto.
- Qed.
-
- Let spec_ww_mod_gt : forall a b,
- [|a|] > [|b|] -> 0 < [|b|] ->
- [|mod_gt a b|] = [|a|] mod [|b|].
- Proof.
- refine (@spec_ww_mod_gt t w_digits w_0 w_WW w_0W w_compare w_eq0
- w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt w_mod_gt
- w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div
- w_zdigits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_div_gt.
- exact ZnZ.spec_div21.
- exact ZnZ.spec_zdigits.
- exact spec_ww_add_mul_div.
- Qed.
-
- Let spec_ww_mod : forall a b, 0 < [|b|] -> [|mod_ a b|] = [|a|] mod [|b|].
- Proof.
- refine (spec_ww_mod w_digits W0 compare mod_gt w_to_Z _ _ _);wwauto.
- Qed.
-
- Let spec_ww_gcd_gt : forall a b, [|a|] > [|b|] ->
- Zis_gcd [|a|] [|b|] [|gcd_gt a b|].
- Proof.
- refine (@spec_ww_gcd_gt t w_digits W0 w_to_Z _
- w_0 w_0 w_eq0 w_gcd_gt _ww_digits
- _ gcd_gt_fix _ _ _ _ gcd_cont _);wwauto.
- refine (@spec_ww_gcd_gt_aux t w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp
- w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0
- w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_div21.
- exact ZnZ.spec_zdigits.
- exact spec_ww_add_mul_div.
- refine (@spec_gcd_cont t w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare
- _ _);wwauto.
- Qed.
-
- Let spec_ww_gcd : forall a b, Zis_gcd [|a|] [|b|] [|gcd a b|].
- Proof.
- refine (@spec_ww_gcd t w_digits W0 compare w_to_Z _ _ w_0 w_0 w_eq0 w_gcd_gt
- _ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);wwauto.
- refine (@spec_ww_gcd_gt_aux t w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp
- w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0
- w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_div21.
- exact ZnZ.spec_zdigits.
- exact spec_ww_add_mul_div.
- refine (@spec_gcd_cont t w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare
- _ _);wwauto.
- Qed.
-
- Let spec_ww_is_even : forall x,
- match is_even x with
- true => [|x|] mod 2 = 0
- | false => [|x|] mod 2 = 1
- end.
- Proof.
- refine (@spec_ww_is_even t w_is_even w_digits _ _ ).
- exact ZnZ.spec_is_even.
- Qed.
-
- Let spec_ww_sqrt2 : forall x y,
- wwB/ 4 <= [|x|] ->
- let (s,r) := sqrt2 x y in
- [[WW x y]] = [|s|] ^ 2 + [+|r|] /\
- [+|r|] <= 2 * [|s|].
- Proof.
- intros x y H.
- refine (@spec_ww_sqrt2 t w_is_even w_compare w_0 w_1 w_Bm1
- w_0W w_sub w_square_c w_div21 w_add_mul_div w_digits w_zdigits
- _ww_zdigits
- w_add_c w_sqrt2 w_pred pred_c pred add_c add sub_c add_mul_div
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto.
- exact ZnZ.spec_zdigits.
- exact ZnZ.spec_more_than_1_digit.
- exact ZnZ.spec_is_even.
- exact ZnZ.spec_div21.
- exact spec_ww_add_mul_div.
- exact ZnZ.spec_sqrt2.
- Qed.
-
- Let spec_ww_sqrt : forall x,
- [|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2.
- Proof.
- refine (@spec_ww_sqrt t w_is_even w_0 w_1 w_Bm1
- w_sub w_add_mul_div w_digits w_zdigits _ww_zdigits
- w_sqrt2 pred add_mul_div head0 compare
- _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto.
- exact ZnZ.spec_zdigits.
- exact ZnZ.spec_more_than_1_digit.
- exact ZnZ.spec_is_even.
- exact spec_ww_add_mul_div.
- exact ZnZ.spec_sqrt2.
- Qed.
-
- Let wB_pos : 0 < wB.
- Proof.
- unfold wB, base; apply Z.pow_pos_nonneg; auto with zarith.
- Qed.
-
- Hint Transparent ww_to_Z.
-
- Let ww_testbit_high n x y : Z.pos w_digits <= n ->
- Z.testbit [|WW x y|] n =
- Z.testbit (ZnZ.to_Z x) (n - Z.pos w_digits).
- Proof.
- intros Hn.
- assert (E : ZnZ.to_Z x = [|WW x y|] / wB).
- { simpl.
- rewrite Z.div_add_l; eauto with ZnZ zarith.
- now rewrite Z.div_small, Z.add_0_r; wwauto. }
- rewrite E.
- unfold wB, base. rewrite Z.div_pow2_bits.
- - f_equal; auto with zarith.
- - easy.
- - auto with zarith.
- Qed.
-
- Let ww_testbit_low n x y : 0 <= n < Z.pos w_digits ->
- Z.testbit [|WW x y|] n = Z.testbit (ZnZ.to_Z y) n.
- Proof.
- intros (Hn,Hn').
- assert (E : ZnZ.to_Z y = [|WW x y|] mod wB).
- { simpl; symmetry.
- rewrite Z.add_comm, Z.mod_add; auto with zarith nocore.
- apply Z.mod_small; eauto with ZnZ zarith. }
- rewrite E.
- unfold wB, base. symmetry. apply Z.mod_pow2_bits_low; auto.
- Qed.
-
- Let spec_lor x y : [|lor x y|] = Z.lor [|x|] [|y|].
- Proof.
- destruct x as [ |hx lx]. trivial.
- destruct y as [ |hy ly]. now rewrite Z.lor_comm.
- change ([|WW (ZnZ.lor hx hy) (ZnZ.lor lx ly)|] =
- Z.lor [|WW hx lx|] [|WW hy ly|]).
- apply Z.bits_inj'; intros n Hn.
- rewrite Z.lor_spec.
- destruct (Z.le_gt_cases (Z.pos w_digits) n) as [LE|GT].
- - now rewrite !ww_testbit_high, ZnZ.spec_lor, Z.lor_spec.
- - rewrite !ww_testbit_low; auto.
- now rewrite ZnZ.spec_lor, Z.lor_spec.
- Qed.
-
- Let spec_land x y : [|land x y|] = Z.land [|x|] [|y|].
- Proof.
- destruct x as [ |hx lx]. trivial.
- destruct y as [ |hy ly]. now rewrite Z.land_comm.
- change ([|WW (ZnZ.land hx hy) (ZnZ.land lx ly)|] =
- Z.land [|WW hx lx|] [|WW hy ly|]).
- apply Z.bits_inj'; intros n Hn.
- rewrite Z.land_spec.
- destruct (Z.le_gt_cases (Z.pos w_digits) n) as [LE|GT].
- - now rewrite !ww_testbit_high, ZnZ.spec_land, Z.land_spec.
- - rewrite !ww_testbit_low; auto.
- now rewrite ZnZ.spec_land, Z.land_spec.
- Qed.
-
- Let spec_lxor x y : [|lxor x y|] = Z.lxor [|x|] [|y|].
- Proof.
- destruct x as [ |hx lx]. trivial.
- destruct y as [ |hy ly]. now rewrite Z.lxor_comm.
- change ([|WW (ZnZ.lxor hx hy) (ZnZ.lxor lx ly)|] =
- Z.lxor [|WW hx lx|] [|WW hy ly|]).
- apply Z.bits_inj'; intros n Hn.
- rewrite Z.lxor_spec.
- destruct (Z.le_gt_cases (Z.pos w_digits) n) as [LE|GT].
- - now rewrite !ww_testbit_high, ZnZ.spec_lxor, Z.lxor_spec.
- - rewrite !ww_testbit_low; auto.
- now rewrite ZnZ.spec_lxor, Z.lxor_spec.
- Qed.
-
- Global Instance mk_zn2z_specs : ZnZ.Specs mk_zn2z_ops.
- Proof.
- apply ZnZ.MkSpecs; auto.
- exact spec_ww_add_mul_div.
-
- refine (@spec_ww_pos_mod t w_0 w_digits w_zdigits w_WW
- w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_zdigits.
- unfold w_to_Z, w_zdigits.
- rewrite ZnZ.spec_zdigits.
- rewrite <- Pos2Z.inj_xO; exact spec_ww_digits.
- Qed.
-
- Global Instance mk_zn2z_specs_karatsuba : ZnZ.Specs mk_zn2z_ops_karatsuba.
- Proof.
- apply ZnZ.MkSpecs; auto.
- exact spec_ww_add_mul_div.
- refine (@spec_ww_pos_mod t w_0 w_digits w_zdigits w_WW
- w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z
- _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
- exact ZnZ.spec_zdigits.
- unfold w_to_Z, w_zdigits.
- rewrite ZnZ.spec_zdigits.
- rewrite <- Pos2Z.inj_xO; exact spec_ww_digits.
- Qed.
-
-End Z_2nZ.
-
-
-Section MulAdd.
-
- Context {t : Type}{ops : ZnZ.Ops t}{specs : ZnZ.Specs ops}.
-
- Definition mul_add:= w_mul_add ZnZ.zero ZnZ.succ ZnZ.add_c ZnZ.mul_c.
-
- Notation "[| x |]" := (ZnZ.to_Z x) (at level 0, x at level 99).
-
- Notation "[|| x ||]" :=
- (zn2z_to_Z (base ZnZ.digits) ZnZ.to_Z x) (at level 0, x at level 99).
-
- Lemma spec_mul_add: forall x y z,
- let (zh, zl) := mul_add x y z in
- [||WW zh zl||] = [|x|] * [|y|] + [|z|].
- Proof.
- intros x y z.
- refine (spec_w_mul_add _ _ _ _ _ _ _ _ _ _ _ _ x y z); auto.
- exact ZnZ.spec_0.
- exact ZnZ.spec_to_Z.
- exact ZnZ.spec_succ.
- exact ZnZ.spec_add_c.
- exact ZnZ.spec_mul_c.
- Qed.
-
-End MulAdd.
-
-
-(** Modular versions of DoubleCyclic *)
-
-Module DoubleCyclic (C:CyclicType) <: CyclicType.
- Definition t := zn2z C.t.
- Instance ops : ZnZ.Ops t := mk_zn2z_ops.
- Instance specs : ZnZ.Specs ops := mk_zn2z_specs.
-End DoubleCyclic.
-
-Module DoubleCyclicKaratsuba (C:CyclicType) <: CyclicType.
- Definition t := zn2z C.t.
- Definition ops : ZnZ.Ops t := mk_zn2z_ops_karatsuba.
- Definition specs : ZnZ.Specs ops := mk_zn2z_specs_karatsuba.
-End DoubleCyclicKaratsuba.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
deleted file mode 100644
index 09d7329b6..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
+++ /dev/null
@@ -1,1494 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-Require Import DoubleDivn1.
-Require Import DoubleAdd.
-Require Import DoubleSub.
-
-Local Open Scope Z_scope.
-
-Ltac zarith := auto with zarith.
-
-
-Section POS_MOD.
-
- Variable w:Type.
- Variable w_0 : w.
- Variable w_digits : positive.
- Variable w_zdigits : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_pos_mod : w -> w -> w.
- Variable w_compare : w -> w -> comparison.
- Variable ww_compare : zn2z w -> zn2z w -> comparison.
- Variable w_0W : w -> zn2z w.
- Variable low: zn2z w -> w.
- Variable ww_sub: zn2z w -> zn2z w -> zn2z w.
- Variable ww_zdigits : zn2z w.
-
-
- Definition ww_pos_mod p x :=
- let zdigits := w_0W w_zdigits in
- match x with
- | W0 => W0
- | WW xh xl =>
- match ww_compare p zdigits with
- | Eq => w_WW w_0 xl
- | Lt => w_WW w_0 (w_pos_mod (low p) xl)
- | Gt =>
- match ww_compare p ww_zdigits with
- | Lt =>
- let n := low (ww_sub p zdigits) in
- w_WW (w_pos_mod n xh) xl
- | _ => x
- end
- end
- end.
-
-
- Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
-
-
- Variable spec_w_0 : [|w_0|] = 0.
-
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
-
- Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB.
-
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
-
- Variable spec_pos_mod : forall w p,
- [|w_pos_mod p w|] = [|w|] mod (2 ^ [|p|]).
-
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_ww_compare : forall x y,
- ww_compare x y = Z.compare [[x]] [[y]].
- Variable spec_ww_sub: forall x y,
- [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
-
- Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits.
- Variable spec_low: forall x, [| low x|] = [[x]] mod wB.
- Variable spec_ww_zdigits : [[ww_zdigits]] = 2 * [|w_zdigits|].
- Variable spec_ww_digits : ww_digits w_digits = xO w_digits.
-
-
- Hint Rewrite spec_w_0 spec_w_WW : w_rewrite.
-
- Lemma spec_ww_pos_mod : forall w p,
- [[ww_pos_mod p w]] = [[w]] mod (2 ^ [[p]]).
- assert (HHHHH:= lt_0_wB w_digits).
- assert (F0: forall x y, x - y + y = x); auto with zarith.
- intros w1 p; case (spec_to_w_Z p); intros HH1 HH2.
- unfold ww_pos_mod; case w1. reflexivity.
- intros xh xl; rewrite spec_ww_compare.
- case Z.compare_spec;
- rewrite spec_w_0W; rewrite spec_zdigits; fold wB;
- intros H1.
- rewrite H1; simpl ww_to_Z.
- autorewrite with w_rewrite rm10.
- rewrite Zplus_mod; auto with zarith.
- rewrite Z_mod_mult; auto with zarith.
- autorewrite with rm10.
- rewrite Zmod_mod; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- autorewrite with w_rewrite rm10.
- simpl ww_to_Z.
- rewrite spec_pos_mod.
- assert (HH0: [|low p|] = [[p]]).
- rewrite spec_low.
- apply Zmod_small; auto with zarith.
- case (spec_to_w_Z p); intros HHH1 HHH2; split; auto with zarith.
- apply Z.lt_le_trans with (1 := H1).
- unfold base; apply Zpower2_le_lin; auto with zarith.
- rewrite HH0.
- rewrite Zplus_mod; auto with zarith.
- unfold base.
- rewrite <- (F0 (Zpos w_digits) [[p]]).
- rewrite Zpower_exp; auto with zarith.
- rewrite Z.mul_assoc.
- rewrite Z_mod_mult; auto with zarith.
- autorewrite with w_rewrite rm10.
- rewrite Zmod_mod; auto with zarith.
- rewrite spec_ww_compare.
- case Z.compare_spec; rewrite spec_ww_zdigits;
- rewrite spec_zdigits; intros H2.
- replace (2^[[p]]) with wwB.
- rewrite Zmod_small; auto with zarith.
- unfold base; rewrite H2.
- rewrite spec_ww_digits; auto.
- assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] =
- [[p]] - Zpos w_digits).
- rewrite spec_low.
- rewrite spec_ww_sub.
- rewrite spec_w_0W; rewrite spec_zdigits.
- rewrite <- Zmod_div_mod; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.lt_le_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_le_lin; auto with zarith.
- exists wB; unfold base; rewrite <- Zpower_exp; auto with zarith.
- rewrite spec_ww_digits;
- apply f_equal with (f := Z.pow 2); rewrite Pos2Z.inj_xO; auto with zarith.
- simpl ww_to_Z; autorewrite with w_rewrite.
- rewrite spec_pos_mod; rewrite HH0.
- pattern [|xh|] at 2;
- rewrite Z_div_mod_eq with (b := 2 ^ ([[p]] - Zpos w_digits));
- auto with zarith.
- rewrite (fun x => (Z.mul_comm (2 ^ x))); rewrite Z.mul_add_distr_r.
- unfold base; rewrite <- Z.mul_assoc; rewrite <- Zpower_exp;
- auto with zarith.
- rewrite F0; auto with zarith.
- rewrite <- Z.add_assoc; rewrite Zplus_mod; auto with zarith.
- rewrite Z_mod_mult; auto with zarith.
- autorewrite with rm10.
- rewrite Zmod_mod; auto with zarith.
- symmetry; apply Zmod_small; auto with zarith.
- case (spec_to_Z xh); intros U1 U2.
- case (spec_to_Z xl); intros U3 U4.
- split; auto with zarith.
- apply Z.add_nonneg_nonneg; auto with zarith.
- apply Z.mul_nonneg_nonneg; auto with zarith.
- match goal with |- 0 <= ?X mod ?Y =>
- case (Z_mod_lt X Y); auto with zarith
- end.
- match goal with |- ?X mod ?Y * ?U + ?Z < ?T =>
- apply Z.le_lt_trans with ((Y - 1) * U + Z );
- [case (Z_mod_lt X Y); auto with zarith | idtac]
- end.
- match goal with |- ?X * ?U + ?Y < ?Z =>
- apply Z.le_lt_trans with (X * U + (U - 1))
- end.
- apply Z.add_le_mono_l; auto with zarith.
- case (spec_to_Z xl); unfold base; auto with zarith.
- rewrite Z.mul_sub_distr_r; rewrite <- Zpower_exp; auto with zarith.
- rewrite F0; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- case (spec_to_w_Z (WW xh xl)); intros U1 U2.
- split; auto with zarith.
- apply Z.lt_le_trans with (1:= U2).
- unfold base; rewrite spec_ww_digits.
- apply Zpower_le_monotone; auto with zarith.
- split; auto with zarith.
- rewrite Pos2Z.inj_xO; auto with zarith.
- Qed.
-
-End POS_MOD.
-
-Section DoubleDiv32.
-
- Variable w : Type.
- Variable w_0 : w.
- Variable w_Bm1 : w.
- Variable w_Bm2 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_compare : w -> w -> comparison.
- Variable w_add_c : w -> w -> carry w.
- Variable w_add_carry_c : w -> w -> carry w.
- Variable w_add : w -> w -> w.
- Variable w_add_carry : w -> w -> w.
- Variable w_pred : w -> w.
- Variable w_sub : w -> w -> w.
- Variable w_mul_c : w -> w -> zn2z w.
- Variable w_div21 : w -> w -> w -> w*w.
- Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w).
-
- Definition w_div32_body a1 a2 a3 b1 b2 :=
- match w_compare a1 b1 with
- | Lt =>
- let (q,r) := w_div21 a1 a2 b1 in
- match ww_sub_c (w_WW r a3) (w_mul_c q b2) with
- | C0 r1 => (q,r1)
- | C1 r1 =>
- let q := w_pred q in
- ww_add_c_cont w_WW w_add_c w_add_carry_c
- (fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2)))
- (fun r2 => (q,r2))
- r1 (WW b1 b2)
- end
- | Eq =>
- ww_add_c_cont w_WW w_add_c w_add_carry_c
- (fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2)))
- (fun r => (w_Bm1,r))
- (WW (w_sub a2 b2) a3) (WW b1 b2)
- | Gt => (w_0, W0) (* cas absurde *)
- end.
-
- Definition w_div32 a1 a2 a3 b1 b2 :=
- Eval lazy beta iota delta [ww_add_c_cont ww_add w_div32_body] in
- w_div32_body a1 a2 a3 b1 b2.
-
- (* Proof *)
-
- Variable w_digits : positive.
- Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
- Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
-
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1.
- Variable spec_w_Bm2 : [|w_Bm2|] = wB - 2.
-
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
-
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_compare :
- forall x y, w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
- Variable spec_w_add_carry_c :
- forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1.
-
- Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB.
- Variable spec_w_add_carry :
- forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB.
-
- Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB.
- Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
-
- Variable spec_mul_c : forall x y, [[ w_mul_c x y ]] = [|x|] * [|y|].
- Variable spec_div21 : forall a1 a2 b,
- wB/2 <= [|b|] ->
- [|a1|] < [|b|] ->
- let (q,r) := w_div21 a1 a2 b in
- [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
-
- Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]].
-
- Ltac Spec_w_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_to_Z x).
- Ltac Spec_ww_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x).
-
- Theorem wB_div2: forall x, wB/2 <= x -> wB <= 2 * x.
- intros x H; rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; auto with zarith.
- Qed.
-
- Lemma Zmult_lt_0_reg_r_2 : forall n m : Z, 0 <= n -> 0 < m * n -> 0 < m.
- Proof.
- intros n m H1 H2;apply Z.mul_pos_cancel_r with n;trivial.
- Z.le_elim H1; trivial.
- subst;rewrite Z.mul_0_r in H2;discriminate H2.
- Qed.
-
- Theorem spec_w_div32 : forall a1 a2 a3 b1 b2,
- wB/2 <= [|b1|] ->
- [[WW a1 a2]] < [[WW b1 b2]] ->
- let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- [|a1|] * wwB + [|a2|] * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
- 0 <= [[r]] < [|b1|] * wB + [|b2|].
- Proof.
- intros a1 a2 a3 b1 b2 Hle Hlt.
- assert (U:= lt_0_wB w_digits); assert (U1:= lt_0_wwB w_digits).
- Spec_w_to_Z a1;Spec_w_to_Z a2;Spec_w_to_Z a3;Spec_w_to_Z b1;Spec_w_to_Z b2.
- rewrite wwB_wBwB; rewrite Z.pow_2_r; rewrite Z.mul_assoc;rewrite <- Z.mul_add_distr_r.
- change (w_div32 a1 a2 a3 b1 b2) with (w_div32_body a1 a2 a3 b1 b2).
- unfold w_div32_body.
- rewrite spec_compare. case Z.compare_spec; intro Hcmp.
- simpl in Hlt.
- rewrite Hcmp in Hlt;assert ([|a2|] < [|b2|]). omega.
- assert ([[WW (w_sub a2 b2) a3]] = ([|a2|]-[|b2|])*wB + [|a3|] + wwB).
- simpl;rewrite spec_sub.
- assert ([|a2|] - [|b2|] = wB*(-1) + ([|a2|] - [|b2|] + wB)). ring.
- assert (0 <= [|a2|] - [|b2|] + wB < wB). omega.
- rewrite <-(Zmod_unique ([|a2|]-[|b2|]) wB (-1) ([|a2|]-[|b2|]+wB) H1 H0).
- rewrite wwB_wBwB;ring.
- assert (U2 := wB_pos w_digits).
- eapply spec_ww_add_c_cont with (P :=
- fun (x y:zn2z w) (res:w*zn2z w) =>
- let (q, r) := res in
- ([|a1|] * wB + [|a2|]) * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
- 0 <= [[r]] < [|b1|] * wB + [|b2|]);eauto.
- rewrite H0;intros r.
- repeat
- (rewrite spec_ww_add;eauto || rewrite spec_w_Bm1 || rewrite spec_w_Bm2);
- simpl ww_to_Z;try rewrite Z.mul_1_l;intros H1.
- assert (0<= ([[r]] + ([|b1|] * wB + [|b2|])) - wwB < [|b1|] * wB + [|b2|]).
- Spec_ww_to_Z r;split;zarith.
- rewrite H1.
- assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB).
- rewrite wwB_wBwB; rewrite Z.pow_2_r; zarith.
- assert (-wwB < ([|a2|] - [|b2|]) * wB + [|a3|] < 0).
- split. apply Z.lt_le_trans with (([|a2|] - [|b2|]) * wB);zarith.
- rewrite wwB_wBwB;replace (-(wB^2)) with (-wB*wB);[zarith | ring].
- apply Z.mul_lt_mono_pos_r;zarith.
- apply Z.le_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith.
- replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with
- (([|a2|] - [|b2|] + 1) * wB + - 1);[zarith | ring].
- assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith.
- replace 0 with (0*wB);zarith.
- replace (([|a2|] - [|b2|]) * wB + [|a3|] + wwB + ([|b1|] * wB + [|b2|]) +
- ([|b1|] * wB + [|b2|]) - wwB) with
- (([|a2|] - [|b2|]) * wB + [|a3|] + 2*[|b1|] * wB + 2*[|b2|]);
- [zarith | ring].
- rewrite <- (Zmod_unique ([[r]] + ([|b1|] * wB + [|b2|])) wwB
- 1 ([[r]] + ([|b1|] * wB + [|b2|]) - wwB));zarith;try (ring;fail).
- split. rewrite H1;rewrite Hcmp;ring. trivial.
- Spec_ww_to_Z (WW b1 b2). simpl in HH4;zarith.
- rewrite H0;intros r;repeat
- (rewrite spec_w_Bm1 || rewrite spec_w_Bm2);
- simpl ww_to_Z;try rewrite Z.mul_1_l;intros H1.
- assert ([[r]]=([|a2|]-[|b2|])*wB+[|a3|]+([|b1|]*wB+[|b2|])). zarith.
- split. rewrite H2;rewrite Hcmp;ring.
- split. Spec_ww_to_Z r;zarith.
- rewrite H2.
- assert (([|a2|] - [|b2|]) * wB + [|a3|] < 0);zarith.
- apply Z.le_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith.
- replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with
- (([|a2|] - [|b2|] + 1) * wB + - 1);[zarith|ring].
- assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith.
- replace 0 with (0*wB);zarith.
- (* Cas Lt *)
- assert (Hdiv21 := spec_div21 a2 Hle Hcmp);
- destruct (w_div21 a1 a2 b1) as (q, r);destruct Hdiv21.
- rewrite H.
- assert (Hq := spec_to_Z q).
- generalize
- (spec_ww_sub_c (w_WW r a3) (w_mul_c q b2));
- destruct (ww_sub_c (w_WW r a3) (w_mul_c q b2))
- as [r1|r1];repeat (rewrite spec_w_WW || rewrite spec_mul_c);
- unfold interp_carry;intros H1.
- rewrite H1.
- split. ring. split.
- rewrite <- H1;destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z r1);trivial.
- apply Z.le_lt_trans with ([|r|] * wB + [|a3|]).
- assert ( 0 <= [|q|] * [|b2|]);zarith.
- apply beta_lex_inv;zarith.
- assert ([[r1]] = [|r|] * wB + [|a3|] - [|q|] * [|b2|] + wwB).
- rewrite <- H1;ring.
- Spec_ww_to_Z r1; assert (0 <= [|r|]*wB). zarith.
- assert (0 < [|q|] * [|b2|]). zarith.
- assert (0 < [|q|]).
- apply Zmult_lt_0_reg_r_2 with [|b2|];zarith.
- eapply spec_ww_add_c_cont with (P :=
- fun (x y:zn2z w) (res:w*zn2z w) =>
- let (q0, r0) := res in
- ([|q|] * [|b1|] + [|r|]) * wB + [|a3|] =
- [|q0|] * ([|b1|] * wB + [|b2|]) + [[r0]] /\
- 0 <= [[r0]] < [|b1|] * wB + [|b2|]);eauto.
- intros r2;repeat (rewrite spec_pred || rewrite spec_ww_add;eauto);
- simpl ww_to_Z;intros H7.
- assert (0 < [|q|] - 1).
- assert (H6 : 1 <= [|q|]) by zarith.
- Z.le_elim H6;zarith.
- rewrite <- H6 in H2;rewrite H2 in H7.
- assert (0 < [|b1|]*wB). apply Z.mul_pos_pos;zarith.
- Spec_ww_to_Z r2. zarith.
- rewrite (Zmod_small ([|q|] -1));zarith.
- rewrite (Zmod_small ([|q|] -1 -1));zarith.
- assert ([[r2]] + ([|b1|] * wB + [|b2|]) =
- wwB * 1 +
- ([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2 * ([|b1|] * wB + [|b2|]))).
- rewrite H7;rewrite H2;ring.
- assert
- ([|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|])
- < [|b1|]*wB + [|b2|]).
- Spec_ww_to_Z r2;omega.
- Spec_ww_to_Z (WW b1 b2). simpl in HH5.
- assert
- (0 <= [|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|])
- < wwB). split;try omega.
- replace (2*([|b1|]*wB+[|b2|])) with ((2*[|b1|])*wB+2*[|b2|]). 2:ring.
- assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB).
- rewrite wwB_wBwB; rewrite Z.pow_2_r; zarith. omega.
- rewrite <- (Zmod_unique
- ([[r2]] + ([|b1|] * wB + [|b2|]))
- wwB
- 1
- ([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2*([|b1|] * wB + [|b2|]))
- H10 H8).
- split. ring. zarith.
- intros r2;repeat (rewrite spec_pred);simpl ww_to_Z;intros H7.
- rewrite (Zmod_small ([|q|] -1));zarith.
- split.
- replace [[r2]] with ([[r1]] + ([|b1|] * wB + [|b2|]) -wwB).
- rewrite H2; ring. rewrite <- H7; ring.
- Spec_ww_to_Z r2;Spec_ww_to_Z r1. omega.
- simpl in Hlt.
- assert ([|a1|] * wB + [|a2|] <= [|b1|] * wB + [|b2|]). zarith.
- assert (H1 := beta_lex _ _ _ _ _ H HH0 HH3). rewrite spec_w_0;simpl;zarith.
- Qed.
-
-
-End DoubleDiv32.
-
-Section DoubleDiv21.
- Variable w : Type.
- Variable w_0 : w.
-
- Variable w_0W : w -> zn2z w.
- Variable w_div32 : w -> w -> w -> w -> w -> w * zn2z w.
-
- Variable ww_1 : zn2z w.
- Variable ww_compare : zn2z w -> zn2z w -> comparison.
- Variable ww_sub : zn2z w -> zn2z w -> zn2z w.
-
-
- Definition ww_div21 a1 a2 b :=
- match a1 with
- | W0 =>
- match ww_compare a2 b with
- | Gt => (ww_1, ww_sub a2 b)
- | Eq => (ww_1, W0)
- | Lt => (W0, a2)
- end
- | WW a1h a1l =>
- match a2 with
- | W0 =>
- match b with
- | W0 => (W0,W0) (* cas absurde *)
- | WW b1 b2 =>
- let (q1, r) := w_div32 a1h a1l w_0 b1 b2 in
- match r with
- | W0 => (WW q1 w_0, W0)
- | WW r1 r2 =>
- let (q2, s) := w_div32 r1 r2 w_0 b1 b2 in
- (WW q1 q2, s)
- end
- end
- | WW a2h a2l =>
- match b with
- | W0 => (W0,W0) (* cas absurde *)
- | WW b1 b2 =>
- let (q1, r) := w_div32 a1h a1l a2h b1 b2 in
- match r with
- | W0 => (WW q1 w_0, w_0W a2l)
- | WW r1 r2 =>
- let (q2, s) := w_div32 r1 r2 a2l b1 b2 in
- (WW q1 q2, s)
- end
- end
- end
- end.
-
- (* Proof *)
-
- Variable w_digits : positive.
- Variable w_to_Z : w -> Z.
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_w_div32 : forall a1 a2 a3 b1 b2,
- wB/2 <= [|b1|] ->
- [[WW a1 a2]] < [[WW b1 b2]] ->
- let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- [|a1|] * wwB + [|a2|] * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
- 0 <= [[r]] < [|b1|] * wB + [|b2|].
- Variable spec_ww_1 : [[ww_1]] = 1.
- Variable spec_ww_compare : forall x y,
- ww_compare x y = Z.compare [[x]] [[y]].
- Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
-
- Theorem wwB_div: wwB = 2 * (wwB / 2).
- Proof.
- rewrite wwB_div_2; rewrite Z.mul_assoc; rewrite wB_div_2; auto.
- rewrite <- Z.pow_2_r; apply wwB_wBwB.
- Qed.
-
- Ltac Spec_w_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_to_Z x).
- Ltac Spec_ww_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x).
-
- Theorem spec_ww_div21 : forall a1 a2 b,
- wwB/2 <= [[b]] ->
- [[a1]] < [[b]] ->
- let (q,r) := ww_div21 a1 a2 b in
- [[a1]] *wwB+[[a2]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]].
- Proof.
- assert (U:= lt_0_wB w_digits).
- assert (U1:= lt_0_wwB w_digits).
- intros a1 a2 b H Hlt; unfold ww_div21.
- Spec_ww_to_Z b; assert (Eq: 0 < [[b]]). Spec_ww_to_Z a1;omega.
- generalize Hlt H ;clear Hlt H;case a1.
- intros H1 H2;simpl in H1;Spec_ww_to_Z a2.
- rewrite spec_ww_compare. case Z.compare_spec;
- simpl;try rewrite spec_ww_1;autorewrite with rm10; intros;zarith.
- rewrite spec_ww_sub;simpl. rewrite Zmod_small;zarith.
- split. ring.
- assert (wwB <= 2*[[b]]);zarith.
- rewrite wwB_div;zarith.
- intros a1h a1l. Spec_w_to_Z a1h;Spec_w_to_Z a1l. Spec_ww_to_Z a2.
- destruct a2 as [ |a3 a4];
- (destruct b as [ |b1 b2];[unfold Z.le in Eq;discriminate Eq|idtac]);
- try (Spec_w_to_Z a3; Spec_w_to_Z a4); Spec_w_to_Z b1; Spec_w_to_Z b2;
- intros Hlt H; match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] =>
- generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U);
- intros q1 r H0
- end; (assert (Eq1: wB / 2 <= [|b1|]);[
- apply (@beta_lex (wB / 2) 0 [|b1|] [|b2|] wB); auto with zarith;
- autorewrite with rm10;repeat rewrite (Z.mul_comm wB);
- rewrite <- wwB_div_2; trivial
- | generalize (H0 Eq1 Hlt);clear H0;destruct r as [ |r1 r2];simpl;
- try rewrite spec_w_0; try rewrite spec_w_0W;repeat rewrite Z.add_0_r;
- intros (H1,H2) ]).
- split;[rewrite wwB_wBwB; rewrite Z.pow_2_r | trivial].
- rewrite Z.mul_assoc;rewrite Z.mul_add_distr_r;rewrite <- Z.mul_assoc;
- rewrite <- Z.pow_2_r; rewrite <- wwB_wBwB;rewrite H1;ring.
- destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] =>
- generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U);
- intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end.
- split;[rewrite wwB_wBwB | trivial].
- rewrite Z.pow_2_r.
- rewrite Z.mul_assoc;rewrite Z.mul_add_distr_r;rewrite <- Z.mul_assoc;
- rewrite <- Z.pow_2_r.
- rewrite <- wwB_wBwB;rewrite H1.
- rewrite spec_w_0 in H4;rewrite Z.add_0_r in H4.
- repeat rewrite Z.mul_add_distr_r. rewrite <- (Z.mul_assoc [|r1|]).
- rewrite <- Z.pow_2_r; rewrite <- wwB_wBwB;rewrite H4;simpl;ring.
- split;[rewrite wwB_wBwB | split;zarith].
- replace (([|a1h|] * wB + [|a1l|]) * wB^2 + ([|a3|] * wB + [|a4|]))
- with (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB+ [|a4|]).
- rewrite H1;ring. rewrite wwB_wBwB;ring.
- change [|a4|] with (0*wB+[|a4|]);apply beta_lex_inv;zarith.
- assert (1 <= wB/2);zarith.
- assert (H_:= wB_pos w_digits);apply Zdiv_le_lower_bound;zarith.
- destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] =>
- generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U);
- intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end.
- split;trivial.
- replace (([|a1h|] * wB + [|a1l|]) * wwB + ([|a3|] * wB + [|a4|])) with
- (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB + [|a4|]);
- [rewrite H1 | rewrite wwB_wBwB;ring].
- replace (([|q1|]*([|b1|]*wB+[|b2|])+([|r1|]*wB+[|r2|]))*wB+[|a4|]) with
- (([|q1|]*([|b1|]*wB+[|b2|]))*wB+([|r1|]*wwB+[|r2|]*wB+[|a4|]));
- [rewrite H4;simpl|rewrite wwB_wBwB];ring.
- Qed.
-
-End DoubleDiv21.
-
-Section DoubleDivGt.
- Variable w : Type.
- Variable w_digits : positive.
- Variable w_0 : w.
-
- Variable w_WW : w -> w -> zn2z w.
- Variable w_0W : w -> zn2z w.
- Variable w_compare : w -> w -> comparison.
- Variable w_eq0 : w -> bool.
- Variable w_opp_c : w -> carry w.
- Variable w_opp w_opp_carry : w -> w.
- Variable w_sub_c : w -> w -> carry w.
- Variable w_sub w_sub_carry : w -> w -> w.
-
- Variable w_div_gt : w -> w -> w*w.
- Variable w_mod_gt : w -> w -> w.
- Variable w_gcd_gt : w -> w -> w.
- Variable w_add_mul_div : w -> w -> w -> w.
- Variable w_head0 : w -> w.
- Variable w_div21 : w -> w -> w -> w * w.
- Variable w_div32 : w -> w -> w -> w -> w -> w * zn2z w.
-
-
- Variable _ww_zdigits : zn2z w.
- Variable ww_1 : zn2z w.
- Variable ww_add_mul_div : zn2z w -> zn2z w -> zn2z w -> zn2z w.
-
- Variable w_zdigits : w.
-
- Definition ww_div_gt_aux ah al bh bl :=
- Eval lazy beta iota delta [ww_sub ww_opp] in
- let p := w_head0 bh in
- match w_compare p w_0 with
- | Gt =>
- let b1 := w_add_mul_div p bh bl in
- let b2 := w_add_mul_div p bl w_0 in
- let a1 := w_add_mul_div p w_0 ah in
- let a2 := w_add_mul_div p ah al in
- let a3 := w_add_mul_div p al w_0 in
- let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- (WW w_0 q, ww_add_mul_div
- (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
- w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r)
- | _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
- w_opp w_sub w_sub_carry (WW ah al) (WW bh bl))
- end.
-
- Definition ww_div_gt a b :=
- Eval lazy beta iota delta [ww_div_gt_aux double_divn1
- double_divn1_p double_divn1_p_aux double_divn1_0 double_divn1_0_aux
- double_split double_0 double_WW] in
- match a, b with
- | W0, _ => (W0,W0)
- | _, W0 => (W0,W0)
- | WW ah al, WW bh bl =>
- if w_eq0 ah then
- let (q,r) := w_div_gt al bl in
- (WW w_0 q, w_0W r)
- else
- match w_compare w_0 bh with
- | Eq =>
- let(q,r):=
- double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 a bl in
- (q, w_0W r)
- | Lt => ww_div_gt_aux ah al bh bl
- | Gt => (W0,W0) (* cas absurde *)
- end
- end.
-
- Definition ww_mod_gt_aux ah al bh bl :=
- Eval lazy beta iota delta [ww_sub ww_opp] in
- let p := w_head0 bh in
- match w_compare p w_0 with
- | Gt =>
- let b1 := w_add_mul_div p bh bl in
- let b2 := w_add_mul_div p bl w_0 in
- let a1 := w_add_mul_div p w_0 ah in
- let a2 := w_add_mul_div p ah al in
- let a3 := w_add_mul_div p al w_0 in
- let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
- w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r
- | _ =>
- ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
- w_opp w_sub w_sub_carry (WW ah al) (WW bh bl)
- end.
-
- Definition ww_mod_gt a b :=
- Eval lazy beta iota delta [ww_mod_gt_aux double_modn1
- double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux
- double_split double_0 double_WW snd] in
- match a, b with
- | W0, _ => W0
- | _, W0 => W0
- | WW ah al, WW bh bl =>
- if w_eq0 ah then w_0W (w_mod_gt al bl)
- else
- match w_compare w_0 bh with
- | Eq =>
- w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 a bl)
- | Lt => ww_mod_gt_aux ah al bh bl
- | Gt => W0 (* cas absurde *)
- end
- end.
-
- Definition ww_gcd_gt_body (cont: w->w->w->w->zn2z w) (ah al bh bl: w) :=
- Eval lazy beta iota delta [ww_mod_gt_aux double_modn1
- double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux
- double_split double_0 double_WW snd] in
- match w_compare w_0 bh with
- | Eq =>
- match w_compare w_0 bl with
- | Eq => WW ah al (* normalement n'arrive pas si forme normale *)
- | Lt =>
- let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 (WW ah al) bl in
- WW w_0 (w_gcd_gt bl m)
- | Gt => W0 (* absurde *)
- end
- | Lt =>
- let m := ww_mod_gt_aux ah al bh bl in
- match m with
- | W0 => WW bh bl
- | WW mh ml =>
- match w_compare w_0 mh with
- | Eq =>
- match w_compare w_0 ml with
- | Eq => WW bh bl
- | _ =>
- let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 (WW bh bl) ml in
- WW w_0 (w_gcd_gt ml r)
- end
- | Lt =>
- let r := ww_mod_gt_aux bh bl mh ml in
- match r with
- | W0 => m
- | WW rh rl => cont mh ml rh rl
- end
- | Gt => W0 (* absurde *)
- end
- end
- | Gt => W0 (* absurde *)
- end.
-
- Fixpoint ww_gcd_gt_aux
- (p:positive) (cont: w -> w -> w -> w -> zn2z w) (ah al bh bl : w)
- {struct p} : zn2z w :=
- ww_gcd_gt_body
- (fun mh ml rh rl => match p with
- | xH => cont mh ml rh rl
- | xO p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl
- | xI p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl
- end) ah al bh bl.
-
-
- (* Proof *)
-
- Variable w_to_Z : w -> Z.
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB.
-
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_compare :
- forall x y, w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0.
-
- Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|].
- Variable spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB.
- Variable spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1.
-
- Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|].
- Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
- Variable spec_sub_carry :
- forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB.
-
- Variable spec_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] ->
- let (q,r) := w_div_gt a b in
- [|a|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
- Variable spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] ->
- [|w_mod_gt a b|] = [|a|] mod [|b|].
- Variable spec_gcd_gt : forall a b, [|a|] > [|b|] ->
- Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|].
-
- Variable spec_add_mul_div : forall x y p,
- [|p|] <= Zpos w_digits ->
- [| w_add_mul_div p x y |] =
- ([|x|] * (2 ^ ([|p|])) +
- [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB.
- Variable spec_head0 : forall x, 0 < [|x|] ->
- wB/ 2 <= 2 ^ [|w_head0 x|] * [|x|] < wB.
-
- Variable spec_div21 : forall a1 a2 b,
- wB/2 <= [|b|] ->
- [|a1|] < [|b|] ->
- let (q,r) := w_div21 a1 a2 b in
- [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
-
- Variable spec_w_div32 : forall a1 a2 a3 b1 b2,
- wB/2 <= [|b1|] ->
- [[WW a1 a2]] < [[WW b1 b2]] ->
- let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- [|a1|] * wwB + [|a2|] * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
- 0 <= [[r]] < [|b1|] * wB + [|b2|].
-
- Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits.
-
- Variable spec_ww_digits_ : [[_ww_zdigits]] = Zpos (xO w_digits).
- Variable spec_ww_1 : [[ww_1]] = 1.
- Variable spec_ww_add_mul_div : forall x y p,
- [[p]] <= Zpos (xO w_digits) ->
- [[ ww_add_mul_div p x y ]] =
- ([[x]] * (2^[[p]]) +
- [[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB.
-
- Ltac Spec_w_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_to_Z x).
-
- Ltac Spec_ww_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x).
-
- Lemma to_Z_div_minus_p : forall x p,
- 0 < [|p|] < Zpos w_digits ->
- 0 <= [|x|] / 2 ^ (Zpos w_digits - [|p|]) < 2 ^ [|p|].
- Proof.
- intros x p H;Spec_w_to_Z x.
- split. apply Zdiv_le_lower_bound;zarith.
- apply Zdiv_lt_upper_bound;zarith.
- rewrite <- Zpower_exp;zarith.
- ring_simplify ([|p|] + (Zpos w_digits - [|p|])); unfold base in HH;zarith.
- Qed.
- Hint Resolve to_Z_div_minus_p : zarith.
-
- Lemma spec_ww_div_gt_aux : forall ah al bh bl,
- [[WW ah al]] > [[WW bh bl]] ->
- 0 < [|bh|] ->
- let (q,r) := ww_div_gt_aux ah al bh bl in
- [[WW ah al]] = [[q]] * [[WW bh bl]] + [[r]] /\
- 0 <= [[r]] < [[WW bh bl]].
- Proof.
- intros ah al bh bl Hgt Hpos;unfold ww_div_gt_aux.
- change
- (let (q, r) := let p := w_head0 bh in
- match w_compare p w_0 with
- | Gt =>
- let b1 := w_add_mul_div p bh bl in
- let b2 := w_add_mul_div p bl w_0 in
- let a1 := w_add_mul_div p w_0 ah in
- let a2 := w_add_mul_div p ah al in
- let a3 := w_add_mul_div p al w_0 in
- let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- (WW w_0 q, ww_add_mul_div
- (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
- w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r)
- | _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
- w_opp w_sub w_sub_carry (WW ah al) (WW bh bl))
- end in [[WW ah al]]=[[q]]*[[WW bh bl]]+[[r]] /\ 0 <=[[r]]< [[WW bh bl]]).
- assert (Hh := spec_head0 Hpos).
- lazy zeta.
- rewrite spec_compare; case Z.compare_spec;
- rewrite spec_w_0; intros HH.
- generalize Hh; rewrite HH; simpl Z.pow;
- rewrite Z.mul_1_l; intros (HH1, HH2); clear HH.
- assert (wwB <= 2*[[WW bh bl]]).
- apply Z.le_trans with (2*[|bh|]*wB).
- rewrite wwB_wBwB; rewrite Z.pow_2_r; apply Z.mul_le_mono_nonneg_r; zarith.
- rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; zarith.
- simpl ww_to_Z;rewrite Z.mul_add_distr_l;rewrite Z.mul_assoc.
- Spec_w_to_Z bl;zarith.
- Spec_ww_to_Z (WW ah al).
- rewrite spec_ww_sub;eauto.
- simpl;rewrite spec_ww_1;rewrite Z.mul_1_l;simpl.
- simpl ww_to_Z in Hgt, H, HH;rewrite Zmod_small;split;zarith.
- case (spec_to_Z (w_head0 bh)); auto with zarith.
- assert ([|w_head0 bh|] < Zpos w_digits).
- destruct (Z_lt_ge_dec [|w_head0 bh|] (Zpos w_digits));trivial.
- exfalso.
- assert (2 ^ [|w_head0 bh|] * [|bh|] >= wB);auto with zarith.
- apply Z.le_ge; replace wB with (wB * 1);try ring.
- Spec_w_to_Z bh;apply Z.mul_le_mono_nonneg;zarith.
- unfold base;apply Zpower_le_monotone;zarith.
- assert (HHHH : 0 < [|w_head0 bh|] < Zpos w_digits); auto with zarith.
- assert (Hb:= Z.lt_le_incl _ _ H).
- generalize (spec_add_mul_div w_0 ah Hb)
- (spec_add_mul_div ah al Hb)
- (spec_add_mul_div al w_0 Hb)
- (spec_add_mul_div bh bl Hb)
- (spec_add_mul_div bl w_0 Hb);
- rewrite spec_w_0; repeat rewrite Z.mul_0_l;repeat rewrite Z.add_0_l;
- rewrite Zdiv_0_l;repeat rewrite Z.add_0_r.
- Spec_w_to_Z ah;Spec_w_to_Z bh.
- unfold base;repeat rewrite Zmod_shift_r;zarith.
- assert (H3:=to_Z_div_minus_p ah HHHH);assert(H4:=to_Z_div_minus_p al HHHH);
- assert (H5:=to_Z_div_minus_p bl HHHH).
- rewrite Z.mul_comm in Hh.
- assert (2^[|w_head0 bh|] < wB). unfold base;apply Zpower_lt_monotone;zarith.
- unfold base in H0;rewrite Zmod_small;zarith.
- fold wB; rewrite (Zmod_small ([|bh|] * 2 ^ [|w_head0 bh|]));zarith.
- intros U1 U2 U3 V1 V2.
- generalize (@spec_w_div32 (w_add_mul_div (w_head0 bh) w_0 ah)
- (w_add_mul_div (w_head0 bh) ah al)
- (w_add_mul_div (w_head0 bh) al w_0)
- (w_add_mul_div (w_head0 bh) bh bl)
- (w_add_mul_div (w_head0 bh) bl w_0)).
- destruct (w_div32 (w_add_mul_div (w_head0 bh) w_0 ah)
- (w_add_mul_div (w_head0 bh) ah al)
- (w_add_mul_div (w_head0 bh) al w_0)
- (w_add_mul_div (w_head0 bh) bh bl)
- (w_add_mul_div (w_head0 bh) bl w_0)) as (q,r).
- rewrite V1;rewrite V2. rewrite Z.mul_add_distr_r.
- rewrite <- (Z.add_assoc ([|bh|] * 2 ^ [|w_head0 bh|] * wB)).
- unfold base;rewrite <- shift_unshift_mod;zarith. fold wB.
- replace ([|bh|] * 2 ^ [|w_head0 bh|] * wB + [|bl|] * 2 ^ [|w_head0 bh|]) with
- ([[WW bh bl]] * 2^[|w_head0 bh|]). 2:simpl;ring.
- fold wwB. rewrite wwB_wBwB. rewrite Z.pow_2_r. rewrite U1;rewrite U2;rewrite U3.
- rewrite Z.mul_assoc. rewrite Z.mul_add_distr_r.
- rewrite (Z.add_assoc ([|ah|] / 2^(Zpos(w_digits) - [|w_head0 bh|])*wB * wB)).
- rewrite <- Z.mul_add_distr_r. rewrite <- Z.add_assoc.
- unfold base;repeat rewrite <- shift_unshift_mod;zarith. fold wB.
- replace ([|ah|] * 2 ^ [|w_head0 bh|] * wB + [|al|] * 2 ^ [|w_head0 bh|]) with
- ([[WW ah al]] * 2^[|w_head0 bh|]). 2:simpl;ring.
- intros Hd;destruct Hd;zarith.
- simpl. apply beta_lex_inv;zarith. rewrite U1;rewrite V1.
- assert ([|ah|] / 2 ^ (Zpos (w_digits) - [|w_head0 bh|]) < wB/2);zarith.
- apply Zdiv_lt_upper_bound;zarith.
- unfold base.
- replace (2^Zpos (w_digits)) with (2^(Zpos (w_digits) - 1)*2).
- rewrite Z_div_mult;zarith. rewrite <- Zpower_exp;zarith.
- apply Z.lt_le_trans with wB;zarith.
- unfold base;apply Zpower_le_monotone;zarith.
- pattern 2 at 2;replace 2 with (2^1);trivial.
- rewrite <- Zpower_exp;zarith. ring_simplify (Zpos (w_digits) - 1 + 1);trivial.
- change [[WW w_0 q]] with ([|w_0|]*wB+[|q|]);rewrite spec_w_0;rewrite
- Z.mul_0_l;rewrite Z.add_0_l.
- replace [[ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry
- _ww_zdigits (w_0W (w_head0 bh))) W0 r]] with ([[r]]/2^[|w_head0 bh|]).
- assert (0 < 2^[|w_head0 bh|]). apply Z.pow_pos_nonneg;zarith.
- split.
- rewrite <- (Z_div_mult [[WW ah al]] (2^[|w_head0 bh|]));zarith.
- rewrite H1;rewrite Z.mul_assoc;apply Z_div_plus_l;trivial.
- split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith.
- rewrite spec_ww_add_mul_div.
- rewrite spec_ww_sub; auto with zarith.
- rewrite spec_ww_digits_.
- change (Zpos (xO (w_digits))) with (2*Zpos (w_digits));zarith.
- simpl ww_to_Z;rewrite Z.mul_0_l;rewrite Z.add_0_l.
- rewrite spec_w_0W.
- rewrite (fun x y => Zmod_small (x-y)); auto with zarith.
- ring_simplify (2 * Zpos w_digits - (2 * Zpos w_digits - [|w_head0 bh|])).
- rewrite Zmod_small;zarith.
- split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith.
- Spec_ww_to_Z r.
- apply Z.lt_le_trans with wwB;zarith.
- rewrite <- (Z.mul_1_r wwB);apply Z.mul_le_mono_nonneg;zarith.
- split; auto with zarith.
- apply Z.le_lt_trans with (2 * Zpos w_digits); auto with zarith.
- unfold base, ww_digits; rewrite (Pos2Z.inj_xO w_digits).
- apply Zpower2_lt_lin; auto with zarith.
- rewrite spec_ww_sub; auto with zarith.
- rewrite spec_ww_digits_; rewrite spec_w_0W.
- rewrite Zmod_small;zarith.
- rewrite Pos2Z.inj_xO; split; auto with zarith.
- apply Z.le_lt_trans with (2 * Zpos w_digits); auto with zarith.
- unfold base, ww_digits; rewrite (Pos2Z.inj_xO w_digits).
- apply Zpower2_lt_lin; auto with zarith.
- Qed.
-
- Lemma spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
- let (q,r) := ww_div_gt a b in
- [[a]] = [[q]] * [[b]] + [[r]] /\
- 0 <= [[r]] < [[b]].
- Proof.
- intros a b Hgt Hpos;unfold ww_div_gt.
- change (let (q,r) := match a, b with
- | W0, _ => (W0,W0)
- | _, W0 => (W0,W0)
- | WW ah al, WW bh bl =>
- if w_eq0 ah then
- let (q,r) := w_div_gt al bl in
- (WW w_0 q, w_0W r)
- else
- match w_compare w_0 bh with
- | Eq =>
- let(q,r):=
- double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 a bl in
- (q, w_0W r)
- | Lt => ww_div_gt_aux ah al bh bl
- | Gt => (W0,W0) (* cas absurde *)
- end
- end in [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]).
- destruct a as [ |ah al]. simpl in Hgt;omega.
- destruct b as [ |bh bl]. simpl in Hpos;omega.
- Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl.
- assert (H:=@spec_eq0 ah);destruct (w_eq0 ah).
- simpl ww_to_Z;rewrite H;trivial. simpl in Hgt;rewrite H in Hgt;trivial.
- assert ([|bh|] <= 0).
- apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith.
- assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;rewrite H1;simpl in Hgt.
- simpl. simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos.
- assert (H2:=spec_div_gt Hgt Hpos);destruct (w_div_gt al bl).
- repeat rewrite spec_w_0W;simpl;rewrite spec_w_0;simpl;trivial.
- clear H.
- rewrite spec_compare; case Z.compare_spec; intros Hcmp.
- rewrite spec_w_0 in Hcmp. change [[WW bh bl]] with ([|bh|]*wB+[|bl|]).
- rewrite <- Hcmp;rewrite Z.mul_0_l;rewrite Z.add_0_l.
- simpl in Hpos;rewrite <- Hcmp in Hpos;simpl in Hpos.
- assert (H2:= @spec_double_divn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div
- w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0
- spec_add_mul_div spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hpos).
- destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
- w_compare w_sub 1
- (WW ah al) bl).
- rewrite spec_w_0W;unfold ww_to_Z;trivial.
- apply spec_ww_div_gt_aux;trivial. rewrite spec_w_0 in Hcmp;trivial.
- rewrite spec_w_0 in Hcmp;exfalso;omega.
- Qed.
-
- Lemma spec_ww_mod_gt_aux_eq : forall ah al bh bl,
- ww_mod_gt_aux ah al bh bl = snd (ww_div_gt_aux ah al bh bl).
- Proof.
- intros ah al bh bl. unfold ww_mod_gt_aux, ww_div_gt_aux.
- case w_compare; auto.
- case w_div32; auto.
- Qed.
-
- Lemma spec_ww_mod_gt_aux : forall ah al bh bl,
- [[WW ah al]] > [[WW bh bl]] ->
- 0 < [|bh|] ->
- [[ww_mod_gt_aux ah al bh bl]] = [[WW ah al]] mod [[WW bh bl]].
- Proof.
- intros. rewrite spec_ww_mod_gt_aux_eq;trivial.
- assert (H3 := spec_ww_div_gt_aux ah al bl H H0).
- destruct (ww_div_gt_aux ah al bh bl) as (q,r);simpl. simpl in H,H3.
- destruct H3;apply Zmod_unique with [[q]];zarith.
- rewrite H1;ring.
- Qed.
-
- Lemma spec_w_mod_gt_eq : forall a b, [|a|] > [|b|] -> 0 <[|b|] ->
- [|w_mod_gt a b|] = [|snd (w_div_gt a b)|].
- Proof.
- intros a b Hgt Hpos.
- rewrite spec_mod_gt;trivial.
- assert (H:=spec_div_gt Hgt Hpos).
- destruct (w_div_gt a b) as (q,r);simpl.
- rewrite Z.mul_comm in H;destruct H.
- symmetry;apply Zmod_unique with [|q|];trivial.
- Qed.
-
- Lemma spec_ww_mod_gt_eq : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
- [[ww_mod_gt a b]] = [[snd (ww_div_gt a b)]].
- Proof.
- intros a b Hgt Hpos.
- change (ww_mod_gt a b) with
- (match a, b with
- | W0, _ => W0
- | _, W0 => W0
- | WW ah al, WW bh bl =>
- if w_eq0 ah then w_0W (w_mod_gt al bl)
- else
- match w_compare w_0 bh with
- | Eq =>
- w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 a bl)
- | Lt => ww_mod_gt_aux ah al bh bl
- | Gt => W0 (* cas absurde *)
- end end).
- change (ww_div_gt a b) with
- (match a, b with
- | W0, _ => (W0,W0)
- | _, W0 => (W0,W0)
- | WW ah al, WW bh bl =>
- if w_eq0 ah then
- let (q,r) := w_div_gt al bl in
- (WW w_0 q, w_0W r)
- else
- match w_compare w_0 bh with
- | Eq =>
- let(q,r):=
- double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 a bl in
- (q, w_0W r)
- | Lt => ww_div_gt_aux ah al bh bl
- | Gt => (W0,W0) (* cas absurde *)
- end
- end).
- destruct a as [ |ah al];trivial.
- destruct b as [ |bh bl];trivial.
- Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl.
- assert (H:=@spec_eq0 ah);destruct (w_eq0 ah).
- simpl in Hgt;rewrite H in Hgt;trivial.
- assert ([|bh|] <= 0).
- apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith.
- assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt.
- simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos.
- rewrite spec_w_0W;rewrite spec_w_mod_gt_eq;trivial.
- destruct (w_div_gt al bl);simpl;rewrite spec_w_0W;trivial.
- clear H.
- rewrite spec_compare; case Z.compare_spec; intros H2.
- rewrite (@spec_double_modn1_aux w w_zdigits w_0 w_WW w_head0 w_add_mul_div
- w_div21 w_compare w_sub w_to_Z spec_w_0 spec_compare 1 (WW ah al) bl).
- destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1
- (WW ah al) bl);simpl;trivial.
- rewrite spec_ww_mod_gt_aux_eq;trivial;symmetry;trivial.
- trivial.
- Qed.
-
- Lemma spec_ww_mod_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
- [[ww_mod_gt a b]] = [[a]] mod [[b]].
- Proof.
- intros a b Hgt Hpos.
- assert (H:= spec_ww_div_gt a b Hgt Hpos).
- rewrite (spec_ww_mod_gt_eq a b Hgt Hpos).
- destruct (ww_div_gt a b)as(q,r);destruct H.
- apply Zmod_unique with[[q]];simpl;trivial.
- rewrite Z.mul_comm;trivial.
- Qed.
-
- Lemma Zis_gcd_mod : forall a b d,
- 0 < b -> Zis_gcd b (a mod b) d -> Zis_gcd a b d.
- Proof.
- intros a b d H H1; apply Zis_gcd_for_euclid with (a/b).
- pattern a at 1;rewrite (Z_div_mod_eq a b).
- ring_simplify (b * (a / b) + a mod b - a / b * b);trivial. zarith.
- Qed.
-
- Lemma spec_ww_gcd_gt_aux_body :
- forall ah al bh bl n cont,
- [[WW bh bl]] <= 2^n ->
- [[WW ah al]] > [[WW bh bl]] ->
- (forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^(n-1) ->
- Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) ->
- Zis_gcd [[WW ah al]] [[WW bh bl]] [[ww_gcd_gt_body cont ah al bh bl]].
- Proof.
- intros ah al bh bl n cont Hlog Hgt Hcont.
- change (ww_gcd_gt_body cont ah al bh bl) with (match w_compare w_0 bh with
- | Eq =>
- match w_compare w_0 bl with
- | Eq => WW ah al (* normalement n'arrive pas si forme normale *)
- | Lt =>
- let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 (WW ah al) bl in
- WW w_0 (w_gcd_gt bl m)
- | Gt => W0 (* absurde *)
- end
- | Lt =>
- let m := ww_mod_gt_aux ah al bh bl in
- match m with
- | W0 => WW bh bl
- | WW mh ml =>
- match w_compare w_0 mh with
- | Eq =>
- match w_compare w_0 ml with
- | Eq => WW bh bl
- | _ =>
- let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
- w_compare w_sub 1 (WW bh bl) ml in
- WW w_0 (w_gcd_gt ml r)
- end
- | Lt =>
- let r := ww_mod_gt_aux bh bl mh ml in
- match r with
- | W0 => m
- | WW rh rl => cont mh ml rh rl
- end
- | Gt => W0 (* absurde *)
- end
- end
- | Gt => W0 (* absurde *)
- end).
- rewrite spec_compare, spec_w_0.
- case Z.compare_spec; intros Hbh.
- simpl ww_to_Z in *. rewrite <- Hbh.
- rewrite Z.mul_0_l;rewrite Z.add_0_l.
- rewrite spec_compare, spec_w_0.
- case Z.compare_spec; intros Hbl.
- rewrite <- Hbl;apply Zis_gcd_0.
- simpl;rewrite spec_w_0;rewrite Z.mul_0_l;rewrite Z.add_0_l.
- apply Zis_gcd_mod;zarith.
- change ([|ah|] * wB + [|al|]) with (double_to_Z w_digits w_to_Z 1 (WW ah al)).
- rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div
- w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div
- spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hbl).
- apply spec_gcd_gt.
- rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial.
- apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
- destruct (Z_mod_lt x y);zarith end.
- Spec_w_to_Z bl;exfalso;omega.
- assert (H:= spec_ww_mod_gt_aux _ _ _ Hgt Hbh).
- assert (H2 : 0 < [[WW bh bl]]).
- simpl;Spec_w_to_Z bl. apply Z.lt_le_trans with ([|bh|]*wB);zarith.
- apply Z.mul_pos_pos;zarith.
- apply Zis_gcd_mod;trivial. rewrite <- H.
- simpl in *;destruct (ww_mod_gt_aux ah al bh bl) as [ |mh ml].
- simpl;apply Zis_gcd_0;zarith.
- rewrite spec_compare, spec_w_0; case Z.compare_spec; intros Hmh.
- simpl;rewrite <- Hmh;simpl.
- rewrite spec_compare, spec_w_0; case Z.compare_spec; intros Hml.
- rewrite <- Hml;simpl;apply Zis_gcd_0.
- simpl; rewrite spec_w_0; simpl.
- apply Zis_gcd_mod;zarith.
- change ([|bh|] * wB + [|bl|]) with (double_to_Z w_digits w_to_Z 1 (WW bh bl)).
- rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div
- w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div
- spec_div21 spec_compare spec_sub 1 (WW bh bl) ml Hml).
- apply spec_gcd_gt.
- rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial.
- apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
- destruct (Z_mod_lt x y);zarith end.
- Spec_w_to_Z ml;exfalso;omega.
- assert ([[WW bh bl]] > [[WW mh ml]]).
- rewrite H;simpl; apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
- destruct (Z_mod_lt x y);zarith end.
- assert (H1:= spec_ww_mod_gt_aux _ _ _ H0 Hmh).
- assert (H3 : 0 < [[WW mh ml]]).
- simpl;Spec_w_to_Z ml. apply Z.lt_le_trans with ([|mh|]*wB);zarith.
- apply Z.mul_pos_pos;zarith.
- apply Zis_gcd_mod;zarith. simpl in *;rewrite <- H1.
- destruct (ww_mod_gt_aux bh bl mh ml) as [ |rh rl]. simpl; apply Zis_gcd_0.
- simpl;apply Hcont. simpl in H1;rewrite H1.
- apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
- destruct (Z_mod_lt x y);zarith end.
- apply Z.le_trans with (2^n/2).
- apply Zdiv_le_lower_bound;zarith.
- apply Z.le_trans with ([|bh|] * wB + [|bl|]);zarith.
- assert (H3' := Z_div_mod_eq [[WW bh bl]] [[WW mh ml]] (Z.lt_gt _ _ H3)).
- assert (H4 : 0 <= [[WW bh bl]]/[[WW mh ml]]).
- apply Z.ge_le;apply Z_div_ge0;zarith. simpl in *;rewrite H1.
- pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3'.
- Z.le_elim H4.
- assert (H6' : [[WW bh bl]] mod [[WW mh ml]] =
- [[WW bh bl]] - [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])).
- simpl;pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3';ring. simpl in H6'.
- assert ([[WW mh ml]] <= [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])).
- simpl;pattern ([|mh|]*wB+[|ml|]) at 1;rewrite <- Z.mul_1_r;zarith.
- simpl in *;assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in H8;
- zarith.
- assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in *;zarith.
- rewrite <- H4 in H3';rewrite Z.mul_0_r in H3';simpl in H3';zarith.
- pattern n at 1;replace n with (n-1+1);try ring.
- rewrite Zpower_exp;zarith. change (2^1) with 2.
- rewrite Z_div_mult;zarith.
- assert (2^1 <= 2^n). change (2^1) with 2;zarith.
- assert (H7 := @Zpower_le_monotone_inv 2 1 n);zarith.
- Spec_w_to_Z mh;exfalso;zarith.
- Spec_w_to_Z bh;exfalso;zarith.
- Qed.
-
- Lemma spec_ww_gcd_gt_aux :
- forall p cont n,
- (forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
- [[WW yh yl]] <= 2^n ->
- Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) ->
- forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] ->
- [[WW bh bl]] <= 2^(Zpos p + n) ->
- Zis_gcd [[WW ah al]] [[WW bh bl]]
- [[ww_gcd_gt_aux p cont ah al bh bl]].
- Proof.
- induction p;intros cont n Hcont ah al bh bl Hgt Hs;simpl ww_gcd_gt_aux.
- assert (0 < Zpos p). unfold Z.lt;reflexivity.
- apply spec_ww_gcd_gt_aux_body with (n := Zpos (xI p) + n);
- trivial;rewrite Pos2Z.inj_xI.
- intros. apply IHp with (n := Zpos p + n);zarith.
- intros. apply IHp with (n := n );zarith.
- apply Z.le_trans with (2 ^ (2* Zpos p + 1+ n -1));zarith.
- apply Z.pow_le_mono_r;zarith.
- assert (0 < Zpos p). unfold Z.lt;reflexivity.
- apply spec_ww_gcd_gt_aux_body with (n := Zpos (xO p) + n );trivial.
- rewrite (Pos2Z.inj_xO p).
- intros. apply IHp with (n := Zpos p + n - 1);zarith.
- intros. apply IHp with (n := n -1 );zarith.
- intros;apply Hcont;zarith.
- apply Z.le_trans with (2^(n-1));zarith.
- apply Z.pow_le_mono_r;zarith.
- apply Z.le_trans with (2 ^ (Zpos p + n -1));zarith.
- apply Z.pow_le_mono_r;zarith.
- apply Z.le_trans with (2 ^ (2*Zpos p + n -1));zarith.
- apply Z.pow_le_mono_r;zarith.
- apply spec_ww_gcd_gt_aux_body with (n := n+1);trivial.
- rewrite Z.add_comm;trivial.
- ring_simplify (n + 1 - 1);trivial.
- Qed.
-
-End DoubleDivGt.
-
-Section DoubleDiv.
-
- Variable w : Type.
- Variable w_digits : positive.
- Variable ww_1 : zn2z w.
- Variable ww_compare : zn2z w -> zn2z w -> comparison.
-
- Variable ww_div_gt : zn2z w -> zn2z w -> zn2z w * zn2z w.
- Variable ww_mod_gt : zn2z w -> zn2z w -> zn2z w.
-
- Definition ww_div a b :=
- match ww_compare a b with
- | Gt => ww_div_gt a b
- | Eq => (ww_1, W0)
- | Lt => (W0, a)
- end.
-
- Definition ww_mod a b :=
- match ww_compare a b with
- | Gt => ww_mod_gt a b
- | Eq => W0
- | Lt => a
- end.
-
- Variable w_to_Z : w -> Z.
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_ww_1 : [[ww_1]] = 1.
- Variable spec_ww_compare : forall x y,
- ww_compare x y = Z.compare [[x]] [[y]].
- Variable spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
- let (q,r) := ww_div_gt a b in
- [[a]] = [[q]] * [[b]] + [[r]] /\
- 0 <= [[r]] < [[b]].
- Variable spec_ww_mod_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
- [[ww_mod_gt a b]] = [[a]] mod [[b]].
-
- Ltac Spec_w_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_to_Z x).
-
- Ltac Spec_ww_to_Z x :=
- let H:= fresh "HH" in
- assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x).
-
- Lemma spec_ww_div : forall a b, 0 < [[b]] ->
- let (q,r) := ww_div a b in
- [[a]] = [[q]] * [[b]] + [[r]] /\
- 0 <= [[r]] < [[b]].
- Proof.
- intros a b Hpos;unfold ww_div.
- rewrite spec_ww_compare; case Z.compare_spec; intros.
- simpl;rewrite spec_ww_1;split;zarith.
- simpl;split;[ring|Spec_ww_to_Z a;zarith].
- apply spec_ww_div_gt;auto with zarith.
- Qed.
-
- Lemma spec_ww_mod : forall a b, 0 < [[b]] ->
- [[ww_mod a b]] = [[a]] mod [[b]].
- Proof.
- intros a b Hpos;unfold ww_mod.
- rewrite spec_ww_compare; case Z.compare_spec; intros.
- simpl;apply Zmod_unique with 1;try rewrite H;zarith.
- Spec_ww_to_Z a;symmetry;apply Zmod_small;zarith.
- apply spec_ww_mod_gt;auto with zarith.
- Qed.
-
-
- Variable w_0 : w.
- Variable w_1 : w.
- Variable w_compare : w -> w -> comparison.
- Variable w_eq0 : w -> bool.
- Variable w_gcd_gt : w -> w -> w.
- Variable _ww_digits : positive.
- Variable spec_ww_digits_ : _ww_digits = xO w_digits.
- Variable ww_gcd_gt_fix :
- positive -> (w -> w -> w -> w -> zn2z w) ->
- w -> w -> w -> w -> zn2z w.
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_1 : [|w_1|] = 1.
- Variable spec_compare :
- forall x y, w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0.
- Variable spec_gcd_gt : forall a b, [|a|] > [|b|] ->
- Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|].
- Variable spec_gcd_gt_fix :
- forall p cont n,
- (forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
- [[WW yh yl]] <= 2^n ->
- Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) ->
- forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] ->
- [[WW bh bl]] <= 2^(Zpos p + n) ->
- Zis_gcd [[WW ah al]] [[WW bh bl]]
- [[ww_gcd_gt_fix p cont ah al bh bl]].
-
- Definition gcd_cont (xh xl yh yl:w) :=
- match w_compare w_1 yl with
- | Eq => ww_1
- | _ => WW xh xl
- end.
-
- Lemma spec_gcd_cont : forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
- [[WW yh yl]] <= 1 ->
- Zis_gcd [[WW xh xl]] [[WW yh yl]] [[gcd_cont xh xl yh yl]].
- Proof.
- intros xh xl yh yl Hgt' Hle. simpl in Hle.
- assert ([|yh|] = 0).
- change 1 with (0*wB+1) in Hle.
- assert (0 <= 1 < wB). split;zarith. apply wB_pos.
- assert (H1:= beta_lex _ _ _ _ _ Hle (spec_to_Z yl) H).
- Spec_w_to_Z yh;zarith.
- unfold gcd_cont; rewrite spec_compare, spec_w_1.
- case Z.compare_spec; intros Hcmpy.
- simpl;rewrite H;simpl;
- rewrite spec_ww_1;rewrite <- Hcmpy;apply Zis_gcd_mod;zarith.
- rewrite <- (Zmod_unique ([|xh|]*wB+[|xl|]) 1 ([|xh|]*wB+[|xl|]) 0);zarith.
- rewrite H in Hle; exfalso;zarith.
- assert (H0 : [|yl|] = 0) by (Spec_w_to_Z yl;zarith).
- simpl. rewrite H0, H;simpl;apply Zis_gcd_0;trivial.
- Qed.
-
-
- Variable cont : w -> w -> w -> w -> zn2z w.
- Variable spec_cont : forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
- [[WW yh yl]] <= 1 ->
- Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]].
-
- Definition ww_gcd_gt a b :=
- match a, b with
- | W0, _ => b
- | _, W0 => a
- | WW ah al, WW bh bl =>
- if w_eq0 ah then (WW w_0 (w_gcd_gt al bl))
- else ww_gcd_gt_fix _ww_digits cont ah al bh bl
- end.
-
- Definition ww_gcd a b :=
- Eval lazy beta delta [ww_gcd_gt] in
- match ww_compare a b with
- | Gt => ww_gcd_gt a b
- | Eq => a
- | Lt => ww_gcd_gt b a
- end.
-
- Lemma spec_ww_gcd_gt : forall a b, [[a]] > [[b]] ->
- Zis_gcd [[a]] [[b]] [[ww_gcd_gt a b]].
- Proof.
- intros a b Hgt;unfold ww_gcd_gt.
- destruct a as [ |ah al]. simpl;apply Zis_gcd_sym;apply Zis_gcd_0.
- destruct b as [ |bh bl]. simpl;apply Zis_gcd_0.
- simpl in Hgt. generalize (@spec_eq0 ah);destruct (w_eq0 ah);intros.
- simpl;rewrite H in Hgt;trivial;rewrite H;trivial;rewrite spec_w_0;simpl.
- assert ([|bh|] <= 0).
- apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith.
- Spec_w_to_Z bh;assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt.
- rewrite H1;simpl;auto. clear H.
- apply spec_gcd_gt_fix with (n:= 0);trivial.
- rewrite Z.add_0_r;rewrite spec_ww_digits_.
- change (2 ^ Zpos (xO w_digits)) with wwB. Spec_ww_to_Z (WW bh bl);zarith.
- Qed.
-
- Lemma spec_ww_gcd : forall a b, Zis_gcd [[a]] [[b]] [[ww_gcd a b]].
- Proof.
- intros a b.
- change (ww_gcd a b) with
- (match ww_compare a b with
- | Gt => ww_gcd_gt a b
- | Eq => a
- | Lt => ww_gcd_gt b a
- end).
- rewrite spec_ww_compare; case Z.compare_spec; intros Hcmp.
- Spec_ww_to_Z b;rewrite Hcmp.
- apply Zis_gcd_for_euclid with 1;zarith.
- ring_simplify ([[b]] - 1 * [[b]]). apply Zis_gcd_0;zarith.
- apply Zis_gcd_sym;apply spec_ww_gcd_gt;zarith.
- apply spec_ww_gcd_gt;zarith.
- Qed.
-
-End DoubleDiv.
-
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
deleted file mode 100644
index 195527dd5..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
+++ /dev/null
@@ -1,519 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith Ndigits.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-
-Local Open Scope Z_scope.
-
-Local Infix "<<" := Pos.shiftl_nat (at level 30).
-
-Section GENDIVN1.
-
- Variable w : Type.
- Variable w_digits : positive.
- Variable w_zdigits : w.
- Variable w_0 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_head0 : w -> w.
- Variable w_add_mul_div : w -> w -> w -> w.
- Variable w_div21 : w -> w -> w -> w * w.
- Variable w_compare : w -> w -> comparison.
- Variable w_sub : w -> w -> w.
-
-
-
- (* ** For proofs ** *)
- Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
-
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x)
- (at level 0, x at level 99).
- Notation "[[ x ]]" := (zn2z_to_Z wB w_to_Z x) (at level 0, x at level 99).
-
- Variable spec_to_Z : forall x, 0 <= [| x |] < wB.
- Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits.
- Variable spec_0 : [|w_0|] = 0.
- Variable spec_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_head0 : forall x, 0 < [|x|] ->
- wB/ 2 <= 2 ^ [|w_head0 x|] * [|x|] < wB.
- Variable spec_add_mul_div : forall x y p,
- [|p|] <= Zpos w_digits ->
- [| w_add_mul_div p x y |] =
- ([|x|] * (2 ^ [|p|]) +
- [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB.
- Variable spec_div21 : forall a1 a2 b,
- wB/2 <= [|b|] ->
- [|a1|] < [|b|] ->
- let (q,r) := w_div21 a1 a2 b in
- [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
- Variable spec_compare :
- forall x y, w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_sub: forall x y,
- [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
-
-
-
- Section DIVAUX.
- Variable b2p : w.
- Variable b2p_le : wB/2 <= [|b2p|].
-
- Definition double_divn1_0_aux n (divn1: w -> word w n -> word w n * w) r h :=
- let (hh,hl) := double_split w_0 n h in
- let (qh,rh) := divn1 r hh in
- let (ql,rl) := divn1 rh hl in
- (double_WW w_WW n qh ql, rl).
-
- Fixpoint double_divn1_0 (n:nat) : w -> word w n -> word w n * w :=
- match n return w -> word w n -> word w n * w with
- | O => fun r x => w_div21 r x b2p
- | S n => double_divn1_0_aux n (double_divn1_0 n)
- end.
-
- Lemma spec_split : forall (n : nat) (x : zn2z (word w n)),
- let (h, l) := double_split w_0 n x in
- [!S n | x!] = [!n | h!] * double_wB w_digits n + [!n | l!].
- Proof (spec_double_split w_0 w_digits w_to_Z spec_0).
-
- Lemma spec_double_divn1_0 : forall n r a,
- [|r|] < [|b2p|] ->
- let (q,r') := double_divn1_0 n r a in
- [|r|] * double_wB w_digits n + [!n|a!] = [!n|q!] * [|b2p|] + [|r'|] /\
- 0 <= [|r'|] < [|b2p|].
- Proof.
- induction n;intros.
- exact (spec_div21 a b2p_le H).
- simpl (double_divn1_0 (S n) r a); unfold double_divn1_0_aux.
- assert (H1 := spec_split n a);destruct (double_split w_0 n a) as (hh,hl).
- rewrite H1.
- assert (H2 := IHn r hh H);destruct (double_divn1_0 n r hh) as (qh,rh).
- destruct H2.
- assert ([|rh|] < [|b2p|]). omega.
- assert (H4 := IHn rh hl H3);destruct (double_divn1_0 n rh hl) as (ql,rl).
- destruct H4;split;trivial.
- rewrite spec_double_WW;trivial.
- rewrite <- double_wB_wwB.
- rewrite Z.mul_assoc;rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- rewrite H0;rewrite Z.mul_add_distr_r;rewrite <- Z.add_assoc.
- rewrite H4;ring.
- Qed.
-
- Definition double_modn1_0_aux n (modn1:w -> word w n -> w) r h :=
- let (hh,hl) := double_split w_0 n h in modn1 (modn1 r hh) hl.
-
- Fixpoint double_modn1_0 (n:nat) : w -> word w n -> w :=
- match n return w -> word w n -> w with
- | O => fun r x => snd (w_div21 r x b2p)
- | S n => double_modn1_0_aux n (double_modn1_0 n)
- end.
-
- Lemma spec_double_modn1_0 : forall n r x,
- double_modn1_0 n r x = snd (double_divn1_0 n r x).
- Proof.
- induction n;simpl;intros;trivial.
- unfold double_modn1_0_aux, double_divn1_0_aux.
- destruct (double_split w_0 n x) as (hh,hl).
- rewrite (IHn r hh).
- destruct (double_divn1_0 n r hh) as (qh,rh);simpl.
- rewrite IHn. destruct (double_divn1_0 n rh hl);trivial.
- Qed.
-
- Variable p : w.
- Variable p_bounded : [|p|] <= Zpos w_digits.
-
- Lemma spec_add_mul_divp : forall x y,
- [| w_add_mul_div p x y |] =
- ([|x|] * (2 ^ [|p|]) +
- [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB.
- Proof.
- intros;apply spec_add_mul_div;auto.
- Qed.
-
- Definition double_divn1_p_aux n
- (divn1 : w -> word w n -> word w n -> word w n * w) r h l :=
- let (hh,hl) := double_split w_0 n h in
- let (lh,ll) := double_split w_0 n l in
- let (qh,rh) := divn1 r hh hl in
- let (ql,rl) := divn1 rh hl lh in
- (double_WW w_WW n qh ql, rl).
-
- Fixpoint double_divn1_p (n:nat) : w -> word w n -> word w n -> word w n * w :=
- match n return w -> word w n -> word w n -> word w n * w with
- | O => fun r h l => w_div21 r (w_add_mul_div p h l) b2p
- | S n => double_divn1_p_aux n (double_divn1_p n)
- end.
-
- Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (w_digits << n).
- Proof.
- induction n;simpl. trivial.
- case (spec_to_Z p); rewrite Pos2Z.inj_xO;auto with zarith.
- Qed.
-
- Lemma spec_double_divn1_p : forall n r h l,
- [|r|] < [|b2p|] ->
- let (q,r') := double_divn1_p n r h l in
- [|r|] * double_wB w_digits n +
- ([!n|h!]*2^[|p|] +
- [!n|l!] / (2^(Zpos(w_digits << n) - [|p|])))
- mod double_wB w_digits n = [!n|q!] * [|b2p|] + [|r'|] /\
- 0 <= [|r'|] < [|b2p|].
- Proof.
- case (spec_to_Z p); intros HH0 HH1.
- induction n;intros.
- simpl (double_divn1_p 0 r h l).
- unfold double_to_Z, double_wB, "<<".
- rewrite <- spec_add_mul_divp.
- exact (spec_div21 (w_add_mul_div p h l) b2p_le H).
- simpl (double_divn1_p (S n) r h l).
- unfold double_divn1_p_aux.
- assert (H1 := spec_split n h);destruct (double_split w_0 n h) as (hh,hl).
- rewrite H1. rewrite <- double_wB_wwB.
- assert (H2 := spec_split n l);destruct (double_split w_0 n l) as (lh,ll).
- rewrite H2.
- replace ([|r|] * (double_wB w_digits n * double_wB w_digits n) +
- (([!n|hh!] * double_wB w_digits n + [!n|hl!]) * 2 ^ [|p|] +
- ([!n|lh!] * double_wB w_digits n + [!n|ll!]) /
- 2^(Zpos (w_digits << (S n)) - [|p|])) mod
- (double_wB w_digits n * double_wB w_digits n)) with
- (([|r|] * double_wB w_digits n + ([!n|hh!] * 2^[|p|] +
- [!n|hl!] / 2^(Zpos (w_digits << n) - [|p|])) mod
- double_wB w_digits n) * double_wB w_digits n +
- ([!n|hl!] * 2^[|p|] +
- [!n|lh!] / 2^(Zpos (w_digits << n) - [|p|])) mod
- double_wB w_digits n).
- generalize (IHn r hh hl H);destruct (double_divn1_p n r hh hl) as (qh,rh);
- intros (H3,H4);rewrite H3.
- assert ([|rh|] < [|b2p|]). omega.
- replace (([!n|qh!] * [|b2p|] + [|rh|]) * double_wB w_digits n +
- ([!n|hl!] * 2 ^ [|p|] +
- [!n|lh!] / 2 ^ (Zpos (w_digits << n) - [|p|])) mod
- double_wB w_digits n) with
- ([!n|qh!] * [|b2p|] *double_wB w_digits n + ([|rh|]*double_wB w_digits n +
- ([!n|hl!] * 2 ^ [|p|] +
- [!n|lh!] / 2 ^ (Zpos (w_digits << n) - [|p|])) mod
- double_wB w_digits n)). 2:ring.
- generalize (IHn rh hl lh H0);destruct (double_divn1_p n rh hl lh) as (ql,rl);
- intros (H5,H6);rewrite H5.
- split;[rewrite spec_double_WW;trivial;ring|trivial].
- assert (Uhh := spec_double_to_Z w_digits w_to_Z spec_to_Z n hh);
- unfold double_wB,base in Uhh.
- assert (Uhl := spec_double_to_Z w_digits w_to_Z spec_to_Z n hl);
- unfold double_wB,base in Uhl.
- assert (Ulh := spec_double_to_Z w_digits w_to_Z spec_to_Z n lh);
- unfold double_wB,base in Ulh.
- assert (Ull := spec_double_to_Z w_digits w_to_Z spec_to_Z n ll);
- unfold double_wB,base in Ull.
- unfold double_wB,base.
- assert (UU:=p_lt_double_digits n).
- rewrite Zdiv_shift_r;auto with zarith.
- 2:change (Zpos (w_digits << (S n)))
- with (2*Zpos (w_digits << n));auto with zarith.
- replace (2 ^ (Zpos (w_digits << (S n)) - [|p|])) with
- (2^(Zpos (w_digits << n) - [|p|])*2^Zpos (w_digits << n)).
- rewrite Zdiv_mult_cancel_r;auto with zarith.
- rewrite Z.mul_add_distr_r with (p:= 2^[|p|]).
- pattern ([!n|hl!] * 2^[|p|]) at 2;
- rewrite (shift_unshift_mod (Zpos(w_digits << n))([|p|])([!n|hl!]));
- auto with zarith.
- rewrite Z.add_assoc.
- replace
- ([!n|hh!] * 2^Zpos (w_digits << n)* 2^[|p|] +
- ([!n|hl!] / 2^(Zpos (w_digits << n)-[|p|])*
- 2^Zpos(w_digits << n)))
- with
- (([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl /
- 2^(Zpos (w_digits << n)-[|p|]))
- * 2^Zpos(w_digits << n));try (ring;fail).
- rewrite <- Z.add_assoc.
- rewrite <- (Zmod_shift_r ([|p|]));auto with zarith.
- replace
- (2 ^ Zpos (w_digits << n) * 2 ^ Zpos (w_digits << n)) with
- (2 ^ (Zpos (w_digits << n) + Zpos (w_digits << n))).
- rewrite (Zmod_shift_r (Zpos (w_digits << n)));auto with zarith.
- replace (2 ^ (Zpos (w_digits << n) + Zpos (w_digits << n)))
- with (2^Zpos(w_digits << n) *2^Zpos(w_digits << n)).
- rewrite (Z.mul_comm (([!n|hh!] * 2 ^ [|p|] +
- [!n|hl!] / 2 ^ (Zpos (w_digits << n) - [|p|])))).
- rewrite Zmult_mod_distr_l;auto with zarith.
- ring.
- rewrite Zpower_exp;auto with zarith.
- assert (0 < Zpos (w_digits << n)). unfold Z.lt;reflexivity.
- auto with zarith.
- apply Z_mod_lt;auto with zarith.
- rewrite Zpower_exp;auto with zarith.
- split;auto with zarith.
- apply Zdiv_lt_upper_bound;auto with zarith.
- rewrite <- Zpower_exp;auto with zarith.
- replace ([|p|] + (Zpos (w_digits << n) - [|p|])) with
- (Zpos(w_digits << n));auto with zarith.
- rewrite <- Zpower_exp;auto with zarith.
- replace (Zpos (w_digits << (S n)) - [|p|]) with
- (Zpos (w_digits << n) - [|p|] +
- Zpos (w_digits << n));trivial.
- change (Zpos (w_digits << (S n))) with
- (2*Zpos (w_digits << n)). ring.
- Qed.
-
- Definition double_modn1_p_aux n (modn1 : w -> word w n -> word w n -> w) r h l:=
- let (hh,hl) := double_split w_0 n h in
- let (lh,ll) := double_split w_0 n l in
- modn1 (modn1 r hh hl) hl lh.
-
- Fixpoint double_modn1_p (n:nat) : w -> word w n -> word w n -> w :=
- match n return w -> word w n -> word w n -> w with
- | O => fun r h l => snd (w_div21 r (w_add_mul_div p h l) b2p)
- | S n => double_modn1_p_aux n (double_modn1_p n)
- end.
-
- Lemma spec_double_modn1_p : forall n r h l ,
- double_modn1_p n r h l = snd (double_divn1_p n r h l).
- Proof.
- induction n;simpl;intros;trivial.
- unfold double_modn1_p_aux, double_divn1_p_aux.
- destruct(double_split w_0 n h)as(hh,hl);destruct(double_split w_0 n l) as (lh,ll).
- rewrite (IHn r hh hl);destruct (double_divn1_p n r hh hl) as (qh,rh).
- rewrite IHn;simpl;destruct (double_divn1_p n rh hl lh);trivial.
- Qed.
-
- End DIVAUX.
-
- Fixpoint high (n:nat) : word w n -> w :=
- match n return word w n -> w with
- | O => fun a => a
- | S n =>
- fun (a:zn2z (word w n)) =>
- match a with
- | W0 => w_0
- | WW h l => high n h
- end
- end.
-
- Lemma spec_double_digits:forall n, Zpos w_digits <= Zpos (w_digits << n).
- Proof.
- induction n;simpl;auto with zarith.
- change (Zpos (xO (w_digits << n))) with
- (2*Zpos (w_digits << n)).
- assert (0 < Zpos w_digits) by reflexivity.
- auto with zarith.
- Qed.
-
- Lemma spec_high : forall n (x:word w n),
- [|high n x|] = [!n|x!] / 2^(Zpos (w_digits << n) - Zpos w_digits).
- Proof.
- induction n;intros.
- unfold high,double_to_Z. rewrite Pshiftl_nat_0.
- replace (Zpos w_digits - Zpos w_digits) with 0;try ring.
- simpl. rewrite <- (Zdiv_unique [|x|] 1 [|x|] 0);auto with zarith.
- assert (U2 := spec_double_digits n).
- assert (U3 : 0 < Zpos w_digits). exact (eq_refl Lt).
- destruct x;unfold high;fold high.
- unfold double_to_Z,zn2z_to_Z;rewrite spec_0.
- rewrite Zdiv_0_l;trivial.
- assert (U0 := spec_double_to_Z w_digits w_to_Z spec_to_Z n w0);
- assert (U1 := spec_double_to_Z w_digits w_to_Z spec_to_Z n w1).
- simpl [!S n|WW w0 w1!].
- unfold double_wB,base;rewrite Zdiv_shift_r;auto with zarith.
- replace (2 ^ (Zpos (w_digits << (S n)) - Zpos w_digits)) with
- (2^(Zpos (w_digits << n) - Zpos w_digits) *
- 2^Zpos (w_digits << n)).
- rewrite Zdiv_mult_cancel_r;auto with zarith.
- rewrite <- Zpower_exp;auto with zarith.
- replace (Zpos (w_digits << n) - Zpos w_digits +
- Zpos (w_digits << n)) with
- (Zpos (w_digits << (S n)) - Zpos w_digits);trivial.
- change (Zpos (w_digits << (S n))) with
- (2*Zpos (w_digits << n));ring.
- change (Zpos (w_digits << (S n))) with
- (2*Zpos (w_digits << n)); auto with zarith.
- Qed.
-
- Definition double_divn1 (n:nat) (a:word w n) (b:w) :=
- let p := w_head0 b in
- match w_compare p w_0 with
- | Gt =>
- let b2p := w_add_mul_div p b w_0 in
- let ha := high n a in
- let k := w_sub w_zdigits p in
- let lsr_n := w_add_mul_div k w_0 in
- let r0 := w_add_mul_div p w_0 ha in
- let (q,r) := double_divn1_p b2p p n r0 a (double_0 w_0 n) in
- (q, lsr_n r)
- | _ => double_divn1_0 b n w_0 a
- end.
-
- Lemma spec_double_divn1 : forall n a b,
- 0 < [|b|] ->
- let (q,r) := double_divn1 n a b in
- [!n|a!] = [!n|q!] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
- Proof.
- intros n a b H. unfold double_divn1.
- case (spec_head0 H); intros H0 H1.
- case (spec_to_Z (w_head0 b)); intros HH1 HH2.
- rewrite spec_compare; case Z.compare_spec;
- rewrite spec_0; intros H2; auto with zarith.
- assert (Hv1: wB/2 <= [|b|]).
- generalize H0; rewrite H2; rewrite Z.pow_0_r;
- rewrite Z.mul_1_l; auto.
- assert (Hv2: [|w_0|] < [|b|]).
- rewrite spec_0; auto.
- generalize (spec_double_divn1_0 Hv1 n a Hv2).
- rewrite spec_0;rewrite Z.mul_0_l; rewrite Z.add_0_l; auto.
- contradict H2; auto with zarith.
- assert (HHHH : 0 < [|w_head0 b|]); auto with zarith.
- assert ([|w_head0 b|] < Zpos w_digits).
- case (Z.le_gt_cases (Zpos w_digits) [|w_head0 b|]); auto; intros HH.
- assert (2 ^ [|w_head0 b|] < wB).
- apply Z.le_lt_trans with (2 ^ [|w_head0 b|] * [|b|]);auto with zarith.
- replace (2 ^ [|w_head0 b|]) with (2^[|w_head0 b|] * 1);try (ring;fail).
- apply Z.mul_le_mono_nonneg;auto with zarith.
- assert (wB <= 2^[|w_head0 b|]).
- unfold base;apply Zpower_le_monotone;auto with zarith. omega.
- assert ([|w_add_mul_div (w_head0 b) b w_0|] =
- 2 ^ [|w_head0 b|] * [|b|]).
- rewrite (spec_add_mul_div b w_0); auto with zarith.
- rewrite spec_0;rewrite Zdiv_0_l; try omega.
- rewrite Z.add_0_r; rewrite Z.mul_comm.
- rewrite Zmod_small; auto with zarith.
- assert (H5 := spec_to_Z (high n a)).
- assert
- ([|w_add_mul_div (w_head0 b) w_0 (high n a)|]
- <[|w_add_mul_div (w_head0 b) b w_0|]).
- rewrite H4.
- rewrite spec_add_mul_div;auto with zarith.
- rewrite spec_0;rewrite Z.mul_0_l;rewrite Z.add_0_l.
- assert (([|high n a|]/2^(Zpos w_digits - [|w_head0 b|])) < wB).
- apply Zdiv_lt_upper_bound;auto with zarith.
- apply Z.lt_le_trans with wB;auto with zarith.
- pattern wB at 1;replace wB with (wB*1);try ring.
- apply Z.mul_le_mono_nonneg;auto with zarith.
- assert (H6 := Z.pow_pos_nonneg 2 (Zpos w_digits - [|w_head0 b|]));
- auto with zarith.
- rewrite Zmod_small;auto with zarith.
- apply Zdiv_lt_upper_bound;auto with zarith.
- apply Z.lt_le_trans with wB;auto with zarith.
- apply Z.le_trans with (2 ^ [|w_head0 b|] * [|b|] * 2).
- rewrite <- wB_div_2; try omega.
- apply Z.mul_le_mono_nonneg;auto with zarith.
- pattern 2 at 1;rewrite <- Z.pow_1_r.
- apply Zpower_le_monotone;split;auto with zarith.
- rewrite <- H4 in H0.
- assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith.
- assert (H7:= spec_double_divn1_p H0 Hb3 n a (double_0 w_0 n) H6).
- destruct (double_divn1_p (w_add_mul_div (w_head0 b) b w_0) (w_head0 b) n
- (w_add_mul_div (w_head0 b) w_0 (high n a)) a
- (double_0 w_0 n)) as (q,r).
- assert (U:= spec_double_digits n).
- rewrite spec_double_0 in H7;trivial;rewrite Zdiv_0_l in H7.
- rewrite Z.add_0_r in H7.
- rewrite spec_add_mul_div in H7;auto with zarith.
- rewrite spec_0 in H7;rewrite Z.mul_0_l in H7;rewrite Z.add_0_l in H7.
- assert (([|high n a|] / 2 ^ (Zpos w_digits - [|w_head0 b|])) mod wB
- = [!n|a!] / 2^(Zpos (w_digits << n) - [|w_head0 b|])).
- rewrite Zmod_small;auto with zarith.
- rewrite spec_high. rewrite Zdiv_Zdiv;auto with zarith.
- rewrite <- Zpower_exp;auto with zarith.
- replace (Zpos (w_digits << n) - Zpos w_digits +
- (Zpos w_digits - [|w_head0 b|]))
- with (Zpos (w_digits << n) - [|w_head0 b|]);trivial;ring.
- assert (H8 := Z.pow_pos_nonneg 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith.
- split;auto with zarith.
- apply Z.le_lt_trans with ([|high n a|]);auto with zarith.
- apply Zdiv_le_upper_bound;auto with zarith.
- pattern ([|high n a|]) at 1;rewrite <- Z.mul_1_r.
- apply Z.mul_le_mono_nonneg;auto with zarith.
- rewrite H8 in H7;unfold double_wB,base in H7.
- rewrite <- shift_unshift_mod in H7;auto with zarith.
- rewrite H4 in H7.
- assert ([|w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r|]
- = [|r|]/2^[|w_head0 b|]).
- rewrite spec_add_mul_div.
- rewrite spec_0;rewrite Z.mul_0_l;rewrite Z.add_0_l.
- replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|])
- with ([|w_head0 b|]).
- rewrite Zmod_small;auto with zarith.
- assert (H9 := spec_to_Z r).
- split;auto with zarith.
- apply Z.le_lt_trans with ([|r|]);auto with zarith.
- apply Zdiv_le_upper_bound;auto with zarith.
- pattern ([|r|]) at 1;rewrite <- Z.mul_1_r.
- apply Z.mul_le_mono_nonneg;auto with zarith.
- assert (H10 := Z.pow_pos_nonneg 2 ([|w_head0 b|]));auto with zarith.
- rewrite spec_sub.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- case (spec_to_Z w_zdigits); auto with zarith.
- rewrite spec_sub.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- case (spec_to_Z w_zdigits); auto with zarith.
- case H7; intros H71 H72.
- split.
- rewrite <- (Z_div_mult [!n|a!] (2^[|w_head0 b|]));auto with zarith.
- rewrite H71;rewrite H9.
- replace ([!n|q!] * (2 ^ [|w_head0 b|] * [|b|]))
- with ([!n|q!] *[|b|] * 2^[|w_head0 b|]);
- try (ring;fail).
- rewrite Z_div_plus_l;auto with zarith.
- assert (H10 := spec_to_Z
- (w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r));split;
- auto with zarith.
- rewrite H9.
- apply Zdiv_lt_upper_bound;auto with zarith.
- rewrite Z.mul_comm;auto with zarith.
- exact (spec_double_to_Z w_digits w_to_Z spec_to_Z n a).
- Qed.
-
-
- Definition double_modn1 (n:nat) (a:word w n) (b:w) :=
- let p := w_head0 b in
- match w_compare p w_0 with
- | Gt =>
- let b2p := w_add_mul_div p b w_0 in
- let ha := high n a in
- let k := w_sub w_zdigits p in
- let lsr_n := w_add_mul_div k w_0 in
- let r0 := w_add_mul_div p w_0 ha in
- let r := double_modn1_p b2p p n r0 a (double_0 w_0 n) in
- lsr_n r
- | _ => double_modn1_0 b n w_0 a
- end.
-
- Lemma spec_double_modn1_aux : forall n a b,
- double_modn1 n a b = snd (double_divn1 n a b).
- Proof.
- intros n a b;unfold double_divn1,double_modn1.
- rewrite spec_compare; case Z.compare_spec;
- rewrite spec_0; intros H2; auto with zarith.
- apply spec_double_modn1_0.
- apply spec_double_modn1_0.
- rewrite spec_double_modn1_p.
- destruct (double_divn1_p (w_add_mul_div (w_head0 b) b w_0) (w_head0 b) n
- (w_add_mul_div (w_head0 b) w_0 (high n a)) a (double_0 w_0 n));simpl;trivial.
- Qed.
-
- Lemma spec_double_modn1 : forall n a b, 0 < [|b|] ->
- [|double_modn1 n a b|] = [!n|a!] mod [|b|].
- Proof.
- intros n a b H;assert (H1 := spec_double_divn1 n a H).
- assert (H2 := spec_double_modn1_aux n a b).
- rewrite H2;destruct (double_divn1 n a b) as (q,r).
- simpl;apply Zmod_unique with (double_to_Z w_digits w_to_Z n q);auto with zarith.
- destruct H1 as (h1,h2);rewrite h1;ring.
- Qed.
-
-End GENDIVN1.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
deleted file mode 100644
index f65b47c8c..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
+++ /dev/null
@@ -1,475 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-
-Local Open Scope Z_scope.
-
-Section DoubleLift.
- Variable w : Type.
- Variable w_0 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_W0 : w -> zn2z w.
- Variable w_0W : w -> zn2z w.
- Variable w_compare : w -> w -> comparison.
- Variable ww_compare : zn2z w -> zn2z w -> comparison.
- Variable w_head0 : w -> w.
- Variable w_tail0 : w -> w.
- Variable w_add: w -> w -> zn2z w.
- Variable w_add_mul_div : w -> w -> w -> w.
- Variable ww_sub: zn2z w -> zn2z w -> zn2z w.
- Variable w_digits : positive.
- Variable ww_Digits : positive.
- Variable w_zdigits : w.
- Variable ww_zdigits : zn2z w.
- Variable low: zn2z w -> w.
-
- Definition ww_head0 x :=
- match x with
- | W0 => ww_zdigits
- | WW xh xl =>
- match w_compare w_0 xh with
- | Eq => w_add w_zdigits (w_head0 xl)
- | _ => w_0W (w_head0 xh)
- end
- end.
-
-
- Definition ww_tail0 x :=
- match x with
- | W0 => ww_zdigits
- | WW xh xl =>
- match w_compare w_0 xl with
- | Eq => w_add w_zdigits (w_tail0 xh)
- | _ => w_0W (w_tail0 xl)
- end
- end.
-
-
- (* 0 < p < ww_digits *)
- Definition ww_add_mul_div p x y :=
- let zdigits := w_0W w_zdigits in
- match x, y with
- | W0, W0 => W0
- | W0, WW yh yl =>
- match ww_compare p zdigits with
- | Eq => w_0W yh
- | Lt => w_0W (w_add_mul_div (low p) w_0 yh)
- | Gt =>
- let n := low (ww_sub p zdigits) in
- w_WW (w_add_mul_div n w_0 yh) (w_add_mul_div n yh yl)
- end
- | WW xh xl, W0 =>
- match ww_compare p zdigits with
- | Eq => w_W0 xl
- | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl w_0)
- | Gt =>
- let n := low (ww_sub p zdigits) in
- w_W0 (w_add_mul_div n xl w_0)
- end
- | WW xh xl, WW yh yl =>
- match ww_compare p zdigits with
- | Eq => w_WW xl yh
- | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl yh)
- | Gt =>
- let n := low (ww_sub p zdigits) in
- w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl)
- end
- end.
-
- Section DoubleProof.
- Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB.
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_compare : forall x y,
- w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_ww_compare : forall x y,
- ww_compare x y = Z.compare [[x]] [[y]].
- Variable spec_ww_digits : ww_Digits = xO w_digits.
- Variable spec_w_head00 : forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits.
- Variable spec_w_head0 : forall x, 0 < [|x|] ->
- wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB.
- Variable spec_w_tail00 : forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits.
- Variable spec_w_tail0 : forall x, 0 < [|x|] ->
- exists y, 0 <= y /\ [|x|] = (2* y + 1) * (2 ^ [|w_tail0 x|]).
- Variable spec_w_add_mul_div : forall x y p,
- [|p|] <= Zpos w_digits ->
- [| w_add_mul_div p x y |] =
- ([|x|] * (2 ^ [|p|]) +
- [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB.
- Variable spec_w_add: forall x y,
- [[w_add x y]] = [|x|] + [|y|].
- Variable spec_ww_sub: forall x y,
- [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
-
- Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits.
- Variable spec_low: forall x, [| low x|] = [[x]] mod wB.
-
- Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos ww_Digits.
-
- Hint Resolve div_le_0 div_lt w_to_Z_wwB: lift.
- Ltac zarith := auto with zarith lift.
-
- Lemma spec_ww_head00 : forall x, [[x]] = 0 -> [[ww_head0 x]] = Zpos ww_Digits.
- Proof.
- intros x; case x; unfold ww_head0.
- intros HH; rewrite spec_ww_zdigits; auto.
- intros xh xl; simpl; intros Hx.
- case (spec_to_Z xh); intros Hx1 Hx2.
- case (spec_to_Z xl); intros Hy1 Hy2.
- assert (F1: [|xh|] = 0).
- { Z.le_elim Hy1; auto.
- - absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- apply Z.lt_le_trans with (1 := Hy1); auto with zarith.
- pattern [|xl|] at 1; rewrite <- (Z.add_0_l [|xl|]).
- apply Z.add_le_mono_r; auto with zarith.
- - Z.le_elim Hx1; auto.
- absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- rewrite <- Hy1; rewrite Z.add_0_r; auto with zarith.
- apply Z.mul_pos_pos; auto with zarith. }
- rewrite spec_compare. case Z.compare_spec.
- intros H; simpl.
- rewrite spec_w_add; rewrite spec_w_head00.
- rewrite spec_zdigits; rewrite spec_ww_digits.
- rewrite Pos2Z.inj_xO; auto with zarith.
- rewrite F1 in Hx; auto with zarith.
- rewrite spec_w_0; auto with zarith.
- rewrite spec_w_0; auto with zarith.
- Qed.
-
- Lemma spec_ww_head0 : forall x, 0 < [[x]] ->
- wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB.
- Proof.
- clear spec_ww_zdigits.
- rewrite wwB_div_2;rewrite Z.mul_comm;rewrite wwB_wBwB.
- assert (U:= lt_0_wB w_digits); destruct x as [ |xh xl];simpl ww_to_Z;intros H.
- unfold Z.lt in H;discriminate H.
- rewrite spec_compare, spec_w_0. case Z.compare_spec; intros H0.
- rewrite <- H0 in *. simpl Z.add. simpl in H.
- case (spec_to_Z w_zdigits);
- case (spec_to_Z (w_head0 xl)); intros HH1 HH2 HH3 HH4.
- rewrite spec_w_add.
- rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith.
- case (spec_w_head0 H); intros H1 H2.
- rewrite Z.pow_2_r; fold wB; rewrite <- Z.mul_assoc; split.
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- apply Z.mul_lt_mono_pos_l; auto with zarith.
- assert (H1 := spec_w_head0 H0).
- rewrite spec_w_0W.
- split.
- rewrite Z.mul_add_distr_l;rewrite Z.mul_assoc.
- apply Z.le_trans with (2 ^ [|w_head0 xh|] * [|xh|] * wB).
- rewrite Z.mul_comm; zarith.
- assert (0 <= 2 ^ [|w_head0 xh|] * [|xl|]);zarith.
- assert (H2:=spec_to_Z xl);apply Z.mul_nonneg_nonneg;zarith.
- case (spec_to_Z (w_head0 xh)); intros H2 _.
- generalize ([|w_head0 xh|]) H1 H2;clear H1 H2;
- intros p H1 H2.
- assert (Eq1 : 2^p < wB).
- rewrite <- (Z.mul_1_r (2^p));apply Z.le_lt_trans with (2^p*[|xh|]);zarith.
- assert (Eq2: p < Zpos w_digits).
- destruct (Z.le_gt_cases (Zpos w_digits) p);trivial;contradict Eq1.
- apply Z.le_ngt;unfold base;apply Zpower_le_monotone;zarith.
- assert (Zpos w_digits = p + (Zpos w_digits - p)). ring.
- rewrite Z.pow_2_r.
- unfold base at 2;rewrite H3;rewrite Zpower_exp;zarith.
- rewrite <- Z.mul_assoc; apply Z.mul_lt_mono_pos_l; zarith.
- rewrite <- (Z.add_0_r (2^(Zpos w_digits - p)*wB));apply beta_lex_inv;zarith.
- apply Z.mul_lt_mono_pos_r with (2 ^ p); zarith.
- rewrite <- Zpower_exp;zarith.
- rewrite Z.mul_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith.
- assert (H1 := spec_to_Z xh);zarith.
- Qed.
-
- Lemma spec_ww_tail00 : forall x, [[x]] = 0 -> [[ww_tail0 x]] = Zpos ww_Digits.
- Proof.
- intros x; case x; unfold ww_tail0.
- intros HH; rewrite spec_ww_zdigits; auto.
- intros xh xl; simpl; intros Hx.
- case (spec_to_Z xh); intros Hx1 Hx2.
- case (spec_to_Z xl); intros Hy1 Hy2.
- assert (F1: [|xh|] = 0).
- { Z.le_elim Hy1; auto.
- - absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- apply Z.lt_le_trans with (1 := Hy1); auto with zarith.
- pattern [|xl|] at 1; rewrite <- (Z.add_0_l [|xl|]).
- apply Z.add_le_mono_r; auto with zarith.
- - Z.le_elim Hx1; auto.
- absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- rewrite <- Hy1; rewrite Z.add_0_r; auto with zarith.
- apply Z.mul_pos_pos; auto with zarith. }
- assert (F2: [|xl|] = 0).
- rewrite F1 in Hx; auto with zarith.
- rewrite spec_compare; case Z.compare_spec.
- intros H; simpl.
- rewrite spec_w_add; rewrite spec_w_tail00; auto.
- rewrite spec_zdigits; rewrite spec_ww_digits.
- rewrite Pos2Z.inj_xO; auto with zarith.
- rewrite spec_w_0; auto with zarith.
- rewrite spec_w_0; auto with zarith.
- Qed.
-
- Lemma spec_ww_tail0 : forall x, 0 < [[x]] ->
- exists y, 0 <= y /\ [[x]] = (2 * y + 1) * 2 ^ [[ww_tail0 x]].
- Proof.
- clear spec_ww_zdigits.
- destruct x as [ |xh xl];simpl ww_to_Z;intros H.
- unfold Z.lt in H;discriminate H.
- rewrite spec_compare, spec_w_0. case Z.compare_spec; intros H0.
- rewrite <- H0; rewrite Z.add_0_r.
- case (spec_to_Z (w_tail0 xh)); intros HH1 HH2.
- generalize H; rewrite <- H0; rewrite Z.add_0_r; clear H; intros H.
- case (@spec_w_tail0 xh).
- apply Z.mul_lt_mono_pos_r with wB; auto with zarith.
- unfold base; auto with zarith.
- intros z (Hz1, Hz2); exists z; split; auto.
- rewrite spec_w_add; rewrite (fun x => Z.add_comm [|x|]).
- rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith.
- rewrite Z.mul_assoc; rewrite <- Hz2; auto.
-
- case (spec_to_Z (w_tail0 xh)); intros HH1 HH2.
- case (spec_w_tail0 H0); intros z (Hz1, Hz2).
- assert (Hp: [|w_tail0 xl|] < Zpos w_digits).
- case (Z.le_gt_cases (Zpos w_digits) [|w_tail0 xl|]); auto; intros H1.
- absurd (2 ^ (Zpos w_digits) <= 2 ^ [|w_tail0 xl|]).
- apply Z.lt_nge.
- case (spec_to_Z xl); intros HH3 HH4.
- apply Z.le_lt_trans with (2 := HH4).
- apply Z.le_trans with (1 * 2 ^ [|w_tail0 xl|]); auto with zarith.
- rewrite Hz2.
- apply Z.mul_le_mono_nonneg_r; auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- exists ([|xh|] * (2 ^ ((Zpos w_digits - [|w_tail0 xl|]) - 1)) + z); split.
- apply Z.add_nonneg_nonneg; auto.
- apply Z.mul_nonneg_nonneg; auto with zarith.
- case (spec_to_Z xh); auto.
- rewrite spec_w_0W.
- rewrite (Z.mul_add_distr_l 2); rewrite <- Z.add_assoc.
- rewrite Z.mul_add_distr_r; rewrite <- Hz2.
- apply f_equal2 with (f := Z.add); auto.
- rewrite (Z.mul_comm 2).
- repeat rewrite <- Z.mul_assoc.
- apply f_equal2 with (f := Z.mul); auto.
- case (spec_to_Z (w_tail0 xl)); intros HH3 HH4.
- pattern 2 at 2; rewrite <- Z.pow_1_r.
- lazy beta; repeat rewrite <- Zpower_exp; auto with zarith.
- unfold base; apply f_equal with (f := Z.pow 2); auto with zarith.
-
- contradict H0; case (spec_to_Z xl); auto with zarith.
- Qed.
-
- Hint Rewrite Zdiv_0_l Z.mul_0_l Z.add_0_l Z.mul_0_r Z.add_0_r
- spec_w_W0 spec_w_0W spec_w_WW spec_w_0
- (wB_div w_digits w_to_Z spec_to_Z)
- (wB_div_plus w_digits w_to_Z spec_to_Z) : w_rewrite.
- Ltac w_rewrite := autorewrite with w_rewrite;trivial.
-
- Lemma spec_ww_add_mul_div_aux : forall xh xl yh yl p,
- let zdigits := w_0W w_zdigits in
- [[p]] <= Zpos (xO w_digits) ->
- [[match ww_compare p zdigits with
- | Eq => w_WW xl yh
- | Lt => w_WW (w_add_mul_div (low p) xh xl)
- (w_add_mul_div (low p) xl yh)
- | Gt =>
- let n := low (ww_sub p zdigits) in
- w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl)
- end]] =
- ([[WW xh xl]] * (2^[[p]]) +
- [[WW yh yl]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB.
- Proof.
- clear spec_ww_zdigits.
- intros xh xl yh yl p zdigits;assert (HwwB := wwB_pos w_digits).
- case (spec_to_w_Z p); intros Hv1 Hv2.
- replace (Zpos (xO w_digits)) with (Zpos w_digits + Zpos w_digits).
- 2 : rewrite Pos2Z.inj_xO;ring.
- replace (Zpos w_digits + Zpos w_digits - [[p]]) with
- (Zpos w_digits + (Zpos w_digits - [[p]])). 2:ring.
- intros Hp; assert (Hxh := spec_to_Z xh);assert (Hxl:=spec_to_Z xl);
- assert (Hx := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl));
- simpl in Hx;assert (Hyh := spec_to_Z yh);assert (Hyl:=spec_to_Z yl);
- assert (Hy:=spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW yh yl));simpl in Hy.
- rewrite spec_ww_compare; case Z.compare_spec; intros H1.
- rewrite H1; unfold zdigits; rewrite spec_w_0W.
- rewrite spec_zdigits; rewrite Z.sub_diag; rewrite Z.add_0_r.
- simpl ww_to_Z; w_rewrite;zarith.
- fold wB.
- rewrite Z.mul_add_distr_r;rewrite <- Z.mul_assoc;rewrite <- Z.add_assoc.
- rewrite <- Z.pow_2_r.
- rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|].
- exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xl yh)). ring.
- simpl ww_to_Z; w_rewrite;zarith.
- assert (HH0: [|low p|] = [[p]]).
- rewrite spec_low.
- apply Zmod_small.
- case (spec_to_w_Z p); intros HH1 HH2; split; auto.
- generalize H1; unfold zdigits; rewrite spec_w_0W;
- rewrite spec_zdigits; intros tmp.
- apply Z.lt_le_trans with (1 := tmp).
- unfold base.
- apply Zpower2_le_lin; auto with zarith.
- 2: generalize H1; unfold zdigits; rewrite spec_w_0W;
- rewrite spec_zdigits; auto with zarith.
- generalize H1; unfold zdigits; rewrite spec_w_0W;
- rewrite spec_zdigits; auto; clear H1; intros H1.
- assert (HH: [|low p|] <= Zpos w_digits).
- rewrite HH0; auto with zarith.
- repeat rewrite spec_w_add_mul_div with (1 := HH).
- rewrite HH0.
- rewrite Z.mul_add_distr_r.
- pattern ([|xl|] * 2 ^ [[p]]) at 2;
- rewrite shift_unshift_mod with (n:= Zpos w_digits);fold wB;zarith.
- replace ([|xh|] * wB * 2^[[p]]) with ([|xh|] * 2^[[p]] * wB). 2:ring.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r. rewrite <- Z.add_assoc.
- unfold base at 5;rewrite <- Zmod_shift_r;zarith.
- unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits));
- fold wB;fold wwB;zarith.
- rewrite wwB_wBwB;rewrite Z.pow_2_r; rewrite Zmult_mod_distr_r;zarith.
- unfold ww_digits;rewrite Pos2Z.inj_xO;zarith. apply Z_mod_lt;zarith.
- split;zarith. apply Zdiv_lt_upper_bound;zarith.
- rewrite <- Zpower_exp;zarith.
- ring_simplify ([[p]] + (Zpos w_digits - [[p]]));fold wB;zarith.
- assert (Hv: [[p]] > Zpos w_digits).
- generalize H1; clear H1.
- unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits; auto with zarith.
- clear H1.
- assert (HH0: [|low (ww_sub p zdigits)|] = [[p]] - Zpos w_digits).
- rewrite spec_low.
- rewrite spec_ww_sub.
- unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits.
- rewrite <- Zmod_div_mod; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_lt_lin; auto with zarith.
- exists wB; unfold base.
- unfold ww_digits; rewrite (Pos2Z.inj_xO w_digits).
- rewrite <- Zpower_exp; auto with zarith.
- apply f_equal with (f := fun x => 2 ^ x); auto with zarith.
- assert (HH: [|low (ww_sub p zdigits)|] <= Zpos w_digits).
- rewrite HH0; auto with zarith.
- replace (Zpos w_digits + (Zpos w_digits - [[p]])) with
- (Zpos w_digits - ([[p]] - Zpos w_digits)); zarith.
- lazy zeta; simpl ww_to_Z; w_rewrite;zarith.
- repeat rewrite spec_w_add_mul_div;zarith.
- rewrite HH0.
- pattern wB at 5;replace wB with
- (2^(([[p]] - Zpos w_digits)
- + (Zpos w_digits - ([[p]] - Zpos w_digits)))).
- rewrite Zpower_exp;zarith. rewrite Z.mul_assoc.
- rewrite Z_div_plus_l;zarith.
- rewrite shift_unshift_mod with (a:= [|yh|]) (p:= [[p]] - Zpos w_digits)
- (n := Zpos w_digits);zarith. fold wB.
- set (u := [[p]] - Zpos w_digits).
- replace [[p]] with (u + Zpos w_digits);zarith.
- rewrite Zpower_exp;zarith. rewrite Z.mul_assoc. fold wB.
- repeat rewrite Z.add_assoc. rewrite <- Z.mul_add_distr_r.
- repeat rewrite <- Z.add_assoc.
- unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits));
- fold wB;fold wwB;zarith.
- unfold base;rewrite Zmod_shift_r with (a:= Zpos w_digits)
- (b:= Zpos w_digits);fold wB;fold wwB;zarith.
- rewrite wwB_wBwB; rewrite Z.pow_2_r; rewrite Zmult_mod_distr_r;zarith.
- rewrite Z.mul_add_distr_r.
- replace ([|xh|] * wB * 2 ^ u) with
- ([|xh|]*2^u*wB). 2:ring.
- repeat rewrite <- Z.add_assoc.
- rewrite (Z.add_comm ([|xh|] * 2 ^ u * wB)).
- rewrite Z_mod_plus;zarith. rewrite Z_mod_mult;zarith.
- unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith.
- unfold u; split;zarith.
- split;zarith. unfold u; apply Zdiv_lt_upper_bound;zarith.
- rewrite <- Zpower_exp;zarith.
- fold u.
- ring_simplify (u + (Zpos w_digits - u)); fold
- wB;zarith. unfold ww_digits;rewrite Pos2Z.inj_xO;zarith.
- unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith.
- unfold u; split;zarith.
- unfold u; split;zarith.
- apply Zdiv_lt_upper_bound;zarith.
- rewrite <- Zpower_exp;zarith.
- fold u.
- ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith.
- unfold u;zarith.
- unfold u;zarith.
- set (u := [[p]] - Zpos w_digits).
- ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith.
- Qed.
-
- Lemma spec_ww_add_mul_div : forall x y p,
- [[p]] <= Zpos (xO w_digits) ->
- [[ ww_add_mul_div p x y ]] =
- ([[x]] * (2^[[p]]) +
- [[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB.
- Proof.
- clear spec_ww_zdigits.
- intros x y p H.
- destruct x as [ |xh xl];
- [assert (H1 := @spec_ww_add_mul_div_aux w_0 w_0)
- |assert (H1 := @spec_ww_add_mul_div_aux xh xl)];
- (destruct y as [ |yh yl];
- [generalize (H1 w_0 w_0 p H) | generalize (H1 yh yl p H)];
- clear H1;w_rewrite);simpl ww_add_mul_div.
- replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial].
- intros Heq;rewrite <- Heq;clear Heq; auto.
- rewrite spec_ww_compare. case Z.compare_spec; intros H1; w_rewrite.
- rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith.
- generalize H1; w_rewrite; rewrite spec_zdigits; clear H1; intros H1.
- assert (HH0: [|low p|] = [[p]]).
- rewrite spec_low.
- apply Zmod_small.
- case (spec_to_w_Z p); intros HH1 HH2; split; auto.
- apply Z.lt_le_trans with (1 := H1).
- unfold base; apply Zpower2_le_lin; auto with zarith.
- rewrite HH0; auto with zarith.
- replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial].
- intros Heq;rewrite <- Heq;clear Heq.
- generalize (spec_ww_compare p (w_0W w_zdigits));
- case ww_compare; intros H1; w_rewrite.
- rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith.
- rewrite Pos2Z.inj_xO in H;zarith.
- assert (HH: [|low (ww_sub p (w_0W w_zdigits)) |] = [[p]] - Zpos w_digits).
- symmetry in H1; change ([[p]] > [[w_0W w_zdigits]]) in H1.
- revert H1.
- rewrite spec_low.
- rewrite spec_ww_sub; w_rewrite; intros H1.
- rewrite <- Zmod_div_mod; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_lt_lin; auto with zarith.
- unfold base; auto with zarith.
- unfold base; auto with zarith.
- exists wB; unfold base.
- unfold ww_digits; rewrite (Pos2Z.inj_xO w_digits).
- rewrite <- Zpower_exp; auto with zarith.
- apply f_equal with (f := fun x => 2 ^ x); auto with zarith.
- case (spec_to_Z xh); auto with zarith.
- Qed.
-
- End DoubleProof.
-
-End DoubleLift.
-
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
deleted file mode 100644
index b99013900..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
+++ /dev/null
@@ -1,621 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-
-Local Open Scope Z_scope.
-
-Section DoubleMul.
- Variable w : Type.
- Variable w_0 : w.
- Variable w_1 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_W0 : w -> zn2z w.
- Variable w_0W : w -> zn2z w.
- Variable w_compare : w -> w -> comparison.
- Variable w_succ : w -> w.
- Variable w_add_c : w -> w -> carry w.
- Variable w_add : w -> w -> w.
- Variable w_sub: w -> w -> w.
- Variable w_mul_c : w -> w -> zn2z w.
- Variable w_mul : w -> w -> w.
- Variable w_square_c : w -> zn2z w.
- Variable ww_add_c : zn2z w -> zn2z w -> carry (zn2z w).
- Variable ww_add : zn2z w -> zn2z w -> zn2z w.
- Variable ww_add_carry : zn2z w -> zn2z w -> zn2z w.
- Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w).
- Variable ww_sub : zn2z w -> zn2z w -> zn2z w.
-
- (* ** Multiplication ** *)
-
- (* (xh*B+xl) (yh*B + yl)
- xh*yh = hh = |hhh|hhl|B2
- xh*yl +xl*yh = cc = |cch|ccl|B
- xl*yl = ll = |llh|lll
- *)
-
- Definition double_mul_c (cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w) x y :=
- match x, y with
- | W0, _ => W0
- | _, W0 => W0
- | WW xh xl, WW yh yl =>
- let hh := w_mul_c xh yh in
- let ll := w_mul_c xl yl in
- let (wc,cc) := cross xh xl yh yl hh ll in
- match cc with
- | W0 => WW (ww_add hh (w_W0 wc)) ll
- | WW cch ccl =>
- match ww_add_c (w_W0 ccl) ll with
- | C0 l => WW (ww_add hh (w_WW wc cch)) l
- | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l
- end
- end
- end.
-
- Definition ww_mul_c :=
- double_mul_c
- (fun xh xl yh yl hh ll=>
- match ww_add_c (w_mul_c xh yl) (w_mul_c xl yh) with
- | C0 cc => (w_0, cc)
- | C1 cc => (w_1, cc)
- end).
-
- Definition w_2 := w_add w_1 w_1.
-
- Definition kara_prod xh xl yh yl hh ll :=
- match ww_add_c hh ll with
- C0 m =>
- match w_compare xl xh with
- Eq => (w_0, m)
- | Lt =>
- match w_compare yl yh with
- Eq => (w_0, m)
- | Lt => (w_0, ww_sub m (w_mul_c (w_sub xh xl) (w_sub yh yl)))
- | Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with
- C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1)
- end
- end
- | Gt =>
- match w_compare yl yh with
- Eq => (w_0, m)
- | Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with
- C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1)
- end
- | Gt => (w_0, ww_sub m (w_mul_c (w_sub xl xh) (w_sub yl yh)))
- end
- end
- | C1 m =>
- match w_compare xl xh with
- Eq => (w_1, m)
- | Lt =>
- match w_compare yl yh with
- Eq => (w_1, m)
- | Lt => match ww_sub_c m (w_mul_c (w_sub xh xl) (w_sub yh yl)) with
- C0 m1 => (w_1, m1) | C1 m1 => (w_0, m1)
- end
- | Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with
- C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1)
- end
- end
- | Gt =>
- match w_compare yl yh with
- Eq => (w_1, m)
- | Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with
- C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1)
- end
- | Gt => match ww_sub_c m (w_mul_c (w_sub xl xh) (w_sub yl yh)) with
- C1 m1 => (w_0, m1) | C0 m1 => (w_1, m1)
- end
- end
- end
- end.
-
- Definition ww_karatsuba_c := double_mul_c kara_prod.
-
- Definition ww_mul x y :=
- match x, y with
- | W0, _ => W0
- | _, W0 => W0
- | WW xh xl, WW yh yl =>
- let ccl := w_add (w_mul xh yl) (w_mul xl yh) in
- ww_add (w_W0 ccl) (w_mul_c xl yl)
- end.
-
- Definition ww_square_c x :=
- match x with
- | W0 => W0
- | WW xh xl =>
- let hh := w_square_c xh in
- let ll := w_square_c xl in
- let xhxl := w_mul_c xh xl in
- let (wc,cc) :=
- match ww_add_c xhxl xhxl with
- | C0 cc => (w_0, cc)
- | C1 cc => (w_1, cc)
- end in
- match cc with
- | W0 => WW (ww_add hh (w_W0 wc)) ll
- | WW cch ccl =>
- match ww_add_c (w_W0 ccl) ll with
- | C0 l => WW (ww_add hh (w_WW wc cch)) l
- | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l
- end
- end
- end.
-
- Section DoubleMulAddn1.
- Variable w_mul_add : w -> w -> w -> w * w.
-
- Fixpoint double_mul_add_n1 (n:nat) : word w n -> w -> w -> w * word w n :=
- match n return word w n -> w -> w -> w * word w n with
- | O => w_mul_add
- | S n1 =>
- let mul_add := double_mul_add_n1 n1 in
- fun x y r =>
- match x with
- | W0 => (w_0,extend w_0W n1 r)
- | WW xh xl =>
- let (rl,l) := mul_add xl y r in
- let (rh,h) := mul_add xh y rl in
- (rh, double_WW w_WW n1 h l)
- end
- end.
-
- End DoubleMulAddn1.
-
- Section DoubleMulAddmn1.
- Variable wn: Type.
- Variable extend_n : w -> wn.
- Variable wn_0W : wn -> zn2z wn.
- Variable wn_WW : wn -> wn -> zn2z wn.
- Variable w_mul_add_n1 : wn -> w -> w -> w*wn.
- Fixpoint double_mul_add_mn1 (m:nat) :
- word wn m -> w -> w -> w*word wn m :=
- match m return word wn m -> w -> w -> w*word wn m with
- | O => w_mul_add_n1
- | S m1 =>
- let mul_add := double_mul_add_mn1 m1 in
- fun x y r =>
- match x with
- | W0 => (w_0,extend wn_0W m1 (extend_n r))
- | WW xh xl =>
- let (rl,l) := mul_add xl y r in
- let (rh,h) := mul_add xh y rl in
- (rh, double_WW wn_WW m1 h l)
- end
- end.
-
- End DoubleMulAddmn1.
-
- Definition w_mul_add x y r :=
- match w_mul_c x y with
- | W0 => (w_0, r)
- | WW h l =>
- match w_add_c l r with
- | C0 lr => (h,lr)
- | C1 lr => (w_succ h, lr)
- end
- end.
-
-
- (*Section DoubleProof. *)
- Variable w_digits : positive.
- Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
- Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
-
- Notation "[|| x ||]" :=
- (zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99).
-
- Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x)
- (at level 0, x at level 99).
-
- Variable spec_more_than_1_digit: 1 < Zpos w_digits.
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_1 : [|w_1|] = 1.
-
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
-
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_w_compare :
- forall x y, w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB.
- Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
- Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB.
- Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
-
- Variable spec_w_mul_c : forall x y, [[ w_mul_c x y ]] = [|x|] * [|y|].
- Variable spec_w_mul : forall x y, [|w_mul x y|] = ([|x|] * [|y|]) mod wB.
- Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|].
-
- Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
- Variable spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB.
- Variable spec_ww_add_carry :
- forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB.
- Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
- Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]].
-
-
- Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB.
- Proof. intros x;apply spec_ww_to_Z;auto. Qed.
-
- Lemma spec_ww_to_Z_wBwB : forall x, 0 <= [[x]] < wB^2.
- Proof. rewrite <- wwB_wBwB;apply spec_ww_to_Z. Qed.
-
- Hint Resolve spec_ww_to_Z spec_ww_to_Z_wBwB : mult.
- Ltac zarith := auto with zarith mult.
-
- Lemma wBwB_lex: forall a b c d,
- a * wB^2 + [[b]] <= c * wB^2 + [[d]] ->
- a <= c.
- Proof.
- intros a b c d H; apply beta_lex with [[b]] [[d]] (wB^2);zarith.
- Qed.
-
- Lemma wBwB_lex_inv: forall a b c d,
- a < c ->
- a * wB^2 + [[b]] < c * wB^2 + [[d]].
- Proof.
- intros a b c d H; apply beta_lex_inv; zarith.
- Qed.
-
- Lemma sum_mul_carry : forall xh xl yh yl wc cc,
- [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] ->
- 0 <= [|wc|] <= 1.
- Proof.
- intros.
- apply (sum_mul_carry [|xh|] [|xl|] [|yh|] [|yl|] [|wc|][[cc]] wB);zarith.
- apply wB_pos.
- Qed.
-
- Theorem mult_add_ineq: forall xH yH crossH,
- 0 <= [|xH|] * [|yH|] + [|crossH|] < wwB.
- Proof.
- intros;rewrite wwB_wBwB;apply mult_add_ineq;zarith.
- Qed.
-
- Hint Resolve mult_add_ineq : mult.
-
- Lemma spec_mul_aux : forall xh xl yh yl wc (cc:zn2z w) hh ll,
- [[hh]] = [|xh|] * [|yh|] ->
- [[ll]] = [|xl|] * [|yl|] ->
- [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] ->
- [||match cc with
- | W0 => WW (ww_add hh (w_W0 wc)) ll
- | WW cch ccl =>
- match ww_add_c (w_W0 ccl) ll with
- | C0 l => WW (ww_add hh (w_WW wc cch)) l
- | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l
- end
- end||] = ([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|]).
- Proof.
- intros;assert (U1 := wB_pos w_digits).
- replace (([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|])) with
- ([|xh|]*[|yh|]*wB^2 + ([|xh|]*[|yl|] + [|xl|]*[|yh|])*wB + [|xl|]*[|yl|]).
- 2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0.
- assert (H2 := sum_mul_carry _ _ _ _ _ _ H1).
- destruct cc as [ | cch ccl]; simpl zn2z_to_Z; simpl ww_to_Z.
- rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_small;
- rewrite wwB_wBwB. ring.
- rewrite <- (Z.add_0_r ([|wc|]*wB));rewrite H;apply mult_add_ineq3;zarith.
- simpl ww_to_Z in H1. assert (U:=spec_to_Z cch).
- assert ([|wc|]*wB + [|cch|] <= 2*wB - 3).
- destruct (Z_le_gt_dec ([|wc|]*wB + [|cch|]) (2*wB - 3)) as [Hle|Hgt];trivial.
- assert ([|xh|] * [|yl|] + [|xl|] * [|yh|] <= (2*wB - 4)*wB + 2).
- ring_simplify ((2*wB - 4)*wB + 2).
- assert (H4 := Zmult_lt_b _ _ _ (spec_to_Z xh) (spec_to_Z yl)).
- assert (H5 := Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)).
- omega.
- generalize H3;clear H3;rewrite <- H1.
- rewrite Z.add_assoc; rewrite Z.pow_2_r; rewrite Z.mul_assoc;
- rewrite <- Z.mul_add_distr_r.
- assert (((2 * wB - 4) + 2)*wB <= ([|wc|] * wB + [|cch|])*wB).
- apply Z.mul_le_mono_nonneg;zarith.
- rewrite Z.mul_add_distr_r in H3.
- intros. assert (U2 := spec_to_Z ccl);omega.
- generalize (spec_ww_add_c (w_W0 ccl) ll);destruct (ww_add_c (w_W0 ccl) ll)
- as [l|l];unfold interp_carry;rewrite spec_w_W0;try rewrite Z.mul_1_l;
- simpl zn2z_to_Z;
- try rewrite spec_ww_add;try rewrite spec_ww_add_carry;rewrite spec_w_WW;
- rewrite Zmod_small;rewrite wwB_wBwB;intros.
- rewrite H4;ring. rewrite H;apply mult_add_ineq2;zarith.
- rewrite Z.add_assoc;rewrite Z.mul_add_distr_r.
- rewrite Z.mul_1_l;rewrite <- Z.add_assoc;rewrite H4;ring.
- repeat rewrite <- Z.add_assoc;rewrite H;apply mult_add_ineq2;zarith.
- Qed.
-
- Lemma spec_double_mul_c : forall cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w,
- (forall xh xl yh yl hh ll,
- [[hh]] = [|xh|]*[|yh|] ->
- [[ll]] = [|xl|]*[|yl|] ->
- let (wc,cc) := cross xh xl yh yl hh ll in
- [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]) ->
- forall x y, [||double_mul_c cross x y||] = [[x]] * [[y]].
- Proof.
- intros cross Hcross x y;destruct x as [ |xh xl];simpl;trivial.
- destruct y as [ |yh yl];simpl. rewrite Z.mul_0_r;trivial.
- assert (H1:= spec_w_mul_c xh yh);assert (H2:= spec_w_mul_c xl yl).
- generalize (Hcross _ _ _ _ _ _ H1 H2).
- destruct (cross xh xl yh yl (w_mul_c xh yh) (w_mul_c xl yl)) as (wc,cc).
- intros;apply spec_mul_aux;trivial.
- rewrite <- wwB_wBwB;trivial.
- Qed.
-
- Lemma spec_ww_mul_c : forall x y, [||ww_mul_c x y||] = [[x]] * [[y]].
- Proof.
- intros x y;unfold ww_mul_c;apply spec_double_mul_c.
- intros xh xl yh yl hh ll H1 H2.
- generalize (spec_ww_add_c (w_mul_c xh yl) (w_mul_c xl yh));
- destruct (ww_add_c (w_mul_c xh yl) (w_mul_c xl yh)) as [c|c];
- unfold interp_carry;repeat rewrite spec_w_mul_c;intros H;
- (rewrite spec_w_0 || rewrite spec_w_1);rewrite H;ring.
- Qed.
-
- Lemma spec_w_2: [|w_2|] = 2.
- unfold w_2; rewrite spec_w_add; rewrite spec_w_1; simpl.
- apply Zmod_small; split; auto with zarith.
- rewrite <- (Z.pow_1_r 2); unfold base; apply Zpower_lt_monotone; auto with zarith.
- Qed.
-
- Lemma kara_prod_aux : forall xh xl yh yl,
- xh*yh + xl*yl - (xh-xl)*(yh-yl) = xh*yl + xl*yh.
- Proof. intros;ring. Qed.
-
- Lemma spec_kara_prod : forall xh xl yh yl hh ll,
- [[hh]] = [|xh|]*[|yh|] ->
- [[ll]] = [|xl|]*[|yl|] ->
- let (wc,cc) := kara_prod xh xl yh yl hh ll in
- [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|].
- Proof.
- intros xh xl yh yl hh ll H H0; rewrite <- kara_prod_aux;
- rewrite <- H; rewrite <- H0; unfold kara_prod.
- assert (Hxh := (spec_to_Z xh)); assert (Hxl := (spec_to_Z xl));
- assert (Hyh := (spec_to_Z yh)); assert (Hyl := (spec_to_Z yl)).
- generalize (spec_ww_add_c hh ll); case (ww_add_c hh ll);
- intros z Hz; rewrite <- Hz; unfold interp_carry; assert (Hz1 := (spec_ww_to_Z z)).
- rewrite spec_w_compare; case Z.compare_spec; intros Hxlh;
- try rewrite Hxlh; try rewrite spec_w_0; try (ring; fail).
- rewrite spec_w_compare; case Z.compare_spec; intros Hylh.
- rewrite Hylh; rewrite spec_w_0; try (ring; fail).
- rewrite spec_w_0; try (ring; fail).
- repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- split; auto with zarith.
- simpl in Hz; rewrite Hz; rewrite H; rewrite H0.
- rewrite kara_prod_aux; apply Z.add_nonneg_nonneg; apply Z.mul_nonneg_nonneg; auto with zarith.
- apply Z.le_lt_trans with ([[z]]-0); auto with zarith.
- unfold Z.sub; apply Z.add_le_mono_l; apply Z.le_0_sub; simpl; rewrite Z.opp_involutive.
- apply Z.mul_nonneg_nonneg; auto with zarith.
- match goal with |- context[ww_add_c ?x ?y] =>
- generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0;
- intros z1 Hz2
- end.
- simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2;
- repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_compare; case Z.compare_spec; intros Hylh.
- rewrite Hylh; rewrite spec_w_0; try (ring; fail).
- match goal with |- context[ww_add_c ?x ?y] =>
- generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0;
- intros z1 Hz2
- end.
- simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2;
- repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_0; try (ring; fail).
- repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- split.
- match goal with |- context[(?x - ?y) * (?z - ?t)] =>
- replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring]
- end.
- simpl in Hz; rewrite Hz; rewrite H; rewrite H0.
- rewrite kara_prod_aux; apply Z.add_nonneg_nonneg; apply Z.mul_nonneg_nonneg; auto with zarith.
- apply Z.le_lt_trans with ([[z]]-0); auto with zarith.
- unfold Z.sub; apply Z.add_le_mono_l; apply Z.le_0_sub; simpl; rewrite Z.opp_involutive.
- apply Z.mul_nonneg_nonneg; auto with zarith.
- (** there is a carry in hh + ll **)
- rewrite Z.mul_1_l.
- rewrite spec_w_compare; case Z.compare_spec; intros Hxlh;
- try rewrite Hxlh; try rewrite spec_w_1; try (ring; fail).
- rewrite spec_w_compare; case Z.compare_spec; intros Hylh;
- try rewrite Hylh; try rewrite spec_w_1; try (ring; fail).
- match goal with |- context[ww_sub_c ?x ?y] =>
- generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1;
- intros z1 Hz2
- end.
- simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_0; rewrite Z.mul_0_l; rewrite Z.add_0_l.
- generalize Hz2; clear Hz2; unfold interp_carry.
- repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- match goal with |- context[ww_add_c ?x ?y] =>
- generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1;
- intros z1 Hz2
- end.
- simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_2; unfold interp_carry in Hz2.
- transitivity (wwB + (1 * wwB + [[z1]])).
- ring.
- rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_compare; case Z.compare_spec; intros Hylh;
- try rewrite Hylh; try rewrite spec_w_1; try (ring; fail).
- match goal with |- context[ww_add_c ?x ?y] =>
- generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1;
- intros z1 Hz2
- end.
- simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_2; unfold interp_carry in Hz2.
- transitivity (wwB + (1 * wwB + [[z1]])).
- ring.
- rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- match goal with |- context[ww_sub_c ?x ?y] =>
- generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1;
- intros z1 Hz2
- end.
- simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_0; rewrite Z.mul_0_l; rewrite Z.add_0_l.
- match goal with |- context[(?x - ?y) * (?z - ?t)] =>
- replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring]
- end.
- generalize Hz2; clear Hz2; unfold interp_carry.
- repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- Qed.
-
- Lemma sub_carry : forall xh xl yh yl z,
- 0 <= z ->
- [|xh|]*[|yl|] + [|xl|]*[|yh|] = wwB + z ->
- z < wwB.
- Proof.
- intros xh xl yh yl z Hle Heq.
- destruct (Z_le_gt_dec wwB z);auto with zarith.
- generalize (Zmult_lt_b _ _ _ (spec_to_Z xh) (spec_to_Z yl)).
- generalize (Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)).
- rewrite <- wwB_wBwB;intros H1 H2.
- assert (H3 := wB_pos w_digits).
- assert (2*wB <= wwB).
- rewrite wwB_wBwB; rewrite Z.pow_2_r; apply Z.mul_le_mono_nonneg;zarith.
- omega.
- Qed.
-
- Ltac Spec_ww_to_Z x :=
- let H:= fresh "H" in
- assert (H:= spec_ww_to_Z x).
-
- Ltac Zmult_lt_b x y :=
- let H := fresh "H" in
- assert (H := Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)).
-
- Lemma spec_ww_karatsuba_c : forall x y, [||ww_karatsuba_c x y||]=[[x]]*[[y]].
- Proof.
- intros x y; unfold ww_karatsuba_c;apply spec_double_mul_c.
- intros; apply spec_kara_prod; auto.
- Qed.
-
- Lemma spec_ww_mul : forall x y, [[ww_mul x y]] = [[x]]*[[y]] mod wwB.
- Proof.
- assert (U:= lt_0_wB w_digits).
- assert (U1:= lt_0_wwB w_digits).
- intros x y; case x; auto; intros xh xl.
- case y; auto.
- simpl; rewrite Z.mul_0_r; rewrite Zmod_small; auto with zarith.
- intros yh yl;simpl.
- repeat (rewrite spec_ww_add || rewrite spec_w_W0 || rewrite spec_w_mul_c
- || rewrite spec_w_add || rewrite spec_w_mul).
- rewrite <- Zplus_mod; auto with zarith.
- repeat (rewrite Z.mul_add_distr_r || rewrite Z.mul_add_distr_l).
- rewrite <- Zmult_mod_distr_r; auto with zarith.
- rewrite <- Z.pow_2_r; rewrite <- wwB_wBwB; auto with zarith.
- rewrite Zplus_mod; auto with zarith.
- rewrite Zmod_mod; auto with zarith.
- rewrite <- Zplus_mod; auto with zarith.
- match goal with |- ?X mod _ = _ =>
- rewrite <- Z_mod_plus with (a := X) (b := [|xh|] * [|yh|])
- end; auto with zarith.
- f_equal; auto; rewrite wwB_wBwB; ring.
- Qed.
-
- Lemma spec_ww_square_c : forall x, [||ww_square_c x||] = [[x]]*[[x]].
- Proof.
- destruct x as [ |xh xl];simpl;trivial.
- case_eq match ww_add_c (w_mul_c xh xl) (w_mul_c xh xl) with
- | C0 cc => (w_0, cc)
- | C1 cc => (w_1, cc)
- end;intros wc cc Heq.
- apply (spec_mul_aux xh xl xh xl wc cc);trivial.
- generalize Heq (spec_ww_add_c (w_mul_c xh xl) (w_mul_c xh xl));clear Heq.
- rewrite spec_w_mul_c;destruct (ww_add_c (w_mul_c xh xl) (w_mul_c xh xl));
- unfold interp_carry;try rewrite Z.mul_1_l;intros Heq Heq';inversion Heq;
- rewrite (Z.mul_comm [|xl|]);subst.
- rewrite spec_w_0;rewrite Z.mul_0_l;rewrite Z.add_0_l;trivial.
- rewrite spec_w_1;rewrite Z.mul_1_l;rewrite <- wwB_wBwB;trivial.
- Qed.
-
- Section DoubleMulAddn1Proof.
-
- Variable w_mul_add : w -> w -> w -> w * w.
- Variable spec_w_mul_add : forall x y r,
- let (h,l):= w_mul_add x y r in
- [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|].
-
- Lemma spec_double_mul_add_n1 : forall n x y r,
- let (h,l) := double_mul_add_n1 w_mul_add n x y r in
- [|h|]*double_wB w_digits n + [!n|l!] = [!n|x!]*[|y|]+[|r|].
- Proof.
- induction n;intros x y r;trivial.
- exact (spec_w_mul_add x y r).
- unfold double_mul_add_n1;destruct x as[ |xh xl];
- fold(double_mul_add_n1 w_mul_add).
- rewrite spec_w_0;rewrite spec_extend;simpl;trivial.
- assert(H:=IHn xl y r);destruct (double_mul_add_n1 w_mul_add n xl y r)as(rl,l).
- assert(U:=IHn xh y rl);destruct(double_mul_add_n1 w_mul_add n xh y rl)as(rh,h).
- rewrite <- double_wB_wwB. rewrite spec_double_WW;simpl;trivial.
- rewrite Z.mul_add_distr_r;rewrite <- Z.add_assoc;rewrite <- H.
- rewrite Z.mul_assoc;rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- rewrite U;ring.
- Qed.
-
- End DoubleMulAddn1Proof.
-
- Lemma spec_w_mul_add : forall x y r,
- let (h,l):= w_mul_add x y r in
- [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|].
- Proof.
- intros x y r;unfold w_mul_add;assert (H:=spec_w_mul_c x y);
- destruct (w_mul_c x y) as [ |h l];simpl;rewrite <- H.
- rewrite spec_w_0;trivial.
- assert (U:=spec_w_add_c l r);destruct (w_add_c l r) as [lr|lr];unfold
- interp_carry in U;try rewrite Z.mul_1_l in H;simpl.
- rewrite U;ring. rewrite spec_w_succ. rewrite Zmod_small.
- rewrite <- Z.add_assoc;rewrite <- U;ring.
- simpl in H;assert (H1:= Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)).
- rewrite <- H in H1.
- assert (H2:=spec_to_Z h);split;zarith.
- case H1;clear H1;intro H1;clear H1.
- replace (wB ^ 2 - 2 * wB) with ((wB - 2)*wB). 2:ring.
- intros H0;assert (U1:= wB_pos w_digits).
- assert (H1 := beta_lex _ _ _ _ _ H0 (spec_to_Z l));zarith.
- Qed.
-
-(* End DoubleProof. *)
-
-End DoubleMul.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
deleted file mode 100644
index d07ce3018..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
+++ /dev/null
@@ -1,1369 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-
-Local Open Scope Z_scope.
-
-Section DoubleSqrt.
- Variable w : Type.
- Variable w_is_even : w -> bool.
- Variable w_compare : w -> w -> comparison.
- Variable w_0 : w.
- Variable w_1 : w.
- Variable w_Bm1 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable w_W0 : w -> zn2z w.
- Variable w_0W : w -> zn2z w.
- Variable w_sub : w -> w -> w.
- Variable w_sub_c : w -> w -> carry w.
- Variable w_square_c : w -> zn2z w.
- Variable w_div21 : w -> w -> w -> w * w.
- Variable w_add_mul_div : w -> w -> w -> w.
- Variable w_digits : positive.
- Variable w_zdigits : w.
- Variable ww_zdigits : zn2z w.
- Variable w_add_c : w -> w -> carry w.
- Variable w_sqrt2 : w -> w -> w * carry w.
- Variable w_pred : w -> w.
- Variable ww_pred_c : zn2z w -> carry (zn2z w).
- Variable ww_pred : zn2z w -> zn2z w.
- Variable ww_add_c : zn2z w -> zn2z w -> carry (zn2z w).
- Variable ww_add : zn2z w -> zn2z w -> zn2z w.
- Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w).
- Variable ww_add_mul_div : zn2z w -> zn2z w -> zn2z w -> zn2z w.
- Variable ww_head0 : zn2z w -> zn2z w.
- Variable ww_compare : zn2z w -> zn2z w -> comparison.
- Variable low : zn2z w -> w.
-
- Let wwBm1 := ww_Bm1 w_Bm1.
-
- Definition ww_is_even x :=
- match x with
- | W0 => true
- | WW xh xl => w_is_even xl
- end.
-
- Let w_div21c x y z :=
- match w_compare x z with
- | Eq =>
- match w_compare y z with
- Eq => (C1 w_1, w_0)
- | Gt => (C1 w_1, w_sub y z)
- | Lt => (C1 w_0, y)
- end
- | Gt =>
- let x1 := w_sub x z in
- let (q, r) := w_div21 x1 y z in
- (C1 q, r)
- | Lt =>
- let (q, r) := w_div21 x y z in
- (C0 q, r)
- end.
-
- Let w_div2s x y s :=
- match x with
- C1 x1 =>
- let x2 := w_sub x1 s in
- let (q, r) := w_div21c x2 y s in
- match q with
- C0 q1 =>
- if w_is_even q1 then
- (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), C0 r)
- else
- (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), w_add_c r s)
- | C1 q1 =>
- if w_is_even q1 then
- (C1 (w_add_mul_div (w_pred w_zdigits) w_0 q1), C0 r)
- else
- (C1 (w_add_mul_div (w_pred w_zdigits) w_0 q1), w_add_c r s)
- end
- | C0 x1 =>
- let (q, r) := w_div21c x1 y s in
- match q with
- C0 q1 =>
- if w_is_even q1 then
- (C0 (w_add_mul_div (w_pred w_zdigits) w_0 q1), C0 r)
- else
- (C0 (w_add_mul_div (w_pred w_zdigits) w_0 q1), w_add_c r s)
- | C1 q1 =>
- if w_is_even q1 then
- (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), C0 r)
- else
- (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), w_add_c r s)
- end
- end.
-
- Definition split x :=
- match x with
- | W0 => (w_0,w_0)
- | WW h l => (h,l)
- end.
-
- Definition ww_sqrt2 x y :=
- let (x1, x2) := split x in
- let (y1, y2) := split y in
- let ( q, r) := w_sqrt2 x1 x2 in
- let (q1, r1) := w_div2s r y1 q in
- match q1 with
- C0 q1 =>
- let q2 := w_square_c q1 in
- let a := WW q q1 in
- match r1 with
- C1 r2 =>
- match ww_sub_c (WW r2 y2) q2 with
- C0 r3 => (a, C1 r3)
- | C1 r3 => (a, C0 r3)
- end
- | C0 r2 =>
- match ww_sub_c (WW r2 y2) q2 with
- C0 r3 => (a, C0 r3)
- | C1 r3 =>
- let a2 := ww_add_mul_div (w_0W w_1) a W0 in
- match ww_pred_c a2 with
- C0 a3 =>
- (ww_pred a, ww_add_c a3 r3)
- | C1 a3 =>
- (ww_pred a, C0 (ww_add a3 r3))
- end
- end
- end
- | C1 q1 =>
- let a1 := WW q w_Bm1 in
- let a2 := ww_add_mul_div (w_0W w_1) a1 wwBm1 in
- (a1, ww_add_c a2 y)
- end.
-
- Definition ww_is_zero x :=
- match ww_compare W0 x with
- Eq => true
- | _ => false
- end.
-
- Definition ww_head1 x :=
- let p := ww_head0 x in
- if (ww_is_even p) then p else ww_pred p.
-
- Definition ww_sqrt x :=
- if (ww_is_zero x) then W0
- else
- let p := ww_head1 x in
- match ww_compare p W0 with
- | Gt =>
- match ww_add_mul_div p x W0 with
- W0 => W0
- | WW x1 x2 =>
- let (r, _) := w_sqrt2 x1 x2 in
- WW w_0 (w_add_mul_div
- (w_sub w_zdigits
- (low (ww_add_mul_div (ww_pred ww_zdigits)
- W0 p))) w_0 r)
- end
- | _ =>
- match x with
- W0 => W0
- | WW x1 x2 => WW w_0 (fst (w_sqrt2 x1 x2))
- end
- end.
-
-
- Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
- Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
-
- Notation "[|| x ||]" :=
- (zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99).
-
- Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x)
- (at level 0, x at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_1 : [|w_1|] = 1.
- Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1.
- Variable spec_w_zdigits : [|w_zdigits|] = Zpos w_digits.
- Variable spec_more_than_1_digit: 1 < Zpos w_digits.
-
- Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos (xO w_digits).
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB.
-
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
- Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
- Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
- Variable spec_w_is_even : forall x,
- if w_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1.
- Variable spec_w_compare : forall x y,
- w_compare x y = Z.compare [|x|] [|y|].
- Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
- Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|].
- Variable spec_w_div21 : forall a1 a2 b,
- wB/2 <= [|b|] ->
- [|a1|] < [|b|] ->
- let (q,r) := w_div21 a1 a2 b in
- [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|].
- Variable spec_w_add_mul_div : forall x y p,
- [|p|] <= Zpos w_digits ->
- [| w_add_mul_div p x y |] =
- ([|x|] * (2 ^ [|p|]) +
- [|y|] / (Z.pow 2 ((Zpos w_digits) - [|p|]))) mod wB.
- Variable spec_ww_add_mul_div : forall x y p,
- [[p]] <= Zpos (xO w_digits) ->
- [[ ww_add_mul_div p x y ]] =
- ([[x]] * (2^ [[p]]) +
- [[y]] / (2^ (Zpos (xO w_digits) - [[p]]))) mod wwB.
- Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
- Variable spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB.
- Variable spec_w_sqrt2 : forall x y,
- wB/ 4 <= [|x|] ->
- let (s,r) := w_sqrt2 x y in
- [[WW x y]] = [|s|] ^ 2 + [+|r|] /\
- [+|r|] <= 2 * [|s|].
- Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]].
- Variable spec_ww_pred_c : forall x, [-[ww_pred_c x]] = [[x]] - 1.
- Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB.
- Variable spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB.
- Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
- Variable spec_ww_compare : forall x y,
- ww_compare x y = Z.compare [[x]] [[y]].
- Variable spec_ww_head0 : forall x, 0 < [[x]] ->
- wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB.
- Variable spec_low: forall x, [|low x|] = [[x]] mod wB.
-
- Let spec_ww_Bm1 : [[wwBm1]] = wwB - 1.
- Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed.
-
- Hint Rewrite spec_w_0 spec_w_1 spec_w_WW spec_w_sub
- spec_w_add_mul_div spec_ww_Bm1 spec_w_add_c : w_rewrite.
-
- Lemma spec_ww_is_even : forall x,
- if ww_is_even x then [[x]] mod 2 = 0 else [[x]] mod 2 = 1.
-clear spec_more_than_1_digit.
-intros x; case x; simpl ww_is_even.
- reflexivity.
- simpl.
- intros w1 w2; simpl.
- unfold base.
- rewrite Zplus_mod; auto with zarith.
- rewrite (fun x y => (Zdivide_mod (x * y))); auto with zarith.
- rewrite Z.add_0_l; rewrite Zmod_mod; auto with zarith.
- apply spec_w_is_even; auto with zarith.
- apply Z.divide_mul_r; apply Zpower_divide; auto with zarith.
- Qed.
-
-
- Theorem spec_w_div21c : forall a1 a2 b,
- wB/2 <= [|b|] ->
- let (q,r) := w_div21c a1 a2 b in
- [|a1|] * wB + [|a2|] = [+|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|].
- intros a1 a2 b Hb; unfold w_div21c.
- assert (H: 0 < [|b|]); auto with zarith.
- assert (U := wB_pos w_digits).
- apply Z.lt_le_trans with (2 := Hb); auto with zarith.
- apply Z.lt_le_trans with 1; auto with zarith.
- apply Zdiv_le_lower_bound; auto with zarith.
- rewrite !spec_w_compare. repeat case Z.compare_spec.
- intros H1 H2; split.
- unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith.
- rewrite H1; rewrite H2; ring.
- autorewrite with w_rewrite; auto with zarith.
- intros H1 H2; split.
- unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith.
- rewrite H2; ring.
- destruct (spec_to_Z a2);auto with zarith.
- intros H1 H2; split.
- unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith.
- rewrite H2; rewrite Zmod_small; auto with zarith.
- ring.
- destruct (spec_to_Z a2);auto with zarith.
- rewrite spec_w_sub; auto with zarith.
- destruct (spec_to_Z a2) as [H3 H4];auto with zarith.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- assert ([|a2|] < 2 * [|b|]); auto with zarith.
- apply Z.lt_le_trans with (2 * (wB / 2)); auto with zarith.
- rewrite wB_div_2; auto.
- intros H1.
- match goal with |- context[w_div21 ?y ?z ?t] =>
- generalize (@spec_w_div21 y z t Hb H1);
- case (w_div21 y z t); simpl; autorewrite with w_rewrite;
- auto
- end.
- intros H1.
- assert (H2: [|w_sub a1 b|] < [|b|]).
- rewrite spec_w_sub; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- assert ([|a1|] < 2 * [|b|]); auto with zarith.
- apply Z.lt_le_trans with (2 * (wB / 2)); auto with zarith.
- rewrite wB_div_2; auto.
- destruct (spec_to_Z a1);auto with zarith.
- destruct (spec_to_Z a1);auto with zarith.
- match goal with |- context[w_div21 ?y ?z ?t] =>
- generalize (@spec_w_div21 y z t Hb H2);
- case (w_div21 y z t); autorewrite with w_rewrite;
- auto
- end.
- intros w0 w1; replace [+|C1 w0|] with (wB + [|w0|]).
- rewrite Zmod_small; auto with zarith.
- intros (H3, H4); split; auto.
- rewrite Z.mul_add_distr_r.
- rewrite <- Z.add_assoc; rewrite <- H3; ring.
- split; auto with zarith.
- assert ([|a1|] < 2 * [|b|]); auto with zarith.
- apply Z.lt_le_trans with (2 * (wB / 2)); auto with zarith.
- rewrite wB_div_2; auto.
- destruct (spec_to_Z a1);auto with zarith.
- destruct (spec_to_Z a1);auto with zarith.
- simpl; case wB; auto.
- Qed.
-
- Theorem C0_id: forall p, [+|C0 p|] = [|p|].
- intros p; simpl; auto.
- Qed.
-
- Theorem add_mult_div_2: forall w,
- [|w_add_mul_div (w_pred w_zdigits) w_0 w|] = [|w|] / 2.
- intros w1.
- assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1).
- rewrite spec_pred; rewrite spec_w_zdigits.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.lt_le_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_le_lin; auto with zarith.
- rewrite spec_w_add_mul_div; auto with zarith.
- autorewrite with w_rewrite rm10.
- match goal with |- context[?X - ?Y] =>
- replace (X - Y) with 1
- end.
- rewrite Z.pow_1_r; rewrite Zmod_small; auto with zarith.
- destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
- split; auto with zarith.
- apply Zdiv_lt_upper_bound; auto with zarith.
- rewrite Hp; ring.
- Qed.
-
- Theorem add_mult_div_2_plus_1: forall w,
- [|w_add_mul_div (w_pred w_zdigits) w_1 w|] =
- [|w|] / 2 + 2 ^ Zpos (w_digits - 1).
- intros w1.
- assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1).
- rewrite spec_pred; rewrite spec_w_zdigits.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.lt_le_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_le_lin; auto with zarith.
- autorewrite with w_rewrite rm10; auto with zarith.
- match goal with |- context[?X - ?Y] =>
- replace (X - Y) with 1
- end; rewrite Hp; try ring.
- rewrite Pos2Z.inj_sub_max; auto with zarith.
- rewrite Z.max_r; auto with zarith.
- rewrite Z.pow_1_r; rewrite Zmod_small; auto with zarith.
- destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
- split; auto with zarith.
- unfold base.
- match goal with |- _ < _ ^ ?X =>
- assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp
- end.
- rewrite Zpower_exp; try rewrite Z.pow_1_r; auto with zarith.
- assert (tmp: forall p, 1 + (p -1) - 1 = p - 1); auto with zarith;
- rewrite tmp; clear tmp; auto with zarith.
- match goal with |- ?X + ?Y < _ =>
- assert (Y < X); auto with zarith
- end.
- apply Zdiv_lt_upper_bound; auto with zarith.
- pattern 2 at 2; rewrite <- Z.pow_1_r; rewrite <- Zpower_exp;
- auto with zarith.
- assert (tmp: forall p, (p - 1) + 1 = p); auto with zarith;
- rewrite tmp; clear tmp; auto with zarith.
- Qed.
-
- Theorem add_mult_mult_2: forall w,
- [|w_add_mul_div w_1 w w_0|] = 2 * [|w|] mod wB.
- intros w1.
- autorewrite with w_rewrite rm10; auto with zarith.
- rewrite Z.pow_1_r; auto with zarith.
- rewrite Z.mul_comm; auto.
- Qed.
-
- Theorem ww_add_mult_mult_2: forall w,
- [[ww_add_mul_div (w_0W w_1) w W0]] = 2 * [[w]] mod wwB.
- intros w1.
- rewrite spec_ww_add_mul_div; auto with zarith.
- autorewrite with w_rewrite rm10.
- rewrite spec_w_0W; rewrite spec_w_1.
- rewrite Z.pow_1_r; auto with zarith.
- rewrite Z.mul_comm; auto.
- rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
- red; simpl; intros; discriminate.
- Qed.
-
- Theorem ww_add_mult_mult_2_plus_1: forall w,
- [[ww_add_mul_div (w_0W w_1) w wwBm1]] =
- (2 * [[w]] + 1) mod wwB.
- intros w1.
- rewrite spec_ww_add_mul_div; auto with zarith.
- rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
- rewrite Z.pow_1_r; auto with zarith.
- f_equal; auto.
- rewrite Z.mul_comm; f_equal; auto.
- autorewrite with w_rewrite rm10.
- unfold ww_digits, base.
- symmetry; apply Zdiv_unique with (r := 2 ^ (Zpos (ww_digits w_digits) - 1) -1);
- auto with zarith.
- unfold ww_digits; split; auto with zarith.
- match goal with |- 0 <= ?X - 1 =>
- assert (0 < X); auto with zarith
- end.
- apply Z.pow_pos_nonneg; auto with zarith.
- match goal with |- 0 <= ?X - 1 =>
- assert (0 < X); auto with zarith; red; reflexivity
- end.
- unfold ww_digits; autorewrite with rm10.
- assert (tmp: forall p q r, p + (q - r) = p + q - r); auto with zarith;
- rewrite tmp; clear tmp.
- assert (tmp: forall p, p + p = 2 * p); auto with zarith;
- rewrite tmp; clear tmp.
- f_equal; auto.
- pattern 2 at 2; rewrite <- Z.pow_1_r; rewrite <- Zpower_exp;
- auto with zarith.
- assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite tmp; clear tmp; auto.
- match goal with |- ?X - 1 >= 0 =>
- assert (0 < X); auto with zarith; red; reflexivity
- end.
- rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
- red; simpl; intros; discriminate.
- Qed.
-
- Theorem Zplus_mod_one: forall a1 b1, 0 < b1 -> (a1 + b1) mod b1 = a1 mod b1.
- intros a1 b1 H; rewrite Zplus_mod; auto with zarith.
- rewrite Z_mod_same; try rewrite Z.add_0_r; auto with zarith.
- apply Zmod_mod; auto.
- Qed.
-
- Lemma C1_plus_wB: forall x, [+|C1 x|] = wB + [|x|].
- unfold interp_carry; auto with zarith.
- Qed.
-
- Theorem spec_w_div2s : forall a1 a2 b,
- wB/2 <= [|b|] -> [+|a1|] <= 2 * [|b|] ->
- let (q,r) := w_div2s a1 a2 b in
- [+|a1|] * wB + [|a2|] = [+|q|] * (2 * [|b|]) + [+|r|] /\ 0 <= [+|r|] < 2 * [|b|].
- intros a1 a2 b H.
- assert (HH: 0 < [|b|]); auto with zarith.
- assert (U := wB_pos w_digits).
- apply Z.lt_le_trans with (2 := H); auto with zarith.
- apply Z.lt_le_trans with 1; auto with zarith.
- apply Zdiv_le_lower_bound; auto with zarith.
- unfold w_div2s; case a1; intros w0 H0.
- match goal with |- context[w_div21c ?y ?z ?t] =>
- generalize (@spec_w_div21c y z t H);
- case (w_div21c y z t); autorewrite with w_rewrite;
- auto
- end.
- intros c w1; case c.
- simpl interp_carry; intros w2 (Hw1, Hw2).
- match goal with |- context[w_is_even ?y] =>
- generalize (spec_w_is_even y);
- case (w_is_even y)
- end.
- repeat rewrite C0_id.
- rewrite add_mult_div_2.
- intros H1; split; auto with zarith.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1; ring.
- repeat rewrite C0_id.
- rewrite add_mult_div_2.
- rewrite spec_w_add_c; auto with zarith.
- intros H1; split; auto with zarith.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1; ring.
- intros w2; rewrite C1_plus_wB.
- intros (Hw1, Hw2).
- match goal with |- context[w_is_even ?y] =>
- generalize (spec_w_is_even y);
- case (w_is_even y)
- end.
- repeat rewrite C0_id.
- intros H1; split; auto with zarith.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1.
- repeat rewrite C0_id.
- rewrite add_mult_div_2_plus_1; unfold base.
- match goal with |- context[_ ^ ?X] =>
- assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Z.pow_1_r; auto with zarith
- end.
- rewrite Pos2Z.inj_sub_max; auto with zarith.
- rewrite Z.max_r; auto with zarith.
- ring.
- repeat rewrite C0_id.
- rewrite spec_w_add_c; auto with zarith.
- intros H1; split; auto with zarith.
- rewrite add_mult_div_2_plus_1.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1.
- unfold base.
- match goal with |- context[_ ^ ?X] =>
- assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Z.pow_1_r; auto with zarith
- end.
- rewrite Pos2Z.inj_sub_max; auto with zarith.
- rewrite Z.max_r; auto with zarith.
- ring.
- repeat rewrite C1_plus_wB in H0.
- rewrite C1_plus_wB.
- match goal with |- context[w_div21c ?y ?z ?t] =>
- generalize (@spec_w_div21c y z t H);
- case (w_div21c y z t); autorewrite with w_rewrite;
- auto
- end.
- intros c w1; case c.
- intros w2 (Hw1, Hw2); rewrite C0_id in Hw1.
- rewrite <- Zplus_mod_one in Hw1; auto with zarith.
- rewrite Zmod_small in Hw1; auto with zarith.
- match goal with |- context[w_is_even ?y] =>
- generalize (spec_w_is_even y);
- case (w_is_even y)
- end.
- repeat rewrite C0_id.
- intros H1; split; auto with zarith.
- rewrite add_mult_div_2_plus_1.
- replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
- auto with zarith.
- rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1; unfold base.
- match goal with |- context[_ ^ ?X] =>
- assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Z.pow_1_r; auto with zarith
- end.
- rewrite Pos2Z.inj_sub_max; auto with zarith.
- rewrite Z.max_r; auto with zarith.
- ring.
- repeat rewrite C0_id.
- rewrite add_mult_div_2_plus_1.
- rewrite spec_w_add_c; auto with zarith.
- intros H1; split; auto with zarith.
- replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
- auto with zarith.
- rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1; unfold base.
- match goal with |- context[_ ^ ?X] =>
- assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Z.pow_1_r; auto with zarith
- end.
- rewrite Pos2Z.inj_sub_max; auto with zarith.
- rewrite Z.max_r; auto with zarith.
- ring.
- split; auto with zarith.
- destruct (spec_to_Z b);auto with zarith.
- destruct (spec_to_Z w0);auto with zarith.
- destruct (spec_to_Z b);auto with zarith.
- destruct (spec_to_Z b);auto with zarith.
- intros w2; rewrite C1_plus_wB.
- rewrite <- Zplus_mod_one; auto with zarith.
- rewrite Zmod_small; auto with zarith.
- intros (Hw1, Hw2).
- match goal with |- context[w_is_even ?y] =>
- generalize (spec_w_is_even y);
- case (w_is_even y)
- end.
- repeat (rewrite C0_id || rewrite C1_plus_wB).
- intros H1; split; auto with zarith.
- rewrite add_mult_div_2.
- replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
- auto with zarith.
- rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1; ring.
- repeat (rewrite C0_id || rewrite C1_plus_wB).
- rewrite spec_w_add_c; auto with zarith.
- intros H1; split; auto with zarith.
- rewrite add_mult_div_2.
- replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
- auto with zarith.
- rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
- rewrite Hw1.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
- auto with zarith.
- rewrite H1; ring.
- split; auto with zarith.
- destruct (spec_to_Z b);auto with zarith.
- destruct (spec_to_Z w0);auto with zarith.
- destruct (spec_to_Z b);auto with zarith.
- destruct (spec_to_Z b);auto with zarith.
- Qed.
-
- Theorem wB_div_4: 4 * (wB / 4) = wB.
- Proof.
- unfold base.
- assert (2 ^ Zpos w_digits =
- 4 * (2 ^ (Zpos w_digits - 2))).
- change 4 with (2 ^ 2).
- rewrite <- Zpower_exp; auto with zarith.
- f_equal; auto with zarith.
- rewrite H.
- rewrite (fun x => (Z.mul_comm 4 (2 ^x))).
- rewrite Z_div_mult; auto with zarith.
- Qed.
-
- Theorem Zsquare_mult: forall p, p ^ 2 = p * p.
- intros p; change 2 with (1 + 1); rewrite Zpower_exp;
- try rewrite Z.pow_1_r; auto with zarith.
- Qed.
-
- Theorem Zsquare_pos: forall p, 0 <= p ^ 2.
- intros p; case (Z.le_gt_cases 0 p); intros H1.
- rewrite Zsquare_mult; apply Z.mul_nonneg_nonneg; auto with zarith.
- rewrite Zsquare_mult; replace (p * p) with ((- p) * (- p)); try ring.
- apply Z.mul_nonneg_nonneg; auto with zarith.
- Qed.
-
- Lemma spec_split: forall x,
- [|fst (split x)|] * wB + [|snd (split x)|] = [[x]].
- intros x; case x; simpl; autorewrite with w_rewrite;
- auto with zarith.
- Qed.
-
- Theorem mult_wwB: forall x y, [|x|] * [|y|] < wwB.
- Proof.
- intros x y; rewrite wwB_wBwB; rewrite Z.pow_2_r.
- generalize (spec_to_Z x); intros U.
- generalize (spec_to_Z y); intros U1.
- apply Z.le_lt_trans with ((wB -1 ) * (wB - 1)); auto with zarith.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- rewrite ?Z.mul_sub_distr_l, ?Z.mul_sub_distr_r; auto with zarith.
- Qed.
- Hint Resolve mult_wwB.
-
- Lemma spec_ww_sqrt2 : forall x y,
- wwB/ 4 <= [[x]] ->
- let (s,r) := ww_sqrt2 x y in
- [||WW x y||] = [[s]] ^ 2 + [+[r]] /\
- [+[r]] <= 2 * [[s]].
- intros x y H; unfold ww_sqrt2.
- repeat match goal with |- context[split ?x] =>
- generalize (spec_split x); case (split x)
- end; simpl @fst; simpl @snd.
- intros w0 w1 Hw0 w2 w3 Hw1.
- assert (U: wB/4 <= [|w2|]).
- case (Z.le_gt_cases (wB / 4) [|w2|]); auto; intros H1.
- contradict H; apply Z.lt_nge.
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- pattern wB at 1; rewrite <- wB_div_4; rewrite <- Z.mul_assoc;
- rewrite Z.mul_comm.
- rewrite Z_div_mult; auto with zarith.
- rewrite <- Hw1.
- match goal with |- _ < ?X =>
- pattern X; rewrite <- Z.add_0_r; apply beta_lex_inv;
- auto with zarith
- end.
- destruct (spec_to_Z w3);auto with zarith.
- generalize (@spec_w_sqrt2 w2 w3 U); case (w_sqrt2 w2 w3).
- intros w4 c (H1, H2).
- assert (U1: wB/2 <= [|w4|]).
- case (Z.le_gt_cases (wB/2) [|w4|]); auto with zarith.
- intros U1.
- assert (U2 : [|w4|] <= wB/2 -1); auto with zarith.
- assert (U3 : [|w4|] ^ 2 <= wB/4 * wB - wB + 1); auto with zarith.
- match goal with |- ?X ^ 2 <= ?Y =>
- rewrite Zsquare_mult;
- replace Y with ((wB/2 - 1) * (wB/2 -1))
- end.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- destruct (spec_to_Z w4);auto with zarith.
- destruct (spec_to_Z w4);auto with zarith.
- pattern wB at 4 5; rewrite <- wB_div_2.
- rewrite Z.mul_assoc.
- replace ((wB / 4) * 2) with (wB / 2).
- ring.
- pattern wB at 1; rewrite <- wB_div_4.
- change 4 with (2 * 2).
- rewrite <- Z.mul_assoc; rewrite (Z.mul_comm 2).
- rewrite Z_div_mult; try ring; auto with zarith.
- assert (U4 : [+|c|] <= wB -2); auto with zarith.
- apply Z.le_trans with (1 := H2).
- match goal with |- ?X <= ?Y =>
- replace Y with (2 * (wB/ 2 - 1)); auto with zarith
- end.
- pattern wB at 2; rewrite <- wB_div_2; auto with zarith.
- match type of H1 with ?X = _ =>
- assert (U5: X < wB / 4 * wB)
- end.
- rewrite H1; auto with zarith.
- contradict U; apply Z.lt_nge.
- apply Z.mul_lt_mono_pos_r with wB; auto with zarith.
- destruct (spec_to_Z w4);auto with zarith.
- apply Z.le_lt_trans with (2 := U5).
- unfold ww_to_Z, zn2z_to_Z.
- destruct (spec_to_Z w3);auto with zarith.
- generalize (@spec_w_div2s c w0 w4 U1 H2).
- case (w_div2s c w0 w4).
- intros c0; case c0; intros w5;
- repeat (rewrite C0_id || rewrite C1_plus_wB).
- intros c1; case c1; intros w6;
- repeat (rewrite C0_id || rewrite C1_plus_wB).
- intros (H3, H4).
- match goal with |- context [ww_sub_c ?y ?z] =>
- generalize (spec_ww_sub_c y z); case (ww_sub_c y z)
- end.
- intros z; change [-[C0 z]] with ([[z]]).
- change [+[C0 z]] with ([[z]]).
- intros H5; rewrite spec_w_square_c in H5;
- auto.
- split.
- unfold zn2z_to_Z; rewrite <- Hw1.
- unfold ww_to_Z, zn2z_to_Z in H1. rewrite H1.
- rewrite <- Hw0.
- match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
- end.
- repeat rewrite Zsquare_mult.
- rewrite wwB_wBwB; ring.
- rewrite H3.
- rewrite H5.
- unfold ww_to_Z, zn2z_to_Z.
- repeat rewrite Zsquare_mult; ring.
- rewrite H5.
- unfold ww_to_Z, zn2z_to_Z.
- match goal with |- ?X - ?Y * ?Y <= _ =>
- assert (V := Zsquare_pos Y);
- rewrite Zsquare_mult in V;
- apply Z.le_trans with X; auto with zarith;
- clear V
- end.
- match goal with |- ?X * wB + ?Y <= 2 * (?Z * wB + ?T) =>
- apply Z.le_trans with ((2 * Z - 1) * wB + wB); auto with zarith
- end.
- destruct (spec_to_Z w1);auto with zarith.
- match goal with |- ?X <= _ =>
- replace X with (2 * [|w4|] * wB); auto with zarith
- end.
- rewrite Z.mul_add_distr_l; rewrite Z.mul_assoc.
- destruct (spec_to_Z w5); auto with zarith.
- ring.
- intros z; replace [-[C1 z]] with (- wwB + [[z]]).
- 2: simpl; case wwB; auto with zarith.
- intros H5; rewrite spec_w_square_c in H5;
- auto.
- match goal with |- context [ww_pred_c ?y] =>
- generalize (spec_ww_pred_c y); case (ww_pred_c y)
- end.
- intros z1; change [-[C0 z1]] with ([[z1]]).
- rewrite ww_add_mult_mult_2.
- rewrite spec_ww_add_c.
- rewrite spec_ww_pred.
- rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]);
- auto with zarith.
- intros Hz1; rewrite Zmod_small; auto with zarith.
- match type of H5 with -?X + ?Y = ?Z =>
- assert (V: Y = Z + X);
- try (rewrite <- H5; ring)
- end.
- split.
- unfold zn2z_to_Z; rewrite <- Hw1.
- unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
- rewrite <- Hw0.
- match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
- end.
- repeat rewrite Zsquare_mult.
- rewrite wwB_wBwB; ring.
- rewrite H3.
- rewrite V.
- rewrite Hz1.
- unfold ww_to_Z; simpl zn2z_to_Z.
- repeat rewrite Zsquare_mult; ring.
- rewrite Hz1.
- destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith.
- assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)).
- assert (0 < [[WW w4 w5]]); auto with zarith.
- apply Z.lt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
- autorewrite with rm10; apply Z.mul_pos_pos; auto with zarith.
- apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
- autorewrite with rm10.
- rewrite Z.mul_comm; rewrite wB_div_2; auto with zarith.
- case (spec_to_Z w5);auto with zarith.
- case (spec_to_Z w5);auto with zarith.
- simpl.
- assert (V2 := spec_to_Z w5);auto with zarith.
- assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith.
- split; auto with zarith.
- assert (wwB <= 2 * [[WW w4 w5]]); auto with zarith.
- apply Z.le_trans with (2 * ([|w4|] * wB)).
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- rewrite Z.mul_assoc; apply Z.mul_le_mono_nonneg_r; auto with zarith.
- assert (V2 := spec_to_Z w5);auto with zarith.
- rewrite <- wB_div_2; auto with zarith.
- simpl ww_to_Z; assert (V2 := spec_to_Z w5);auto with zarith.
- assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith.
- intros z1; change [-[C1 z1]] with (-wwB + [[z1]]).
- match goal with |- context[([+[C0 ?z]])] =>
- change [+[C0 z]] with ([[z]])
- end.
- rewrite spec_ww_add; auto with zarith.
- rewrite spec_ww_pred; auto with zarith.
- rewrite ww_add_mult_mult_2.
- rename V1 into VV1.
- assert (VV2: 0 < [[WW w4 w5]]); auto with zarith.
- apply Z.lt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
- autorewrite with rm10; apply Z.mul_pos_pos; auto with zarith.
- apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
- autorewrite with rm10.
- rewrite Z.mul_comm; rewrite wB_div_2; auto with zarith.
- assert (VV3 := spec_to_Z w5);auto with zarith.
- assert (VV3 := spec_to_Z w5);auto with zarith.
- simpl.
- assert (VV3 := spec_to_Z w5);auto with zarith.
- assert (VV3: wwB <= 2 * [[WW w4 w5]]); auto with zarith.
- apply Z.le_trans with (2 * ([|w4|] * wB)).
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- rewrite Z.mul_assoc; apply Z.mul_le_mono_nonneg_r; auto with zarith.
- case (spec_to_Z w5);auto with zarith.
- rewrite <- wB_div_2; auto with zarith.
- simpl ww_to_Z; assert (V4 := spec_to_Z w5);auto with zarith.
- rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]);
- auto with zarith.
- intros Hz1; rewrite Zmod_small; auto with zarith.
- match type of H5 with -?X + ?Y = ?Z =>
- assert (V: Y = Z + X);
- try (rewrite <- H5; ring)
- end.
- match type of Hz1 with -?X + ?Y = -?X + ?Z - 1 =>
- assert (V1: Y = Z - 1);
- [replace (Z - 1) with (X + (-X + Z -1));
- [rewrite <- Hz1 | idtac]; ring
- | idtac]
- end.
- rewrite <- Zmod_unique with (q := 1) (r := -wwB + [[z1]] + [[z]]);
- auto with zarith.
- unfold zn2z_to_Z; rewrite <- Hw1.
- unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
- rewrite <- Hw0.
- split.
- match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
- end.
- repeat rewrite Zsquare_mult.
- rewrite wwB_wBwB; ring.
- rewrite H3.
- rewrite V.
- rewrite Hz1.
- unfold ww_to_Z; simpl zn2z_to_Z.
- repeat rewrite Zsquare_mult; ring.
- assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith.
- assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith.
- assert (V3 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z1);auto with zarith.
- split; auto with zarith.
- rewrite (Z.add_comm (-wwB)); rewrite <- Z.add_assoc.
- rewrite H5.
- match goal with |- 0 <= ?X + (?Y - ?Z) =>
- apply Z.le_trans with (X - Z); auto with zarith
- end.
- 2: generalize (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w6 w1)); unfold ww_to_Z; auto with zarith.
- rewrite V1.
- match goal with |- 0 <= ?X - 1 - ?Y =>
- assert (Y < X); auto with zarith
- end.
- apply Z.lt_le_trans with wwB; auto with zarith.
- intros (H3, H4).
- match goal with |- context [ww_sub_c ?y ?z] =>
- generalize (spec_ww_sub_c y z); case (ww_sub_c y z)
- end.
- intros z; change [-[C0 z]] with ([[z]]).
- match goal with |- context[([+[C1 ?z]])] =>
- replace [+[C1 z]] with (wwB + [[z]])
- end.
- 2: simpl; case wwB; auto.
- intros H5; rewrite spec_w_square_c in H5;
- auto.
- split.
- change ([||WW x y||]) with ([[x]] * wwB + [[y]]).
- rewrite <- Hw1.
- unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
- rewrite <- Hw0.
- match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
- end.
- repeat rewrite Zsquare_mult.
- rewrite wwB_wBwB; ring.
- rewrite H3.
- rewrite H5.
- unfold ww_to_Z; simpl zn2z_to_Z.
- rewrite wwB_wBwB.
- repeat rewrite Zsquare_mult; ring.
- simpl ww_to_Z.
- rewrite H5.
- simpl ww_to_Z.
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- match goal with |- ?X * ?Y + (?Z * ?Y + ?T - ?U) <= _ =>
- apply Z.le_trans with (X * Y + (Z * Y + T - 0));
- auto with zarith
- end.
- assert (V := Zsquare_pos [|w5|]);
- rewrite Zsquare_mult in V; auto with zarith.
- autorewrite with rm10.
- match goal with |- _ <= 2 * (?U * ?V + ?W) =>
- apply Z.le_trans with (2 * U * V + 0);
- auto with zarith
- end.
- match goal with |- ?X * ?Y + (?Z * ?Y + ?T) <= _ =>
- replace (X * Y + (Z * Y + T)) with ((X + Z) * Y + T);
- try ring
- end.
- apply Z.lt_le_incl; apply beta_lex_inv; auto with zarith.
- destruct (spec_to_Z w1);auto with zarith.
- destruct (spec_to_Z w5);auto with zarith.
- rewrite Z.mul_add_distr_l; auto with zarith.
- rewrite Z.mul_assoc; auto with zarith.
- intros z; replace [-[C1 z]] with (- wwB + [[z]]).
- 2: simpl; case wwB; auto with zarith.
- intros H5; rewrite spec_w_square_c in H5;
- auto.
- match goal with |- context[([+[C0 ?z]])] =>
- change [+[C0 z]] with ([[z]])
- end.
- match type of H5 with -?X + ?Y = ?Z =>
- assert (V: Y = Z + X);
- try (rewrite <- H5; ring)
- end.
- change ([||WW x y||]) with ([[x]] * wwB + [[y]]).
- simpl ww_to_Z.
- rewrite <- Hw1.
- simpl ww_to_Z in H1; rewrite H1.
- rewrite <- Hw0.
- split.
- match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
- end.
- repeat rewrite Zsquare_mult.
- rewrite wwB_wBwB; ring.
- rewrite H3.
- rewrite V.
- simpl ww_to_Z.
- rewrite wwB_wBwB.
- repeat rewrite Zsquare_mult; ring.
- rewrite V.
- simpl ww_to_Z.
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- match goal with |- (?Z * ?Y + ?T - ?U) + ?X * ?Y <= _ =>
- apply Z.le_trans with ((Z * Y + T - 0) + X * Y);
- auto with zarith
- end.
- assert (V1 := Zsquare_pos [|w5|]);
- rewrite Zsquare_mult in V1; auto with zarith.
- autorewrite with rm10.
- match goal with |- _ <= 2 * (?U * ?V + ?W) =>
- apply Z.le_trans with (2 * U * V + 0);
- auto with zarith
- end.
- match goal with |- (?Z * ?Y + ?T) + ?X * ?Y <= _ =>
- replace ((Z * Y + T) + X * Y) with ((X + Z) * Y + T);
- try ring
- end.
- apply Z.lt_le_incl; apply beta_lex_inv; auto with zarith.
- destruct (spec_to_Z w1);auto with zarith.
- destruct (spec_to_Z w5);auto with zarith.
- rewrite Z.mul_add_distr_l; auto with zarith.
- rewrite Z.mul_assoc; auto with zarith.
- Z.le_elim H2.
- intros c1 (H3, H4).
- match type of H3 with ?X = ?Y => absurd (X < Y) end.
- apply Z.le_ngt; rewrite <- H3; auto with zarith.
- rewrite Z.mul_add_distr_r.
- apply Z.lt_le_trans with ((2 * [|w4|]) * wB + 0);
- auto with zarith.
- apply beta_lex_inv; auto with zarith.
- destruct (spec_to_Z w0);auto with zarith.
- assert (V1 := spec_to_Z w5);auto with zarith.
- rewrite (Z.mul_comm wB); auto with zarith.
- assert (0 <= [|w5|] * (2 * [|w4|])); auto with zarith.
- intros c1 (H3, H4); rewrite H2 in H3.
- match type of H3 with ?X + ?Y = (?Z + ?T) * ?U + ?V =>
- assert (VV: (Y = (T * U) + V));
- [replace Y with ((X + Y) - X);
- [rewrite H3; ring | ring] | idtac]
- end.
- assert (V1 := spec_to_Z w0);auto with zarith.
- assert (V2 := spec_to_Z w5);auto with zarith.
- case V2; intros V3 _.
- Z.le_elim V3; auto with zarith.
- match type of VV with ?X = ?Y => absurd (X < Y) end.
- apply Z.le_ngt; rewrite <- VV; auto with zarith.
- apply Z.lt_le_trans with wB; auto with zarith.
- match goal with |- _ <= ?X + _ =>
- apply Z.le_trans with X; auto with zarith
- end.
- match goal with |- _ <= _ * ?X =>
- apply Z.le_trans with (1 * X); auto with zarith
- end.
- autorewrite with rm10.
- rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; auto with zarith.
- rewrite <- V3 in VV; generalize VV; autorewrite with rm10;
- clear VV; intros VV.
- rewrite spec_ww_add_c; auto with zarith.
- rewrite ww_add_mult_mult_2_plus_1.
- match goal with |- context[?X mod wwB] =>
- rewrite <- Zmod_unique with (q := 1) (r := -wwB + X)
- end; auto with zarith.
- simpl ww_to_Z.
- rewrite spec_w_Bm1; auto with zarith.
- split.
- change ([||WW x y||]) with ([[x]] * wwB + [[y]]).
- rewrite <- Hw1.
- simpl ww_to_Z in H1; rewrite H1.
- rewrite <- Hw0.
- match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
- end.
- repeat rewrite Zsquare_mult.
- rewrite wwB_wBwB; ring.
- rewrite H2.
- rewrite wwB_wBwB.
- repeat rewrite Zsquare_mult; ring.
- assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z y);auto with zarith.
- assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z y);auto with zarith.
- simpl ww_to_Z; unfold ww_to_Z.
- rewrite spec_w_Bm1; auto with zarith.
- split.
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- match goal with |- _ <= -?X + (2 * (?Z * ?T + ?U) + ?V) =>
- assert (X <= 2 * Z * T); auto with zarith
- end.
- apply Z.mul_le_mono_nonneg_r; auto with zarith.
- rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; auto with zarith.
- rewrite Z.mul_add_distr_l; auto with zarith.
- rewrite Z.mul_assoc; auto with zarith.
- match goal with |- _ + ?X < _ =>
- replace X with ((2 * (([|w4|]) + 1) * wB) - 1); try ring
- end.
- assert (2 * ([|w4|] + 1) * wB <= 2 * wwB); auto with zarith.
- rewrite <- Z.mul_assoc; apply Z.mul_le_mono_nonneg_l; auto with zarith.
- rewrite wwB_wBwB; rewrite Z.pow_2_r.
- apply Z.mul_le_mono_nonneg_r; auto with zarith.
- case (spec_to_Z w4);auto with zarith.
-Qed.
-
- Lemma spec_ww_is_zero: forall x,
- if ww_is_zero x then [[x]] = 0 else 0 < [[x]].
- intro x; unfold ww_is_zero.
- rewrite spec_ww_compare. case Z.compare_spec;
- auto with zarith.
- simpl ww_to_Z.
- assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z x);auto with zarith.
- Qed.
-
- Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2.
- pattern wwB at 1; rewrite wwB_wBwB; rewrite Z.pow_2_r.
- rewrite <- wB_div_2.
- match goal with |- context[(2 * ?X) * (2 * ?Z)] =>
- replace ((2 * X) * (2 * Z)) with ((X * Z) * 4); try ring
- end.
- rewrite Z_div_mult; auto with zarith.
- rewrite Z.mul_assoc; rewrite wB_div_2.
- rewrite wwB_div_2; ring.
- Qed.
-
-
- Lemma spec_ww_head1
- : forall x : zn2z w,
- (ww_is_even (ww_head1 x) = true) /\
- (0 < [[x]] -> wwB / 4 <= 2 ^ [[ww_head1 x]] * [[x]] < wwB).
- assert (U := wB_pos w_digits).
- intros x; unfold ww_head1.
- generalize (spec_ww_is_even (ww_head0 x)); case_eq (ww_is_even (ww_head0 x)).
- intros HH H1; rewrite HH; split; auto.
- intros H2.
- generalize (spec_ww_head0 x H2); case (ww_head0 x); autorewrite with rm10.
- intros (H3, H4); split; auto with zarith.
- apply Z.le_trans with (2 := H3).
- apply Zdiv_le_compat_l; auto with zarith.
- intros xh xl (H3, H4); split; auto with zarith.
- apply Z.le_trans with (2 := H3).
- apply Zdiv_le_compat_l; auto with zarith.
- intros H1.
- case (spec_to_w_Z (ww_head0 x)); intros Hv1 Hv2.
- assert (Hp0: 0 < [[ww_head0 x]]).
- generalize (spec_ww_is_even (ww_head0 x)); rewrite H1.
- generalize Hv1; case [[ww_head0 x]].
- rewrite Zmod_small; auto with zarith.
- intros; assert (0 < Zpos p); auto with zarith.
- red; simpl; auto.
- intros p H2; case H2; auto.
- assert (Hp: [[ww_pred (ww_head0 x)]] = [[ww_head0 x]] - 1).
- rewrite spec_ww_pred.
- rewrite Zmod_small; auto with zarith.
- intros H2; split.
- generalize (spec_ww_is_even (ww_pred (ww_head0 x)));
- case ww_is_even; auto.
- rewrite Hp.
- rewrite Zminus_mod; auto with zarith.
- rewrite H2; repeat rewrite Zmod_small; auto with zarith.
- intros H3; rewrite Hp.
- case (spec_ww_head0 x); auto; intros Hv3 Hv4.
- assert (Hu: forall u, 0 < u -> 2 * 2 ^ (u - 1) = 2 ^u).
- intros u Hu.
- pattern 2 at 1; rewrite <- Z.pow_1_r.
- rewrite <- Zpower_exp; auto with zarith.
- ring_simplify (1 + (u - 1)); auto with zarith.
- split; auto with zarith.
- apply Z.mul_le_mono_pos_r with 2; auto with zarith.
- repeat rewrite (fun x => Z.mul_comm x 2).
- rewrite wwB_4_2.
- rewrite Z.mul_assoc; rewrite Hu; auto with zarith.
- apply Z.le_lt_trans with (2 * 2 ^ ([[ww_head0 x]] - 1) * [[x]]); auto with zarith;
- rewrite Hu; auto with zarith.
- apply Z.mul_le_mono_nonneg_r; auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- Qed.
-
- Theorem wwB_4_wB_4: wwB / 4 = wB / 4 * wB.
- Proof.
- symmetry; apply Zdiv_unique with 0; auto with zarith.
- rewrite Z.mul_assoc; rewrite wB_div_4; auto with zarith.
- rewrite wwB_wBwB; ring.
- Qed.
-
- Lemma spec_ww_sqrt : forall x,
- [[ww_sqrt x]] ^ 2 <= [[x]] < ([[ww_sqrt x]] + 1) ^ 2.
- assert (U := wB_pos w_digits).
- intro x; unfold ww_sqrt.
- generalize (spec_ww_is_zero x); case (ww_is_zero x).
- simpl ww_to_Z; simpl Z.pow; unfold Z.pow_pos; simpl;
- auto with zarith.
- intros H1.
- rewrite spec_ww_compare. case Z.compare_spec;
- simpl ww_to_Z; autorewrite with rm10.
- generalize H1; case x.
- intros HH; contradict HH; simpl ww_to_Z; auto with zarith.
- intros w0 w1; simpl ww_to_Z; autorewrite with w_rewrite rm10.
- intros H2; case (spec_ww_head1 (WW w0 w1)); intros H3 H4 H5.
- generalize (H4 H2); clear H4; rewrite H5; clear H5; autorewrite with rm10.
- intros (H4, H5).
- assert (V: wB/4 <= [|w0|]).
- apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10.
- rewrite <- wwB_4_wB_4; auto.
- generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith.
- case (w_sqrt2 w0 w1); intros w2 c.
- simpl ww_to_Z; simpl @fst.
- case c; unfold interp_carry; autorewrite with rm10.
- intros w3 (H6, H7); rewrite H6.
- assert (V1 := spec_to_Z w3);auto with zarith.
- split; auto with zarith.
- apply Z.le_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith.
- match goal with |- ?X < ?Z =>
- replace Z with (X + 1); auto with zarith
- end.
- repeat rewrite Zsquare_mult; ring.
- intros w3 (H6, H7); rewrite H6.
- assert (V1 := spec_to_Z w3);auto with zarith.
- split; auto with zarith.
- apply Z.le_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith.
- match goal with |- ?X < ?Z =>
- replace Z with (X + 1); auto with zarith
- end.
- repeat rewrite Zsquare_mult; ring.
- intros HH; case (spec_to_w_Z (ww_head1 x)); auto with zarith.
- intros Hv1.
- case (spec_ww_head1 x); intros Hp1 Hp2.
- generalize (Hp2 H1); clear Hp2; intros Hp2.
- assert (Hv2: [[ww_head1 x]] <= Zpos (xO w_digits)).
- case (Z.le_gt_cases (Zpos (xO w_digits)) [[ww_head1 x]]); auto with zarith; intros HH1.
- case Hp2; intros _ HH2; contradict HH2.
- apply Z.le_ngt; unfold base.
- apply Z.le_trans with (2 ^ [[ww_head1 x]]).
- apply Zpower_le_monotone; auto with zarith.
- pattern (2 ^ [[ww_head1 x]]) at 1;
- rewrite <- (Z.mul_1_r (2 ^ [[ww_head1 x]])).
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- generalize (spec_ww_add_mul_div x W0 (ww_head1 x) Hv2);
- case ww_add_mul_div.
- simpl ww_to_Z; autorewrite with w_rewrite rm10.
- rewrite Zmod_small; auto with zarith.
- intros H2. symmetry in H2. rewrite Z.mul_eq_0 in H2. destruct H2 as [H2|H2].
- rewrite H2; unfold Z.pow, Z.pow_pos; simpl; auto with zarith.
- match type of H2 with ?X = ?Y =>
- absurd (Y < X); try (rewrite H2; auto with zarith; fail)
- end.
- apply Z.pow_pos_nonneg; auto with zarith.
- split; auto with zarith.
- case Hp2; intros _ tmp; apply Z.le_lt_trans with (2 := tmp);
- clear tmp.
- rewrite Z.mul_comm; apply Z.mul_le_mono_nonneg_r; auto with zarith.
- assert (Hv0: [[ww_head1 x]] = 2 * ([[ww_head1 x]]/2)).
- pattern [[ww_head1 x]] at 1; rewrite (Z_div_mod_eq [[ww_head1 x]] 2);
- auto with zarith.
- generalize (spec_ww_is_even (ww_head1 x)); rewrite Hp1;
- intros tmp; rewrite tmp; rewrite Z.add_0_r; auto.
- intros w0 w1; autorewrite with w_rewrite rm10.
- rewrite Zmod_small; auto with zarith.
- 2: rewrite Z.mul_comm; auto with zarith.
- intros H2.
- assert (V: wB/4 <= [|w0|]).
- apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10.
- simpl ww_to_Z in H2; rewrite H2.
- rewrite <- wwB_4_wB_4; auto with zarith.
- rewrite Z.mul_comm; auto with zarith.
- assert (V1 := spec_to_Z w1);auto with zarith.
- generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith.
- case (w_sqrt2 w0 w1); intros w2 c.
- case (spec_to_Z w2); intros HH1 HH2.
- simpl ww_to_Z; simpl @fst.
- assert (Hv3: [[ww_pred ww_zdigits]]
- = Zpos (xO w_digits) - 1).
- rewrite spec_ww_pred; rewrite spec_ww_zdigits.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.lt_le_trans with (Zpos (xO w_digits)); auto with zarith.
- unfold base; apply Zpower2_le_lin; auto with zarith.
- assert (Hv4: [[ww_head1 x]]/2 < wB).
- apply Z.le_lt_trans with (Zpos w_digits).
- apply Z.mul_le_mono_pos_r with 2; auto with zarith.
- repeat rewrite (fun x => Z.mul_comm x 2).
- rewrite <- Hv0; rewrite <- Pos2Z.inj_xO; auto.
- unfold base; apply Zpower2_lt_lin; auto with zarith.
- assert (Hv5: [[(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))]]
- = [[ww_head1 x]]/2).
- rewrite spec_ww_add_mul_div.
- simpl ww_to_Z; autorewrite with rm10.
- rewrite Hv3.
- ring_simplify (Zpos (xO w_digits) - (Zpos (xO w_digits) - 1)).
- rewrite Z.pow_1_r.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- apply Z.lt_le_trans with (1 := Hv4); auto with zarith.
- unfold base; apply Zpower_le_monotone; auto with zarith.
- split; unfold ww_digits; try rewrite Pos2Z.inj_xO; auto with zarith.
- rewrite Hv3; auto with zarith.
- assert (Hv6: [|low(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))|]
- = [[ww_head1 x]]/2).
- rewrite spec_low.
- rewrite Hv5; rewrite Zmod_small; auto with zarith.
- rewrite spec_w_add_mul_div; auto with zarith.
- rewrite spec_w_sub; auto with zarith.
- rewrite spec_w_0.
- simpl ww_to_Z; autorewrite with rm10.
- rewrite Hv6; rewrite spec_w_zdigits.
- rewrite (fun x y => Zmod_small (x - y)).
- ring_simplify (Zpos w_digits - (Zpos w_digits - [[ww_head1 x]] / 2)).
- rewrite Zmod_small.
- simpl ww_to_Z in H2; rewrite H2; auto with zarith.
- intros (H4, H5); split.
- apply Z.mul_le_mono_pos_r with (2 ^ [[ww_head1 x]]); auto with zarith.
- rewrite H4.
- apply Z.le_trans with ([|w2|] ^ 2); auto with zarith.
- rewrite Z.mul_comm.
- pattern [[ww_head1 x]] at 1;
- rewrite Hv0; auto with zarith.
- rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r;
- auto with zarith.
- assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2);
- try (intros; repeat rewrite Zsquare_mult; ring);
- rewrite tmp; clear tmp.
- apply Zpower_le_monotone3; auto with zarith.
- split; auto with zarith.
- pattern [|w2|] at 2;
- rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]] / 2)));
- auto with zarith.
- match goal with |- ?X <= ?X + ?Y =>
- assert (0 <= Y); auto with zarith
- end.
- case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]] / 2))); auto with zarith.
- case c; unfold interp_carry; autorewrite with rm10;
- intros w3; assert (V3 := spec_to_Z w3);auto with zarith.
- apply Z.mul_lt_mono_pos_r with (2 ^ [[ww_head1 x]]); auto with zarith.
- rewrite H4.
- apply Z.le_lt_trans with ([|w2|] ^ 2 + 2 * [|w2|]); auto with zarith.
- apply Z.lt_le_trans with (([|w2|] + 1) ^ 2); auto with zarith.
- match goal with |- ?X < ?Y =>
- replace Y with (X + 1); auto with zarith
- end.
- repeat rewrite (Zsquare_mult); ring.
- rewrite Z.mul_comm.
- pattern [[ww_head1 x]] at 1; rewrite Hv0.
- rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r;
- auto with zarith.
- assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2);
- try (intros; repeat rewrite Zsquare_mult; ring);
- rewrite tmp; clear tmp.
- apply Zpower_le_monotone3; auto with zarith.
- split; auto with zarith.
- pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]]/2)));
- auto with zarith.
- rewrite <- Z.add_assoc; rewrite Z.mul_add_distr_l.
- autorewrite with rm10; apply Z.add_le_mono_l; auto with zarith.
- case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]]/2))); auto with zarith.
- split; auto with zarith.
- apply Z.le_lt_trans with ([|w2|]); auto with zarith.
- apply Zdiv_le_upper_bound; auto with zarith.
- pattern [|w2|] at 1; replace [|w2|] with ([|w2|] * 2 ^0);
- auto with zarith.
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- rewrite Z.pow_0_r; autorewrite with rm10; auto.
- split; auto with zarith.
- rewrite Hv0 in Hv2; rewrite (Pos2Z.inj_xO w_digits) in Hv2; auto with zarith.
- apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_lt_lin; auto with zarith.
- rewrite spec_w_sub; auto with zarith.
- rewrite Hv6; rewrite spec_w_zdigits; auto with zarith.
- assert (Hv7: 0 < [[ww_head1 x]]/2); auto with zarith.
- rewrite Zmod_small; auto with zarith.
- split; auto with zarith.
- assert ([[ww_head1 x]]/2 <= Zpos w_digits); auto with zarith.
- apply Z.mul_le_mono_pos_r with 2; auto with zarith.
- repeat rewrite (fun x => Z.mul_comm x 2).
- rewrite <- Hv0; rewrite <- Pos2Z.inj_xO; auto with zarith.
- apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
- unfold base; apply Zpower2_lt_lin; auto with zarith.
- Qed.
-
-End DoubleSqrt.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
deleted file mode 100644
index a2df26002..000000000
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
+++ /dev/null
@@ -1,356 +0,0 @@
-
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Set Implicit Arguments.
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import DoubleType.
-Require Import DoubleBase.
-
-Local Open Scope Z_scope.
-
-Section DoubleSub.
- Variable w : Type.
- Variable w_0 : w.
- Variable w_Bm1 : w.
- Variable w_WW : w -> w -> zn2z w.
- Variable ww_Bm1 : zn2z w.
- Variable w_opp_c : w -> carry w.
- Variable w_opp_carry : w -> w.
- Variable w_pred_c : w -> carry w.
- Variable w_sub_c : w -> w -> carry w.
- Variable w_sub_carry_c : w -> w -> carry w.
- Variable w_opp : w -> w.
- Variable w_pred : w -> w.
- Variable w_sub : w -> w -> w.
- Variable w_sub_carry : w -> w -> w.
-
- (* ** Opposites ** *)
- Definition ww_opp_c x :=
- match x with
- | W0 => C0 W0
- | WW xh xl =>
- match w_opp_c xl with
- | C0 _ =>
- match w_opp_c xh with
- | C0 h => C0 W0
- | C1 h => C1 (WW h w_0)
- end
- | C1 l => C1 (WW (w_opp_carry xh) l)
- end
- end.
-
- Definition ww_opp x :=
- match x with
- | W0 => W0
- | WW xh xl =>
- match w_opp_c xl with
- | C0 _ => WW (w_opp xh) w_0
- | C1 l => WW (w_opp_carry xh) l
- end
- end.
-
- Definition ww_opp_carry x :=
- match x with
- | W0 => ww_Bm1
- | WW xh xl => w_WW (w_opp_carry xh) (w_opp_carry xl)
- end.
-
- Definition ww_pred_c x :=
- match x with
- | W0 => C1 ww_Bm1
- | WW xh xl =>
- match w_pred_c xl with
- | C0 l => C0 (w_WW xh l)
- | C1 _ =>
- match w_pred_c xh with
- | C0 h => C0 (WW h w_Bm1)
- | C1 _ => C1 ww_Bm1
- end
- end
- end.
-
- Definition ww_pred x :=
- match x with
- | W0 => ww_Bm1
- | WW xh xl =>
- match w_pred_c xl with
- | C0 l => w_WW xh l
- | C1 l => WW (w_pred xh) w_Bm1
- end
- end.
-
- Definition ww_sub_c x y :=
- match y, x with
- | W0, _ => C0 x
- | WW yh yl, W0 => ww_opp_c (WW yh yl)
- | WW yh yl, WW xh xl =>
- match w_sub_c xl yl with
- | C0 l =>
- match w_sub_c xh yh with
- | C0 h => C0 (w_WW h l)
- | C1 h => C1 (WW h l)
- end
- | C1 l =>
- match w_sub_carry_c xh yh with
- | C0 h => C0 (WW h l)
- | C1 h => C1 (WW h l)
- end
- end
- end.
-
- Definition ww_sub x y :=
- match y, x with
- | W0, _ => x
- | WW yh yl, W0 => ww_opp (WW yh yl)
- | WW yh yl, WW xh xl =>
- match w_sub_c xl yl with
- | C0 l => w_WW (w_sub xh yh) l
- | C1 l => WW (w_sub_carry xh yh) l
- end
- end.
-
- Definition ww_sub_carry_c x y :=
- match y, x with
- | W0, W0 => C1 ww_Bm1
- | W0, WW xh xl => ww_pred_c (WW xh xl)
- | WW yh yl, W0 => C1 (ww_opp_carry (WW yh yl))
- | WW yh yl, WW xh xl =>
- match w_sub_carry_c xl yl with
- | C0 l =>
- match w_sub_c xh yh with
- | C0 h => C0 (w_WW h l)
- | C1 h => C1 (WW h l)
- end
- | C1 l =>
- match w_sub_carry_c xh yh with
- | C0 h => C0 (w_WW h l)
- | C1 h => C1 (w_WW h l)
- end
- end
- end.
-
- Definition ww_sub_carry x y :=
- match y, x with
- | W0, W0 => ww_Bm1
- | W0, WW xh xl => ww_pred (WW xh xl)
- | WW yh yl, W0 => ww_opp_carry (WW yh yl)
- | WW yh yl, WW xh xl =>
- match w_sub_carry_c xl yl with
- | C0 l => w_WW (w_sub xh yh) l
- | C1 l => w_WW (w_sub_carry xh yh) l
- end
- end.
-
- (*Section DoubleProof.*)
- Variable w_digits : positive.
- Variable w_to_Z : w -> Z.
-
-
- Notation wB := (base w_digits).
- Notation wwB := (base (ww_digits w_digits)).
- Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
- Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
-
- Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, c at level 99).
-
- Variable spec_w_0 : [|w_0|] = 0.
- Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1.
- Variable spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1.
- Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
- Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
-
- Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|].
- Variable spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB.
- Variable spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1.
-
- Variable spec_pred_c : forall x, [-|w_pred_c x|] = [|x|] - 1.
- Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|].
- Variable spec_sub_carry_c :
- forall x y, [-|w_sub_carry_c x y|] = [|x|] - [|y|] - 1.
-
- Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB.
- Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
- Variable spec_sub_carry :
- forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB.
-
-
- Lemma spec_ww_opp_c : forall x, [-[ww_opp_c x]] = -[[x]].
- Proof.
- destruct x as [ |xh xl];simpl. reflexivity.
- rewrite Z.opp_add_distr;generalize (spec_opp_c xl);destruct (w_opp_c xl)
- as [l|l];intros H;unfold interp_carry in H;rewrite <- H;
- rewrite <- Z.mul_opp_l.
- assert ([|l|] = 0).
- assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
- rewrite H0;generalize (spec_opp_c xh);destruct (w_opp_c xh)
- as [h|h];intros H1;unfold interp_carry in *;rewrite <- H1.
- assert ([|h|] = 0).
- assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega.
- rewrite H2;reflexivity.
- simpl ww_to_Z;rewrite wwB_wBwB;rewrite spec_w_0;ring.
- unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB;rewrite spec_opp_carry;
- ring.
- Qed.
-
- Lemma spec_ww_opp : forall x, [[ww_opp x]] = (-[[x]]) mod wwB.
- Proof.
- destruct x as [ |xh xl];simpl. reflexivity.
- rewrite Z.opp_add_distr, <- Z.mul_opp_l.
- generalize (spec_opp_c xl);destruct (w_opp_c xl)
- as [l|l];intros H;unfold interp_carry in H;rewrite <- H;simpl ww_to_Z.
- rewrite spec_w_0;rewrite Z.add_0_r;rewrite wwB_wBwB.
- assert ([|l|] = 0).
- assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
- rewrite H0;rewrite Z.add_0_r; rewrite Z.pow_2_r;
- rewrite Zmult_mod_distr_r;try apply lt_0_wB.
- rewrite spec_opp;trivial.
- apply Zmod_unique with (q:= -1).
- exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW (w_opp_carry xh) l)).
- rewrite spec_opp_carry;rewrite wwB_wBwB;ring.
- Qed.
-
- Lemma spec_ww_opp_carry : forall x, [[ww_opp_carry x]] = wwB - [[x]] - 1.
- Proof.
- destruct x as [ |xh xl];simpl. rewrite spec_ww_Bm1;ring.
- rewrite spec_w_WW;simpl;repeat rewrite spec_opp_carry;rewrite wwB_wBwB;ring.
- Qed.
-
- Lemma spec_ww_pred_c : forall x, [-[ww_pred_c x]] = [[x]] - 1.
- Proof.
- destruct x as [ |xh xl];unfold ww_pred_c.
- unfold interp_carry;rewrite spec_ww_Bm1;simpl ww_to_Z;ring.
- simpl ww_to_Z;replace (([|xh|]*wB+[|xl|])-1) with ([|xh|]*wB+([|xl|]-1)).
- 2:ring. generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];
- intros H;unfold interp_carry in H;rewrite <- H. simpl;apply spec_w_WW.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- assert ([|l|] = wB - 1).
- assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
- rewrite H0;change ([|xh|] + -1) with ([|xh|] - 1).
- generalize (spec_pred_c xh);destruct (w_pred_c xh) as [h|h];
- intros H1;unfold interp_carry in H1;rewrite <- H1.
- simpl;rewrite spec_w_Bm1;ring.
- assert ([|h|] = wB - 1).
- assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega.
- rewrite H2;unfold interp_carry;rewrite spec_ww_Bm1;rewrite wwB_wBwB;ring.
- Qed.
-
- Lemma spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]].
- Proof.
- destruct y as [ |yh yl];simpl. ring.
- destruct x as [ |xh xl];simpl. exact (spec_ww_opp_c (WW yh yl)).
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
- with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|])). 2:ring.
- generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl) as [l|l];intros H;
- unfold interp_carry in H;rewrite <- H.
- generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1;
- unfold interp_carry in H1;rewrite <- H1;unfold interp_carry;
- try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1).
- generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h];
- intros H1;unfold interp_carry in *;rewrite <- H1;simpl ww_to_Z;
- try rewrite wwB_wBwB;ring.
- Qed.
-
- Lemma spec_ww_sub_carry_c :
- forall x y, [-[ww_sub_carry_c x y]] = [[x]] - [[y]] - 1.
- Proof.
- destruct y as [ |yh yl];simpl.
- unfold Z.sub;simpl;rewrite Z.add_0_r;exact (spec_ww_pred_c x).
- destruct x as [ |xh xl].
- unfold interp_carry;rewrite spec_w_WW;simpl ww_to_Z;rewrite wwB_wBwB;
- repeat rewrite spec_opp_carry;ring.
- simpl ww_to_Z.
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1)
- with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|]-1)). 2:ring.
- generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl)
- as [l|l];intros H;unfold interp_carry in H;rewrite <- H.
- generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1;
- unfold interp_carry in H1;rewrite <- H1;unfold interp_carry;
- try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1).
- generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h];
- intros H1;unfold interp_carry in *;rewrite <- H1;try rewrite spec_w_WW;
- simpl ww_to_Z; try rewrite wwB_wBwB;ring.
- Qed.
-
- Lemma spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB.
- Proof.
- destruct x as [ |xh xl];simpl.
- apply Zmod_unique with (-1). apply spec_ww_to_Z;trivial.
- rewrite spec_ww_Bm1;ring.
- replace ([|xh|]*wB + [|xl|] - 1) with ([|xh|]*wB + ([|xl|] - 1)). 2:ring.
- generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];intro H;
- unfold interp_carry in H;rewrite <- H;simpl ww_to_Z.
- rewrite Zmod_small. apply spec_w_WW.
- exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh l)).
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- change ([|xh|] + -1) with ([|xh|] - 1).
- assert ([|l|] = wB - 1).
- assert (H1:= spec_to_Z l);assert (H2:= spec_to_Z xl);omega.
- rewrite (mod_wwB w_digits w_to_Z);trivial.
- rewrite spec_pred;rewrite spec_w_Bm1;rewrite <- H0;trivial.
- Qed.
-
- Lemma spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
- Proof.
- destruct y as [ |yh yl];simpl.
- ring_simplify ([[x]] - 0);rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
- destruct x as [ |xh xl];simpl. exact (spec_ww_opp (WW yh yl)).
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
- with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|])). 2:ring.
- generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl)as[l|l];intros H;
- unfold interp_carry in H;rewrite <- H.
- rewrite spec_w_WW;rewrite (mod_wwB w_digits w_to_Z spec_to_Z).
- rewrite spec_sub;trivial.
- simpl ww_to_Z;rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial.
- Qed.
-
- Lemma spec_ww_sub_carry :
- forall x y, [[ww_sub_carry x y]] = ([[x]] - [[y]] - 1) mod wwB.
- Proof.
- destruct y as [ |yh yl];simpl.
- ring_simplify ([[x]] - 0);exact (spec_ww_pred x).
- destruct x as [ |xh xl];simpl.
- apply Zmod_unique with (-1).
- apply spec_ww_to_Z;trivial.
- fold (ww_opp_carry (WW yh yl)).
- rewrite (spec_ww_opp_carry (WW yh yl));simpl ww_to_Z;ring.
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1)
- with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|] - 1)). 2:ring.
- generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl)as[l|l];
- intros H;unfold interp_carry in H;rewrite <- H;rewrite spec_w_WW.
- rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub;trivial.
- rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
- rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial.
- Qed.
-
-(* End DoubleProof. *)
-
-End DoubleSub.
-
-
-
-
-
diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v
index 0e58b8155..ba55003f7 100644
--- a/theories/Numbers/Cyclic/Int31/Cyclic31.v
+++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v
@@ -18,13 +18,16 @@ Require Export Int31.
Require Import Znumtheory.
Require Import Zgcd_alt.
Require Import Zpow_facts.
-Require Import BigNumPrelude.
Require Import CyclicAxioms.
Require Import ROmega.
+Declare ML Module "int31_syntax_plugin".
+
Local Open Scope nat_scope.
Local Open Scope int31_scope.
+Local Hint Resolve Z.lt_gt Z.div_pos : zarith.
+
Section Basics.
(** * Basic results about [iszero], [shiftl], [shiftr] *)
@@ -455,12 +458,19 @@ Section Basics.
rewrite Z.succ_double_spec; auto with zarith.
Qed.
- Lemma phi_bounded : forall x, (0 <= phi x < 2 ^ (Z.of_nat size))%Z.
+ Lemma phi_nonneg : forall x, (0 <= phi x)%Z.
Proof.
intros.
rewrite <- phibis_aux_equiv.
- split.
apply phibis_aux_pos.
+ Qed.
+
+ Hint Resolve phi_nonneg : zarith.
+
+ Lemma phi_bounded : forall x, (0 <= phi x < 2 ^ (Z.of_nat size))%Z.
+ Proof.
+ intros. split; [auto with zarith|].
+ rewrite <- phibis_aux_equiv.
change x with (nshiftr x (size-size)).
apply phibis_aux_bounded; auto.
Qed.
@@ -1624,6 +1634,37 @@ Section Int31_Specs.
rewrite Z.mul_comm, Z_div_mult; auto with zarith.
Qed.
+ Lemma shift_unshift_mod_2 : forall n p a, 0 <= p <= n ->
+ ((a * 2 ^ (n - p)) mod (2^n) / 2 ^ (n - p)) mod (2^n) =
+ a mod 2 ^ p.
+ Proof.
+ intros.
+ rewrite Zmod_small.
+ rewrite Zmod_eq by (auto with zarith).
+ unfold Z.sub at 1.
+ rewrite Z_div_plus_full_l
+ by (cut (0 < 2^(n-p)); auto with zarith).
+ assert (2^n = 2^(n-p)*2^p).
+ rewrite <- Zpower_exp by (auto with zarith).
+ replace (n-p+p) with n; auto with zarith.
+ rewrite H0.
+ rewrite <- Zdiv_Zdiv, Z_div_mult by (auto with zarith).
+ rewrite (Z.mul_comm (2^(n-p))), Z.mul_assoc.
+ rewrite <- Z.mul_opp_l.
+ rewrite Z_div_mult by (auto with zarith).
+ symmetry; apply Zmod_eq; auto with zarith.
+
+ remember (a * 2 ^ (n - p)) as b.
+ destruct (Z_mod_lt b (2^n)); auto with zarith.
+ split.
+ apply Z_div_pos; auto with zarith.
+ apply Zdiv_lt_upper_bound; auto with zarith.
+ apply Z.lt_le_trans with (2^n); auto with zarith.
+ rewrite <- (Z.mul_1_r (2^n)) at 1.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ cut (0 < 2 ^ (n-p)); auto with zarith.
+ Qed.
+
Lemma spec_pos_mod : forall w p,
[|ZnZ.pos_mod p w|] = [|w|] mod (2 ^ [|p|]).
Proof.
@@ -1654,7 +1695,7 @@ Section Int31_Specs.
rewrite spec_add_mul_div by (rewrite H4; auto with zarith).
change [|0|] with 0%Z; rewrite Zdiv_0_l, Z.add_0_r.
rewrite H4.
- apply shift_unshift_mod_2; auto with zarith.
+ apply shift_unshift_mod_2; simpl; auto with zarith.
Qed.
@@ -1973,32 +2014,24 @@ Section Int31_Specs.
assert (Hp2: 0 < [|2|]) by exact (eq_refl Lt).
intros Hi Hj Hij H31 Hrec; rewrite sqrt31_step_def.
rewrite spec_compare, div31_phi; auto.
- case Z.compare_spec; auto; intros Hc;
+ case Z.compare_spec; auto; intros Hc;
try (split; auto; apply sqrt_test_true; auto with zarith; fail).
- apply Hrec; repeat rewrite div31_phi; auto with zarith.
- replace [|(j + fst (i / j)%int31)|] with ([|j|] + [|i|] / [|j|]).
- split.
+ assert (E : [|(j + fst (i / j)%int31)|] = [|j|] + [|i|] / [|j|]).
+ { rewrite spec_add, div31_phi; auto using Z.mod_small with zarith. }
+ apply Hrec; rewrite !div31_phi, E; auto using sqrt_main with zarith.
+ split; try apply sqrt_test_false; auto with zarith.
apply Z.le_succ_l in Hj. change (1 <= [|j|]) in Hj.
Z.le_elim Hj.
- replace ([|j|] + [|i|]/[|j|]) with
- (1 * 2 + (([|j|] - 2) + [|i|] / [|j|])); try ring.
- rewrite Z_div_plus_full_l; auto with zarith.
- assert (0 <= [|i|]/ [|j|]) by (apply Z_div_pos; auto with zarith).
- assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / [|2|]) ; auto with zarith.
- rewrite <- Hj, Zdiv_1_r.
- replace (1 + [|i|])%Z with (1 * 2 + ([|i|] - 1))%Z; try ring.
- rewrite Z_div_plus_full_l; auto with zarith.
- assert (0 <= ([|i|] - 1) /2)%Z by (apply Z_div_pos; auto with zarith).
- change ([|2|]) with 2%Z; auto with zarith.
- apply sqrt_test_false; auto with zarith.
- rewrite spec_add, div31_phi; auto.
- symmetry; apply Zmod_small.
- split; auto with zarith.
- replace [|j + fst (i / j)%int31|] with ([|j|] + [|i|] / [|j|]).
- apply sqrt_main; auto with zarith.
- rewrite spec_add, div31_phi; auto.
- symmetry; apply Zmod_small.
- split; auto with zarith.
+ - replace ([|j|] + [|i|]/[|j|]) with
+ (1 * 2 + (([|j|] - 2) + [|i|] / [|j|])) by ring.
+ rewrite Z_div_plus_full_l; auto with zarith.
+ assert (0 <= [|i|]/ [|j|]) by auto with zarith.
+ assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / [|2|]); auto with zarith.
+ - rewrite <- Hj, Zdiv_1_r.
+ replace (1 + [|i|]) with (1 * 2 + ([|i|] - 1)) by ring.
+ rewrite Z_div_plus_full_l; auto with zarith.
+ assert (0 <= ([|i|] - 1) /2) by auto with zarith.
+ change ([|2|]) with 2; auto with zarith.
Qed.
Lemma iter31_sqrt_correct n rec i j: 0 < [|i|] -> 0 < [|j|] ->
@@ -2078,11 +2111,12 @@ Section Int31_Specs.
case (phi_bounded j); intros Hbj _.
case (phi_bounded il); intros Hbil _.
case (phi_bounded ih); intros Hbih Hbih1.
- assert (([|ih|] < [|j|] + 1)%Z); auto with zarith.
+ assert ([|ih|] < [|j|] + 1); auto with zarith.
apply Z.square_lt_simpl_nonneg; auto with zarith.
- repeat rewrite <-Z.pow_2_r; apply Z.le_lt_trans with (2 := H1).
- apply Z.le_trans with ([|ih|] * base)%Z; unfold phi2, base;
- try rewrite Z.pow_2_r; auto with zarith.
+ rewrite <- ?Z.pow_2_r; apply Z.le_lt_trans with (2 := H1).
+ apply Z.le_trans with ([|ih|] * wB).
+ - rewrite ? Z.pow_2_r; auto with zarith.
+ - unfold phi2. change base with wB; auto with zarith.
Qed.
Lemma div312_phi ih il j: (2^30 <= [|j|] -> [|ih|] < [|j|] ->
@@ -2104,90 +2138,89 @@ Section Int31_Specs.
Proof.
assert (Hp2: (0 < [|2|])%Z) by exact (eq_refl Lt).
intros Hih Hj Hij Hrec; rewrite sqrt312_step_def.
- assert (H1: ([|ih|] <= [|j|])%Z) by (apply sqrt312_lower_bound with il; auto).
+ assert (H1: ([|ih|] <= [|j|])) by (apply sqrt312_lower_bound with il; auto).
case (phi_bounded ih); intros Hih1 _.
case (phi_bounded il); intros Hil1 _.
case (phi_bounded j); intros _ Hj1.
assert (Hp3: (0 < phi2 ih il)).
- unfold phi2; apply Z.lt_le_trans with ([|ih|] * base)%Z; auto with zarith.
- apply Z.mul_pos_pos; auto with zarith.
- apply Z.lt_le_trans with (2:= Hih); auto with zarith.
+ { unfold phi2; apply Z.lt_le_trans with ([|ih|] * base); auto with zarith.
+ apply Z.mul_pos_pos; auto with zarith.
+ apply Z.lt_le_trans with (2:= Hih); auto with zarith. }
rewrite spec_compare. case Z.compare_spec; intros Hc1.
- split; auto.
- apply sqrt_test_true; auto.
- unfold phi2, base; auto with zarith.
- unfold phi2; rewrite Hc1.
- assert (0 <= [|il|]/[|j|]) by (apply Z_div_pos; auto with zarith).
- rewrite Z.mul_comm, Z_div_plus_full_l; unfold base; auto with zarith.
- simpl wB in Hj1. unfold Z.pow_pos in Hj1. simpl in Hj1. auto with zarith.
- case (Z.le_gt_cases (2 ^ 30) [|j|]); intros Hjj.
- rewrite spec_compare; case Z.compare_spec;
- rewrite div312_phi; auto; intros Hc;
- try (split; auto; apply sqrt_test_true; auto with zarith; fail).
- apply Hrec.
- assert (Hf1: 0 <= phi2 ih il/ [|j|]) by (apply Z_div_pos; auto with zarith).
- apply Z.le_succ_l in Hj. change (1 <= [|j|]) in Hj.
- Z.le_elim Hj.
- 2: contradict Hc; apply Z.le_ngt; rewrite <- Hj, Zdiv_1_r; auto with zarith.
- assert (Hf3: 0 < ([|j|] + phi2 ih il / [|j|]) / 2).
- replace ([|j|] + phi2 ih il/ [|j|])%Z with
- (1 * 2 + (([|j|] - 2) + phi2 ih il / [|j|])); try ring.
- rewrite Z_div_plus_full_l; auto with zarith.
- assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ; auto with zarith.
- assert (Hf4: ([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]).
- apply sqrt_test_false; auto with zarith.
- generalize (spec_add_c j (fst (div3121 ih il j))).
- unfold interp_carry; case add31c; intros r;
- rewrite div312_phi; auto with zarith.
- rewrite div31_phi; change [|2|] with 2%Z; auto with zarith.
- intros HH; rewrite HH; clear HH; auto with zarith.
- rewrite spec_add, div31_phi; change [|2|] with 2%Z; auto.
- rewrite Z.mul_1_l; intros HH.
- rewrite Z.add_comm, <- Z_div_plus_full_l; auto with zarith.
- change (phi v30 * 2) with (2 ^ Z.of_nat size).
- rewrite HH, Zmod_small; auto with zarith.
- replace (phi
- match j +c fst (div3121 ih il j) with
- | C0 m1 => fst (m1 / 2)%int31
- | C1 m1 => fst (m1 / 2)%int31 + v30
- end) with ((([|j|] + (phi2 ih il)/([|j|]))/2)).
- apply sqrt_main; auto with zarith.
- generalize (spec_add_c j (fst (div3121 ih il j))).
- unfold interp_carry; case add31c; intros r;
- rewrite div312_phi; auto with zarith.
- rewrite div31_phi; auto with zarith.
- intros HH; rewrite HH; auto with zarith.
- intros HH; rewrite <- HH.
- change (1 * 2 ^ Z.of_nat size) with (phi (v30) * 2).
- rewrite Z_div_plus_full_l; auto with zarith.
- rewrite Z.add_comm.
- rewrite spec_add, Zmod_small.
- rewrite div31_phi; auto.
- split; auto with zarith.
- case (phi_bounded (fst (r/2)%int31));
- case (phi_bounded v30); auto with zarith.
- rewrite div31_phi; change (phi 2) with 2%Z; auto.
- change (2 ^Z.of_nat size) with (base/2 + phi v30).
- assert (phi r / 2 < base/2); auto with zarith.
- apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
- change (base/2 * 2) with base.
- apply Z.le_lt_trans with (phi r).
- rewrite Z.mul_comm; apply Z_mult_div_ge; auto with zarith.
- case (phi_bounded r); auto with zarith.
- contradict Hij; apply Z.le_ngt.
- assert ((1 + [|j|]) <= 2 ^ 30); auto with zarith.
- apply Z.le_trans with ((2 ^ 30) * (2 ^ 30)); auto with zarith.
- assert (0 <= 1 + [|j|]); auto with zarith.
- apply Z.mul_le_mono_nonneg; auto with zarith.
- change ((2 ^ 30) * (2 ^ 30)) with ((2 ^ 29) * base).
- apply Z.le_trans with ([|ih|] * base); auto with zarith.
- unfold phi2, base; auto with zarith.
- split; auto.
- apply sqrt_test_true; auto.
- unfold phi2, base; auto with zarith.
- apply Z.le_ge; apply Z.le_trans with (([|j|] * base)/[|j|]).
- rewrite Z.mul_comm, Z_div_mult; auto with zarith.
- apply Z.ge_le; apply Z_div_ge; auto with zarith.
+ - split; auto.
+ apply sqrt_test_true; auto.
+ + unfold phi2, base; auto with zarith.
+ + unfold phi2; rewrite Hc1.
+ assert (0 <= [|il|]/[|j|]) by (apply Z_div_pos; auto with zarith).
+ rewrite Z.mul_comm, Z_div_plus_full_l; auto with zarith.
+ change base with wB. auto with zarith.
+ - case (Z.le_gt_cases (2 ^ 30) [|j|]); intros Hjj.
+ + rewrite spec_compare; case Z.compare_spec;
+ rewrite div312_phi; auto; intros Hc;
+ try (split; auto; apply sqrt_test_true; auto with zarith; fail).
+ apply Hrec.
+ * assert (Hf1: 0 <= phi2 ih il/ [|j|]) by auto with zarith.
+ apply Z.le_succ_l in Hj. change (1 <= [|j|]) in Hj.
+ Z.le_elim Hj;
+ [ | contradict Hc; apply Z.le_ngt;
+ rewrite <- Hj, Zdiv_1_r; auto with zarith ].
+ assert (Hf3: 0 < ([|j|] + phi2 ih il / [|j|]) / 2).
+ { replace ([|j|] + phi2 ih il/ [|j|]) with
+ (1 * 2 + (([|j|] - 2) + phi2 ih il / [|j|])); try ring.
+ rewrite Z_div_plus_full_l; auto with zarith.
+ assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ;
+ auto with zarith. }
+ assert (Hf4: ([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]).
+ { apply sqrt_test_false; auto with zarith. }
+ generalize (spec_add_c j (fst (div3121 ih il j))).
+ unfold interp_carry; case add31c; intros r;
+ rewrite div312_phi; auto with zarith.
+ { rewrite div31_phi; change [|2|] with 2; auto with zarith.
+ intros HH; rewrite HH; clear HH; auto with zarith. }
+ { rewrite spec_add, div31_phi; change [|2|] with 2; auto.
+ rewrite Z.mul_1_l; intros HH.
+ rewrite Z.add_comm, <- Z_div_plus_full_l; auto with zarith.
+ change (phi v30 * 2) with (2 ^ Z.of_nat size).
+ rewrite HH, Zmod_small; auto with zarith. }
+ * replace (phi _) with (([|j|] + (phi2 ih il)/([|j|]))/2);
+ [ apply sqrt_main; auto with zarith | ].
+ generalize (spec_add_c j (fst (div3121 ih il j))).
+ unfold interp_carry; case add31c; intros r;
+ rewrite div312_phi; auto with zarith.
+ { rewrite div31_phi; auto with zarith.
+ intros HH; rewrite HH; auto with zarith. }
+ { intros HH; rewrite <- HH.
+ change (1 * 2 ^ Z.of_nat size) with (phi (v30) * 2).
+ rewrite Z_div_plus_full_l; auto with zarith.
+ rewrite Z.add_comm.
+ rewrite spec_add, Zmod_small.
+ - rewrite div31_phi; auto.
+ - split; auto with zarith.
+ + case (phi_bounded (fst (r/2)%int31));
+ case (phi_bounded v30); auto with zarith.
+ + rewrite div31_phi; change (phi 2) with 2; auto.
+ change (2 ^Z.of_nat size) with (base/2 + phi v30).
+ assert (phi r / 2 < base/2); auto with zarith.
+ apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
+ change (base/2 * 2) with base.
+ apply Z.le_lt_trans with (phi r).
+ * rewrite Z.mul_comm; apply Z_mult_div_ge; auto with zarith.
+ * case (phi_bounded r); auto with zarith. }
+ + contradict Hij; apply Z.le_ngt.
+ assert ((1 + [|j|]) <= 2 ^ 30); auto with zarith.
+ apply Z.le_trans with ((2 ^ 30) * (2 ^ 30)); auto with zarith.
+ * assert (0 <= 1 + [|j|]); auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ * change ((2 ^ 30) * (2 ^ 30)) with ((2 ^ 29) * base).
+ apply Z.le_trans with ([|ih|] * base);
+ change wB with base in *; auto with zarith.
+ unfold phi2, base; auto with zarith.
+ - split; auto.
+ apply sqrt_test_true; auto.
+ + unfold phi2, base; auto with zarith.
+ + apply Z.le_ge; apply Z.le_trans with (([|j|] * base)/[|j|]).
+ * rewrite Z.mul_comm, Z_div_mult; auto with zarith.
+ * apply Z.ge_le; apply Z_div_ge; auto with zarith.
Qed.
Lemma iter312_sqrt_correct n rec ih il j:
@@ -2209,7 +2242,7 @@ Section Int31_Specs.
intros j3 Hj3 Hpj3.
apply HHrec; auto.
rewrite Nat2Z.inj_succ, Z.pow_succ_r.
- apply Z.le_trans with (2 ^Z.of_nat n + [|j2|])%Z; auto with zarith.
+ apply Z.le_trans with (2 ^Z.of_nat n + [|j2|]); auto with zarith.
apply Nat2Z.is_nonneg.
Qed.
diff --git a/theories/Numbers/Cyclic/ZModulo/ZModulo.v b/theories/Numbers/Cyclic/ZModulo/ZModulo.v
index 04fc5a8df..a3d7edbf4 100644
--- a/theories/Numbers/Cyclic/ZModulo/ZModulo.v
+++ b/theories/Numbers/Cyclic/ZModulo/ZModulo.v
@@ -18,7 +18,7 @@ Set Implicit Arguments.
Require Import Bool.
Require Import ZArith.
Require Import Znumtheory.
-Require Import BigNumPrelude.
+Require Import Zpow_facts.
Require Import DoubleType.
Require Import CyclicAxioms.
@@ -48,13 +48,14 @@ Section ZModulo.
Lemma spec_more_than_1_digit: 1 < Zpos digits.
Proof.
- generalize digits_ne_1; destruct digits; auto.
+ generalize digits_ne_1; destruct digits; red; auto.
destruct 1; auto.
Qed.
Let digits_gt_1 := spec_more_than_1_digit.
Lemma wB_pos : wB > 0.
Proof.
+ apply Z.lt_gt.
unfold wB, base; auto with zarith.
Qed.
Hint Resolve wB_pos.
@@ -558,7 +559,7 @@ Section ZModulo.
apply Zmod_small.
generalize (Z_mod_lt [|w|] (2 ^ [|p|])); intros.
split.
- destruct H; auto with zarith.
+ destruct H; auto using Z.lt_gt with zarith.
apply Z.le_lt_trans with [|w|]; auto with zarith.
apply Zmod_le; auto with zarith.
Qed.
diff --git a/theories/Numbers/Integer/BigZ/BigZ.v b/theories/Numbers/Integer/BigZ/BigZ.v
deleted file mode 100644
index 7c76011f2..000000000
--- a/theories/Numbers/Integer/BigZ/BigZ.v
+++ /dev/null
@@ -1,208 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Require Export BigN.
-Require Import ZProperties ZDivFloor ZSig ZSigZAxioms ZMake.
-
-(** * [BigZ] : arbitrary large efficient integers.
-
- The following [BigZ] module regroups both the operations and
- all the abstract properties:
-
- - [ZMake.Make BigN] provides the operations and basic specs w.r.t. ZArith
- - [ZTypeIsZAxioms] shows (mainly) that these operations implement
- the interface [ZAxioms]
- - [ZProp] adds all generic properties derived from [ZAxioms]
- - [MinMax*Properties] provides properties of [min] and [max]
-
-*)
-
-Delimit Scope bigZ_scope with bigZ.
-
-Module BigZ <: ZType <: OrderedTypeFull <: TotalOrder :=
- ZMake.Make BigN
- <+ ZTypeIsZAxioms
- <+ ZBasicProp [no inline] <+ ZExtraProp [no inline]
- <+ HasEqBool2Dec [no inline]
- <+ MinMaxLogicalProperties [no inline]
- <+ MinMaxDecProperties [no inline].
-
-(** For precision concerning the above scope handling, see comment in BigN *)
-
-(** Notations about [BigZ] *)
-
-Local Open Scope bigZ_scope.
-
-Notation bigZ := BigZ.t.
-Bind Scope bigZ_scope with bigZ BigZ.t BigZ.t_.
-Arguments BigZ.Pos _%bigN.
-Arguments BigZ.Neg _%bigN.
-Local Notation "0" := BigZ.zero : bigZ_scope.
-Local Notation "1" := BigZ.one : bigZ_scope.
-Local Notation "2" := BigZ.two : bigZ_scope.
-Infix "+" := BigZ.add : bigZ_scope.
-Infix "-" := BigZ.sub : bigZ_scope.
-Notation "- x" := (BigZ.opp x) : bigZ_scope.
-Infix "*" := BigZ.mul : bigZ_scope.
-Infix "/" := BigZ.div : bigZ_scope.
-Infix "^" := BigZ.pow : bigZ_scope.
-Infix "?=" := BigZ.compare : bigZ_scope.
-Infix "=?" := BigZ.eqb (at level 70, no associativity) : bigZ_scope.
-Infix "<=?" := BigZ.leb (at level 70, no associativity) : bigZ_scope.
-Infix "<?" := BigZ.ltb (at level 70, no associativity) : bigZ_scope.
-Infix "==" := BigZ.eq (at level 70, no associativity) : bigZ_scope.
-Notation "x != y" := (~x==y) (at level 70, no associativity) : bigZ_scope.
-Infix "<" := BigZ.lt : bigZ_scope.
-Infix "<=" := BigZ.le : bigZ_scope.
-Notation "x > y" := (y < x) (only parsing) : bigZ_scope.
-Notation "x >= y" := (y <= x) (only parsing) : bigZ_scope.
-Notation "x < y < z" := (x<y /\ y<z) : bigZ_scope.
-Notation "x < y <= z" := (x<y /\ y<=z) : bigZ_scope.
-Notation "x <= y < z" := (x<=y /\ y<z) : bigZ_scope.
-Notation "x <= y <= z" := (x<=y /\ y<=z) : bigZ_scope.
-Notation "[ i ]" := (BigZ.to_Z i) : bigZ_scope.
-Infix "mod" := BigZ.modulo (at level 40, no associativity) : bigZ_scope.
-Infix "÷" := BigZ.quot (at level 40, left associativity) : bigZ_scope.
-
-(** Some additional results about [BigZ] *)
-
-Theorem spec_to_Z: forall n : bigZ,
- BigN.to_Z (BigZ.to_N n) = ((Z.sgn [n]) * [n])%Z.
-Proof.
-intros n; case n; simpl; intros p;
- generalize (BigN.spec_pos p); case (BigN.to_Z p); auto.
-intros p1 H1; case H1; auto.
-intros p1 H1; case H1; auto.
-Qed.
-
-Theorem spec_to_N n:
- ([n] = Z.sgn [n] * (BigN.to_Z (BigZ.to_N n)))%Z.
-Proof.
-case n; simpl; intros p;
- generalize (BigN.spec_pos p); case (BigN.to_Z p); auto.
-intros p1 H1; case H1; auto.
-intros p1 H1; case H1; auto.
-Qed.
-
-Theorem spec_to_Z_pos: forall n, (0 <= [n])%Z ->
- BigN.to_Z (BigZ.to_N n) = [n].
-Proof.
-intros n; case n; simpl; intros p;
- generalize (BigN.spec_pos p); case (BigN.to_Z p); auto.
-intros p1 _ H1; case H1; auto.
-intros p1 H1; case H1; auto.
-Qed.
-
-(** [BigZ] is a ring *)
-
-Lemma BigZring :
- ring_theory 0 1 BigZ.add BigZ.mul BigZ.sub BigZ.opp BigZ.eq.
-Proof.
-constructor.
-exact BigZ.add_0_l. exact BigZ.add_comm. exact BigZ.add_assoc.
-exact BigZ.mul_1_l. exact BigZ.mul_comm. exact BigZ.mul_assoc.
-exact BigZ.mul_add_distr_r.
-symmetry. apply BigZ.add_opp_r.
-exact BigZ.add_opp_diag_r.
-Qed.
-
-Lemma BigZeqb_correct : forall x y, (x =? y) = true -> x==y.
-Proof. now apply BigZ.eqb_eq. Qed.
-
-Definition BigZ_of_N n := BigZ.of_Z (Z.of_N n).
-
-Lemma BigZpower : power_theory 1 BigZ.mul BigZ.eq BigZ_of_N BigZ.pow.
-Proof.
-constructor.
-intros. unfold BigZ.eq, BigZ_of_N. rewrite BigZ.spec_pow, BigZ.spec_of_Z.
-rewrite Zpower_theory.(rpow_pow_N).
-destruct n; simpl. reflexivity.
-induction p; simpl; intros; BigZ.zify; rewrite ?IHp; auto.
-Qed.
-
-Lemma BigZdiv : div_theory BigZ.eq BigZ.add BigZ.mul (@id _)
- (fun a b => if b =? 0 then (0,a) else BigZ.div_eucl a b).
-Proof.
-constructor. unfold id. intros a b.
-BigZ.zify.
-case Z.eqb_spec.
-BigZ.zify. auto with zarith.
-intros NEQ.
-generalize (BigZ.spec_div_eucl a b).
-generalize (Z_div_mod_full [a] [b] NEQ).
-destruct BigZ.div_eucl as (q,r), Z.div_eucl as (q',r').
-intros (EQ,_). injection 1 as EQr EQq.
-BigZ.zify. rewrite EQr, EQq; auto.
-Qed.
-
-(** Detection of constants *)
-
-Ltac isBigZcst t :=
- match t with
- | BigZ.Pos ?t => isBigNcst t
- | BigZ.Neg ?t => isBigNcst t
- | BigZ.zero => constr:(true)
- | BigZ.one => constr:(true)
- | BigZ.two => constr:(true)
- | BigZ.minus_one => constr:(true)
- | _ => constr:(false)
- end.
-
-Ltac BigZcst t :=
- match isBigZcst t with
- | true => constr:(t)
- | false => constr:(NotConstant)
- end.
-
-Ltac BigZ_to_N t :=
- match t with
- | BigZ.Pos ?t => BigN_to_N t
- | BigZ.zero => constr:(0%N)
- | BigZ.one => constr:(1%N)
- | BigZ.two => constr:(2%N)
- | _ => constr:(NotConstant)
- end.
-
-(** Registration for the "ring" tactic *)
-
-Add Ring BigZr : BigZring
- (decidable BigZeqb_correct,
- constants [BigZcst],
- power_tac BigZpower [BigZ_to_N],
- div BigZdiv).
-
-Section TestRing.
-Let test : forall x y, 1 + x*y + x^2 + 1 == 1*1 + 1 + (y + 1*x)*x.
-Proof.
-intros. ring_simplify. reflexivity.
-Qed.
-Let test' : forall x y, 1 + x*y + x^2 - 1*1 - y*x + 1*(-x)*x == 0.
-Proof.
-intros. ring_simplify. reflexivity.
-Qed.
-End TestRing.
-
-(** [BigZ] also benefits from an "order" tactic *)
-
-Ltac bigZ_order := BigZ.order.
-
-Section TestOrder.
-Let test : forall x y : bigZ, x<=y -> y<=x -> x==y.
-Proof. bigZ_order. Qed.
-End TestOrder.
-
-(** We can use at least a bit of (r)omega by translating to [Z]. *)
-
-Section TestOmega.
-Let test : forall x y : bigZ, x<=y -> y<=x -> x==y.
-Proof. intros x y. BigZ.zify. omega. Qed.
-End TestOmega.
-
-(** Todo: micromega *)
diff --git a/theories/Numbers/Integer/BigZ/ZMake.v b/theories/Numbers/Integer/BigZ/ZMake.v
deleted file mode 100644
index fec6e0683..000000000
--- a/theories/Numbers/Integer/BigZ/ZMake.v
+++ /dev/null
@@ -1,759 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Require Import ZArith.
-Require Import BigNumPrelude.
-Require Import NSig.
-Require Import ZSig.
-
-Open Scope Z_scope.
-
-(** * ZMake
-
- A generic transformation from a structure of natural numbers
- [NSig.NType] to a structure of integers [ZSig.ZType].
-*)
-
-Module Make (NN:NType) <: ZType.
-
- Inductive t_ :=
- | Pos : NN.t -> t_
- | Neg : NN.t -> t_.
-
- Definition t := t_.
-
- Definition zero := Pos NN.zero.
- Definition one := Pos NN.one.
- Definition two := Pos NN.two.
- Definition minus_one := Neg NN.one.
-
- Definition of_Z x :=
- match x with
- | Zpos x => Pos (NN.of_N (Npos x))
- | Z0 => zero
- | Zneg x => Neg (NN.of_N (Npos x))
- end.
-
- Definition to_Z x :=
- match x with
- | Pos nx => NN.to_Z nx
- | Neg nx => Z.opp (NN.to_Z nx)
- end.
-
- Theorem spec_of_Z: forall x, to_Z (of_Z x) = x.
- Proof.
- intros x; case x; unfold to_Z, of_Z, zero.
- exact NN.spec_0.
- intros; rewrite NN.spec_of_N; auto.
- intros; rewrite NN.spec_of_N; auto.
- Qed.
-
- Definition eq x y := (to_Z x = to_Z y).
-
- Theorem spec_0: to_Z zero = 0.
- exact NN.spec_0.
- Qed.
-
- Theorem spec_1: to_Z one = 1.
- exact NN.spec_1.
- Qed.
-
- Theorem spec_2: to_Z two = 2.
- exact NN.spec_2.
- Qed.
-
- Theorem spec_m1: to_Z minus_one = -1.
- simpl; rewrite NN.spec_1; auto.
- Qed.
-
- Definition compare x y :=
- match x, y with
- | Pos nx, Pos ny => NN.compare nx ny
- | Pos nx, Neg ny =>
- match NN.compare nx NN.zero with
- | Gt => Gt
- | _ => NN.compare ny NN.zero
- end
- | Neg nx, Pos ny =>
- match NN.compare NN.zero nx with
- | Lt => Lt
- | _ => NN.compare NN.zero ny
- end
- | Neg nx, Neg ny => NN.compare ny nx
- end.
-
- Theorem spec_compare :
- forall x y, compare x y = Z.compare (to_Z x) (to_Z y).
- Proof.
- unfold compare, to_Z.
- destruct x as [x|x], y as [y|y];
- rewrite ?NN.spec_compare, ?NN.spec_0, ?Z.compare_opp; auto;
- assert (Hx:=NN.spec_pos x); assert (Hy:=NN.spec_pos y);
- set (X:=NN.to_Z x) in *; set (Y:=NN.to_Z y) in *; clearbody X Y.
- - destruct (Z.compare_spec X 0) as [EQ|LT|GT].
- + rewrite <- Z.opp_0 in EQ. now rewrite EQ, Z.compare_opp.
- + exfalso. omega.
- + symmetry. change (X > -Y). omega.
- - destruct (Z.compare_spec 0 X) as [EQ|LT|GT].
- + rewrite <- EQ, Z.opp_0; auto.
- + symmetry. change (-X < Y). omega.
- + exfalso. omega.
- Qed.
-
- Definition eqb x y :=
- match compare x y with
- | Eq => true
- | _ => false
- end.
-
- Theorem spec_eqb x y : eqb x y = Z.eqb (to_Z x) (to_Z y).
- Proof.
- apply Bool.eq_iff_eq_true.
- unfold eqb. rewrite Z.eqb_eq, <- Z.compare_eq_iff, spec_compare.
- split; [now destruct Z.compare | now intros ->].
- Qed.
-
- Definition lt n m := to_Z n < to_Z m.
- Definition le n m := to_Z n <= to_Z m.
-
-
- Definition ltb (x y : t) : bool :=
- match compare x y with
- | Lt => true
- | _ => false
- end.
-
- Theorem spec_ltb x y : ltb x y = Z.ltb (to_Z x) (to_Z y).
- Proof.
- apply Bool.eq_iff_eq_true.
- rewrite Z.ltb_lt. unfold Z.lt, ltb. rewrite spec_compare.
- split; [now destruct Z.compare | now intros ->].
- Qed.
-
- Definition leb (x y : t) : bool :=
- match compare x y with
- | Gt => false
- | _ => true
- end.
-
- Theorem spec_leb x y : leb x y = Z.leb (to_Z x) (to_Z y).
- Proof.
- apply Bool.eq_iff_eq_true.
- rewrite Z.leb_le. unfold Z.le, leb. rewrite spec_compare.
- now destruct Z.compare; split.
- Qed.
-
- Definition min n m := match compare n m with Gt => m | _ => n end.
- Definition max n m := match compare n m with Lt => m | _ => n end.
-
- Theorem spec_min : forall n m, to_Z (min n m) = Z.min (to_Z n) (to_Z m).
- Proof.
- unfold min, Z.min. intros. rewrite spec_compare. destruct Z.compare; auto.
- Qed.
-
- Theorem spec_max : forall n m, to_Z (max n m) = Z.max (to_Z n) (to_Z m).
- Proof.
- unfold max, Z.max. intros. rewrite spec_compare. destruct Z.compare; auto.
- Qed.
-
- Definition to_N x :=
- match x with
- | Pos nx => nx
- | Neg nx => nx
- end.
-
- Definition abs x := Pos (to_N x).
-
- Theorem spec_abs: forall x, to_Z (abs x) = Z.abs (to_Z x).
- Proof.
- intros x; case x; clear x; intros x; assert (F:=NN.spec_pos x).
- simpl; rewrite Z.abs_eq; auto.
- simpl; rewrite Z.abs_neq; simpl; auto with zarith.
- Qed.
-
- Definition opp x :=
- match x with
- | Pos nx => Neg nx
- | Neg nx => Pos nx
- end.
-
- Theorem spec_opp: forall x, to_Z (opp x) = - to_Z x.
- Proof.
- intros x; case x; simpl; auto with zarith.
- Qed.
-
- Definition succ x :=
- match x with
- | Pos n => Pos (NN.succ n)
- | Neg n =>
- match NN.compare NN.zero n with
- | Lt => Neg (NN.pred n)
- | _ => one
- end
- end.
-
- Theorem spec_succ: forall n, to_Z (succ n) = to_Z n + 1.
- Proof.
- intros x; case x; clear x; intros x.
- exact (NN.spec_succ x).
- simpl. rewrite NN.spec_compare. case Z.compare_spec; rewrite ?NN.spec_0; simpl.
- intros HH; rewrite <- HH; rewrite NN.spec_1; ring.
- intros HH; rewrite NN.spec_pred, Z.max_r; auto with zarith.
- generalize (NN.spec_pos x); auto with zarith.
- Qed.
-
- Definition add x y :=
- match x, y with
- | Pos nx, Pos ny => Pos (NN.add nx ny)
- | Pos nx, Neg ny =>
- match NN.compare nx ny with
- | Gt => Pos (NN.sub nx ny)
- | Eq => zero
- | Lt => Neg (NN.sub ny nx)
- end
- | Neg nx, Pos ny =>
- match NN.compare nx ny with
- | Gt => Neg (NN.sub nx ny)
- | Eq => zero
- | Lt => Pos (NN.sub ny nx)
- end
- | Neg nx, Neg ny => Neg (NN.add nx ny)
- end.
-
- Theorem spec_add: forall x y, to_Z (add x y) = to_Z x + to_Z y.
- Proof.
- unfold add, to_Z; intros [x | x] [y | y];
- try (rewrite NN.spec_add; auto with zarith);
- rewrite NN.spec_compare; case Z.compare_spec;
- unfold zero; rewrite ?NN.spec_0, ?NN.spec_sub; omega with *.
- Qed.
-
- Definition pred x :=
- match x with
- | Pos nx =>
- match NN.compare NN.zero nx with
- | Lt => Pos (NN.pred nx)
- | _ => minus_one
- end
- | Neg nx => Neg (NN.succ nx)
- end.
-
- Theorem spec_pred: forall x, to_Z (pred x) = to_Z x - 1.
- Proof.
- unfold pred, to_Z, minus_one; intros [x | x];
- try (rewrite NN.spec_succ; ring).
- rewrite NN.spec_compare; case Z.compare_spec;
- rewrite ?NN.spec_0, ?NN.spec_1, ?NN.spec_pred;
- generalize (NN.spec_pos x); omega with *.
- Qed.
-
- Definition sub x y :=
- match x, y with
- | Pos nx, Pos ny =>
- match NN.compare nx ny with
- | Gt => Pos (NN.sub nx ny)
- | Eq => zero
- | Lt => Neg (NN.sub ny nx)
- end
- | Pos nx, Neg ny => Pos (NN.add nx ny)
- | Neg nx, Pos ny => Neg (NN.add nx ny)
- | Neg nx, Neg ny =>
- match NN.compare nx ny with
- | Gt => Neg (NN.sub nx ny)
- | Eq => zero
- | Lt => Pos (NN.sub ny nx)
- end
- end.
-
- Theorem spec_sub: forall x y, to_Z (sub x y) = to_Z x - to_Z y.
- Proof.
- unfold sub, to_Z; intros [x | x] [y | y];
- try (rewrite NN.spec_add; auto with zarith);
- rewrite NN.spec_compare; case Z.compare_spec;
- unfold zero; rewrite ?NN.spec_0, ?NN.spec_sub; omega with *.
- Qed.
-
- Definition mul x y :=
- match x, y with
- | Pos nx, Pos ny => Pos (NN.mul nx ny)
- | Pos nx, Neg ny => Neg (NN.mul nx ny)
- | Neg nx, Pos ny => Neg (NN.mul nx ny)
- | Neg nx, Neg ny => Pos (NN.mul nx ny)
- end.
-
- Theorem spec_mul: forall x y, to_Z (mul x y) = to_Z x * to_Z y.
- Proof.
- unfold mul, to_Z; intros [x | x] [y | y]; rewrite NN.spec_mul; ring.
- Qed.
-
- Definition square x :=
- match x with
- | Pos nx => Pos (NN.square nx)
- | Neg nx => Pos (NN.square nx)
- end.
-
- Theorem spec_square: forall x, to_Z (square x) = to_Z x * to_Z x.
- Proof.
- unfold square, to_Z; intros [x | x]; rewrite NN.spec_square; ring.
- Qed.
-
- Definition pow_pos x p :=
- match x with
- | Pos nx => Pos (NN.pow_pos nx p)
- | Neg nx =>
- match p with
- | xH => x
- | xO _ => Pos (NN.pow_pos nx p)
- | xI _ => Neg (NN.pow_pos nx p)
- end
- end.
-
- Theorem spec_pow_pos: forall x n, to_Z (pow_pos x n) = to_Z x ^ Zpos n.
- Proof.
- assert (F0: forall x, (-x)^2 = x^2).
- intros x; rewrite Z.pow_2_r; ring.
- unfold pow_pos, to_Z; intros [x | x] [p | p |];
- try rewrite NN.spec_pow_pos; try ring.
- assert (F: 0 <= 2 * Zpos p).
- assert (0 <= Zpos p); auto with zarith.
- rewrite Pos2Z.inj_xI; repeat rewrite Zpower_exp; auto with zarith.
- repeat rewrite Z.pow_mul_r; auto with zarith.
- rewrite F0; ring.
- assert (F: 0 <= 2 * Zpos p).
- assert (0 <= Zpos p); auto with zarith.
- rewrite Pos2Z.inj_xO; repeat rewrite Zpower_exp; auto with zarith.
- repeat rewrite Z.pow_mul_r; auto with zarith.
- rewrite F0; ring.
- Qed.
-
- Definition pow_N x n :=
- match n with
- | N0 => one
- | Npos p => pow_pos x p
- end.
-
- Theorem spec_pow_N: forall x n, to_Z (pow_N x n) = to_Z x ^ Z.of_N n.
- Proof.
- destruct n; simpl. apply NN.spec_1.
- apply spec_pow_pos.
- Qed.
-
- Definition pow x y :=
- match to_Z y with
- | Z0 => one
- | Zpos p => pow_pos x p
- | Zneg p => zero
- end.
-
- Theorem spec_pow: forall x y, to_Z (pow x y) = to_Z x ^ to_Z y.
- Proof.
- intros. unfold pow. destruct (to_Z y); simpl.
- apply NN.spec_1.
- apply spec_pow_pos.
- apply NN.spec_0.
- Qed.
-
- Definition log2 x :=
- match x with
- | Pos nx => Pos (NN.log2 nx)
- | Neg nx => zero
- end.
-
- Theorem spec_log2: forall x, to_Z (log2 x) = Z.log2 (to_Z x).
- Proof.
- intros. destruct x as [p|p]; simpl. apply NN.spec_log2.
- rewrite NN.spec_0.
- destruct (Z_le_lt_eq_dec _ _ (NN.spec_pos p)) as [LT|EQ].
- rewrite Z.log2_nonpos; auto with zarith.
- now rewrite <- EQ.
- Qed.
-
- Definition sqrt x :=
- match x with
- | Pos nx => Pos (NN.sqrt nx)
- | Neg nx => Neg NN.zero
- end.
-
- Theorem spec_sqrt: forall x, to_Z (sqrt x) = Z.sqrt (to_Z x).
- Proof.
- destruct x as [p|p]; simpl.
- apply NN.spec_sqrt.
- rewrite NN.spec_0.
- destruct (Z_le_lt_eq_dec _ _ (NN.spec_pos p)) as [LT|EQ].
- rewrite Z.sqrt_neg; auto with zarith.
- now rewrite <- EQ.
- Qed.
-
- Definition div_eucl x y :=
- match x, y with
- | Pos nx, Pos ny =>
- let (q, r) := NN.div_eucl nx ny in
- (Pos q, Pos r)
- | Pos nx, Neg ny =>
- let (q, r) := NN.div_eucl nx ny in
- if NN.eqb NN.zero r
- then (Neg q, zero)
- else (Neg (NN.succ q), Neg (NN.sub ny r))
- | Neg nx, Pos ny =>
- let (q, r) := NN.div_eucl nx ny in
- if NN.eqb NN.zero r
- then (Neg q, zero)
- else (Neg (NN.succ q), Pos (NN.sub ny r))
- | Neg nx, Neg ny =>
- let (q, r) := NN.div_eucl nx ny in
- (Pos q, Neg r)
- end.
-
- Ltac break_nonneg x px EQx :=
- let H := fresh "H" in
- assert (H:=NN.spec_pos x);
- destruct (NN.to_Z x) as [|px|px] eqn:EQx;
- [clear H|clear H|elim H; reflexivity].
-
- Theorem spec_div_eucl: forall x y,
- let (q,r) := div_eucl x y in
- (to_Z q, to_Z r) = Z.div_eucl (to_Z x) (to_Z y).
- Proof.
- unfold div_eucl, to_Z. intros [x | x] [y | y].
- (* Pos Pos *)
- generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y); auto.
- (* Pos Neg *)
- generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y) as (q,r).
- break_nonneg x px EQx; break_nonneg y py EQy;
- try (injection 1 as Hq Hr; rewrite NN.spec_eqb, NN.spec_0, Hr;
- simpl; rewrite Hq, NN.spec_0; auto).
- change (- Zpos py) with (Zneg py).
- assert (GT : Zpos py > 0) by (compute; auto).
- generalize (Z_div_mod (Zpos px) (Zpos py) GT).
- unfold Z.div_eucl. destruct (Z.pos_div_eucl px (Zpos py)) as (q',r').
- intros (EQ,MOD). injection 1 as Hq' Hr'.
- rewrite NN.spec_eqb, NN.spec_0, Hr'.
- break_nonneg r pr EQr.
- subst; simpl. rewrite NN.spec_0; auto.
- subst. lazy iota beta delta [Z.eqb].
- rewrite NN.spec_sub, NN.spec_succ, EQy, EQr. f_equal. omega with *.
- (* Neg Pos *)
- generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y) as (q,r).
- break_nonneg x px EQx; break_nonneg y py EQy;
- try (injection 1 as Hq Hr; rewrite NN.spec_eqb, NN.spec_0, Hr;
- simpl; rewrite Hq, NN.spec_0; auto).
- change (- Zpos px) with (Zneg px).
- assert (GT : Zpos py > 0) by (compute; auto).
- generalize (Z_div_mod (Zpos px) (Zpos py) GT).
- unfold Z.div_eucl. destruct (Z.pos_div_eucl px (Zpos py)) as (q',r').
- intros (EQ,MOD). injection 1 as Hq' Hr'.
- rewrite NN.spec_eqb, NN.spec_0, Hr'.
- break_nonneg r pr EQr.
- subst; simpl. rewrite NN.spec_0; auto.
- subst. lazy iota beta delta [Z.eqb].
- rewrite NN.spec_sub, NN.spec_succ, EQy, EQr. f_equal. omega with *.
- (* Neg Neg *)
- generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y) as (q,r).
- break_nonneg x px EQx; break_nonneg y py EQy;
- try (injection 1 as -> ->; auto).
- simpl. intros <-; auto.
- Qed.
-
- Definition div x y := fst (div_eucl x y).
-
- Definition spec_div: forall x y,
- to_Z (div x y) = to_Z x / to_Z y.
- Proof.
- intros x y; generalize (spec_div_eucl x y); unfold div, Z.div.
- case div_eucl; case Z.div_eucl; simpl; auto.
- intros q r q11 r1 H; injection H; auto.
- Qed.
-
- Definition modulo x y := snd (div_eucl x y).
-
- Theorem spec_modulo:
- forall x y, to_Z (modulo x y) = to_Z x mod to_Z y.
- Proof.
- intros x y; generalize (spec_div_eucl x y); unfold modulo, Z.modulo.
- case div_eucl; case Z.div_eucl; simpl; auto.
- intros q r q11 r1 H; injection H; auto.
- Qed.
-
- Definition quot x y :=
- match x, y with
- | Pos nx, Pos ny => Pos (NN.div nx ny)
- | Pos nx, Neg ny => Neg (NN.div nx ny)
- | Neg nx, Pos ny => Neg (NN.div nx ny)
- | Neg nx, Neg ny => Pos (NN.div nx ny)
- end.
-
- Definition rem x y :=
- if eqb y zero then x
- else
- match x, y with
- | Pos nx, Pos ny => Pos (NN.modulo nx ny)
- | Pos nx, Neg ny => Pos (NN.modulo nx ny)
- | Neg nx, Pos ny => Neg (NN.modulo nx ny)
- | Neg nx, Neg ny => Neg (NN.modulo nx ny)
- end.
-
- Lemma spec_quot : forall x y, to_Z (quot x y) = (to_Z x) ÷ (to_Z y).
- Proof.
- intros [x|x] [y|y]; simpl; symmetry; rewrite NN.spec_div;
- (* Nota: we rely here on [forall a b, a ÷ 0 = b / 0] *)
- destruct (Z.eq_dec (NN.to_Z y) 0) as [EQ|NEQ];
- try (rewrite EQ; now destruct (NN.to_Z x));
- rewrite ?Z.quot_opp_r, ?Z.quot_opp_l, ?Z.opp_involutive, ?Z.opp_inj_wd;
- trivial; apply Z.quot_div_nonneg;
- generalize (NN.spec_pos x) (NN.spec_pos y); Z.order.
- Qed.
-
- Lemma spec_rem : forall x y,
- to_Z (rem x y) = Z.rem (to_Z x) (to_Z y).
- Proof.
- intros x y. unfold rem. rewrite spec_eqb, spec_0.
- case Z.eqb_spec; intros Hy.
- (* Nota: we rely here on [Z.rem a 0 = a] *)
- rewrite Hy. now destruct (to_Z x).
- destruct x as [x|x], y as [y|y]; simpl in *; symmetry;
- rewrite ?Z.eq_opp_l, ?Z.opp_0 in Hy;
- rewrite NN.spec_modulo, ?Z.rem_opp_r, ?Z.rem_opp_l, ?Z.opp_involutive,
- ?Z.opp_inj_wd;
- trivial; apply Z.rem_mod_nonneg;
- generalize (NN.spec_pos x) (NN.spec_pos y); Z.order.
- Qed.
-
- Definition gcd x y :=
- match x, y with
- | Pos nx, Pos ny => Pos (NN.gcd nx ny)
- | Pos nx, Neg ny => Pos (NN.gcd nx ny)
- | Neg nx, Pos ny => Pos (NN.gcd nx ny)
- | Neg nx, Neg ny => Pos (NN.gcd nx ny)
- end.
-
- Theorem spec_gcd: forall a b, to_Z (gcd a b) = Z.gcd (to_Z a) (to_Z b).
- Proof.
- unfold gcd, Z.gcd, to_Z; intros [x | x] [y | y]; rewrite NN.spec_gcd; unfold Z.gcd;
- auto; case NN.to_Z; simpl; auto with zarith;
- try rewrite Z.abs_opp; auto;
- case NN.to_Z; simpl; auto with zarith.
- Qed.
-
- Definition sgn x :=
- match compare zero x with
- | Lt => one
- | Eq => zero
- | Gt => minus_one
- end.
-
- Lemma spec_sgn : forall x, to_Z (sgn x) = Z.sgn (to_Z x).
- Proof.
- intros. unfold sgn. rewrite spec_compare. case Z.compare_spec.
- rewrite spec_0. intros <-; auto.
- rewrite spec_0, spec_1. symmetry. rewrite Z.sgn_pos_iff; auto.
- rewrite spec_0, spec_m1. symmetry. rewrite Z.sgn_neg_iff; auto with zarith.
- Qed.
-
- Definition even z :=
- match z with
- | Pos n => NN.even n
- | Neg n => NN.even n
- end.
-
- Definition odd z :=
- match z with
- | Pos n => NN.odd n
- | Neg n => NN.odd n
- end.
-
- Lemma spec_even : forall z, even z = Z.even (to_Z z).
- Proof.
- intros [n|n]; simpl; rewrite NN.spec_even; trivial.
- destruct (NN.to_Z n) as [|p|p]; now try destruct p.
- Qed.
-
- Lemma spec_odd : forall z, odd z = Z.odd (to_Z z).
- Proof.
- intros [n|n]; simpl; rewrite NN.spec_odd; trivial.
- destruct (NN.to_Z n) as [|p|p]; now try destruct p.
- Qed.
-
- Definition norm_pos z :=
- match z with
- | Pos _ => z
- | Neg n => if NN.eqb n NN.zero then Pos n else z
- end.
-
- Definition testbit a n :=
- match norm_pos n, norm_pos a with
- | Pos p, Pos a => NN.testbit a p
- | Pos p, Neg a => negb (NN.testbit (NN.pred a) p)
- | Neg p, _ => false
- end.
-
- Definition shiftl a n :=
- match norm_pos a, n with
- | Pos a, Pos n => Pos (NN.shiftl a n)
- | Pos a, Neg n => Pos (NN.shiftr a n)
- | Neg a, Pos n => Neg (NN.shiftl a n)
- | Neg a, Neg n => Neg (NN.succ (NN.shiftr (NN.pred a) n))
- end.
-
- Definition shiftr a n := shiftl a (opp n).
-
- Definition lor a b :=
- match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (NN.lor a b)
- | Neg a, Pos b => Neg (NN.succ (NN.ldiff (NN.pred a) b))
- | Pos a, Neg b => Neg (NN.succ (NN.ldiff (NN.pred b) a))
- | Neg a, Neg b => Neg (NN.succ (NN.land (NN.pred a) (NN.pred b)))
- end.
-
- Definition land a b :=
- match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (NN.land a b)
- | Neg a, Pos b => Pos (NN.ldiff b (NN.pred a))
- | Pos a, Neg b => Pos (NN.ldiff a (NN.pred b))
- | Neg a, Neg b => Neg (NN.succ (NN.lor (NN.pred a) (NN.pred b)))
- end.
-
- Definition ldiff a b :=
- match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (NN.ldiff a b)
- | Neg a, Pos b => Neg (NN.succ (NN.lor (NN.pred a) b))
- | Pos a, Neg b => Pos (NN.land a (NN.pred b))
- | Neg a, Neg b => Pos (NN.ldiff (NN.pred b) (NN.pred a))
- end.
-
- Definition lxor a b :=
- match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (NN.lxor a b)
- | Neg a, Pos b => Neg (NN.succ (NN.lxor (NN.pred a) b))
- | Pos a, Neg b => Neg (NN.succ (NN.lxor a (NN.pred b)))
- | Neg a, Neg b => Pos (NN.lxor (NN.pred a) (NN.pred b))
- end.
-
- Definition div2 x := shiftr x one.
-
- Lemma Zlnot_alt1 : forall x, -(x+1) = Z.lnot x.
- Proof.
- unfold Z.lnot, Z.pred; auto with zarith.
- Qed.
-
- Lemma Zlnot_alt2 : forall x, Z.lnot (x-1) = -x.
- Proof.
- unfold Z.lnot, Z.pred; auto with zarith.
- Qed.
-
- Lemma Zlnot_alt3 : forall x, Z.lnot (-x) = x-1.
- Proof.
- unfold Z.lnot, Z.pred; auto with zarith.
- Qed.
-
- Lemma spec_norm_pos : forall x, to_Z (norm_pos x) = to_Z x.
- Proof.
- intros [x|x]; simpl; trivial.
- rewrite NN.spec_eqb, NN.spec_0.
- case Z.eqb_spec; simpl; auto with zarith.
- Qed.
-
- Lemma spec_norm_pos_pos : forall x y, norm_pos x = Neg y ->
- 0 < NN.to_Z y.
- Proof.
- intros [x|x] y; simpl; try easy.
- rewrite NN.spec_eqb, NN.spec_0.
- case Z.eqb_spec; simpl; try easy.
- inversion 2. subst. generalize (NN.spec_pos y); auto with zarith.
- Qed.
-
- Ltac destr_norm_pos x :=
- rewrite <- (spec_norm_pos x);
- let H := fresh in
- let x' := fresh x in
- assert (H := spec_norm_pos_pos x);
- destruct (norm_pos x) as [x'|x'];
- specialize (H x' (eq_refl _)) || clear H.
-
- Lemma spec_testbit: forall x p, testbit x p = Z.testbit (to_Z x) (to_Z p).
- Proof.
- intros x p. unfold testbit.
- destr_norm_pos p; simpl. destr_norm_pos x; simpl.
- apply NN.spec_testbit.
- rewrite NN.spec_testbit, NN.spec_pred, Z.max_r by auto with zarith.
- symmetry. apply Z.bits_opp. apply NN.spec_pos.
- symmetry. apply Z.testbit_neg_r; auto with zarith.
- Qed.
-
- Lemma spec_shiftl: forall x p, to_Z (shiftl x p) = Z.shiftl (to_Z x) (to_Z p).
- Proof.
- intros x p. unfold shiftl.
- destr_norm_pos x; destruct p as [p|p]; simpl;
- assert (Hp := NN.spec_pos p).
- apply NN.spec_shiftl.
- rewrite Z.shiftl_opp_r. apply NN.spec_shiftr.
- rewrite !NN.spec_shiftl.
- rewrite !Z.shiftl_mul_pow2 by apply NN.spec_pos.
- symmetry. apply Z.mul_opp_l.
- rewrite Z.shiftl_opp_r, NN.spec_succ, NN.spec_shiftr, NN.spec_pred, Z.max_r
- by auto with zarith.
- now rewrite Zlnot_alt1, Z.lnot_shiftr, Zlnot_alt2.
- Qed.
-
- Lemma spec_shiftr: forall x p, to_Z (shiftr x p) = Z.shiftr (to_Z x) (to_Z p).
- Proof.
- intros. unfold shiftr. rewrite spec_shiftl, spec_opp.
- apply Z.shiftl_opp_r.
- Qed.
-
- Lemma spec_land: forall x y, to_Z (land x y) = Z.land (to_Z x) (to_Z y).
- Proof.
- intros x y. unfold land.
- destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?NN.spec_succ, ?NN.spec_land, ?NN.spec_ldiff, ?NN.spec_lor,
- ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1; auto with zarith.
- now rewrite Z.ldiff_land, Zlnot_alt2.
- now rewrite Z.ldiff_land, Z.land_comm, Zlnot_alt2.
- now rewrite Z.lnot_lor, !Zlnot_alt2.
- Qed.
-
- Lemma spec_lor: forall x y, to_Z (lor x y) = Z.lor (to_Z x) (to_Z y).
- Proof.
- intros x y. unfold lor.
- destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?NN.spec_succ, ?NN.spec_land, ?NN.spec_ldiff, ?NN.spec_lor,
- ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1; auto with zarith.
- now rewrite Z.lnot_ldiff, Z.lor_comm, Zlnot_alt2.
- now rewrite Z.lnot_ldiff, Zlnot_alt2.
- now rewrite Z.lnot_land, !Zlnot_alt2.
- Qed.
-
- Lemma spec_ldiff: forall x y, to_Z (ldiff x y) = Z.ldiff (to_Z x) (to_Z y).
- Proof.
- intros x y. unfold ldiff.
- destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?NN.spec_succ, ?NN.spec_land, ?NN.spec_ldiff, ?NN.spec_lor,
- ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1; auto with zarith.
- now rewrite Z.ldiff_land, Zlnot_alt3.
- now rewrite Z.lnot_lor, Z.ldiff_land, <- Zlnot_alt2.
- now rewrite 2 Z.ldiff_land, Zlnot_alt2, Z.land_comm, Zlnot_alt3.
- Qed.
-
- Lemma spec_lxor: forall x y, to_Z (lxor x y) = Z.lxor (to_Z x) (to_Z y).
- Proof.
- intros x y. unfold lxor.
- destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?NN.spec_succ, ?NN.spec_lxor, ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1;
- auto with zarith.
- now rewrite !Z.lnot_lxor_r, Zlnot_alt2.
- now rewrite !Z.lnot_lxor_l, Zlnot_alt2.
- now rewrite <- Z.lxor_lnot_lnot, !Zlnot_alt2.
- Qed.
-
- Lemma spec_div2: forall x, to_Z (div2 x) = Z.div2 (to_Z x).
- Proof.
- intros x. unfold div2. now rewrite spec_shiftr, Z.div2_spec, spec_1.
- Qed.
-
-End Make.
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSig.v b/theories/Numbers/Integer/SpecViaZ/ZSig.v
deleted file mode 100644
index a360327a4..000000000
--- a/theories/Numbers/Integer/SpecViaZ/ZSig.v
+++ /dev/null
@@ -1,135 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Require Import BinInt.
-
-Open Scope Z_scope.
-
-(** * ZSig *)
-
-(** Interface of a rich structure about integers.
- Specifications are written via translation to Z.
-*)
-
-Module Type ZType.
-
- Parameter t : Type.
-
- Parameter to_Z : t -> Z.
- Local Notation "[ x ]" := (to_Z x).
-
- Definition eq x y := [x] = [y].
- Definition lt x y := [x] < [y].
- Definition le x y := [x] <= [y].
-
- Parameter of_Z : Z -> t.
- Parameter spec_of_Z: forall x, to_Z (of_Z x) = x.
-
- Parameter compare : t -> t -> comparison.
- Parameter eqb : t -> t -> bool.
- Parameter ltb : t -> t -> bool.
- Parameter leb : t -> t -> bool.
- Parameter min : t -> t -> t.
- Parameter max : t -> t -> t.
- Parameter zero : t.
- Parameter one : t.
- Parameter two : t.
- Parameter minus_one : t.
- Parameter succ : t -> t.
- Parameter add : t -> t -> t.
- Parameter pred : t -> t.
- Parameter sub : t -> t -> t.
- Parameter opp : t -> t.
- Parameter mul : t -> t -> t.
- Parameter square : t -> t.
- Parameter pow_pos : t -> positive -> t.
- Parameter pow_N : t -> N -> t.
- Parameter pow : t -> t -> t.
- Parameter sqrt : t -> t.
- Parameter log2 : t -> t.
- Parameter div_eucl : t -> t -> t * t.
- Parameter div : t -> t -> t.
- Parameter modulo : t -> t -> t.
- Parameter quot : t -> t -> t.
- Parameter rem : t -> t -> t.
- Parameter gcd : t -> t -> t.
- Parameter sgn : t -> t.
- Parameter abs : t -> t.
- Parameter even : t -> bool.
- Parameter odd : t -> bool.
- Parameter testbit : t -> t -> bool.
- Parameter shiftr : t -> t -> t.
- Parameter shiftl : t -> t -> t.
- Parameter land : t -> t -> t.
- Parameter lor : t -> t -> t.
- Parameter ldiff : t -> t -> t.
- Parameter lxor : t -> t -> t.
- Parameter div2 : t -> t.
-
- Parameter spec_compare: forall x y, compare x y = ([x] ?= [y]).
- Parameter spec_eqb : forall x y, eqb x y = ([x] =? [y]).
- Parameter spec_ltb : forall x y, ltb x y = ([x] <? [y]).
- Parameter spec_leb : forall x y, leb x y = ([x] <=? [y]).
- Parameter spec_min : forall x y, [min x y] = Z.min [x] [y].
- Parameter spec_max : forall x y, [max x y] = Z.max [x] [y].
- Parameter spec_0: [zero] = 0.
- Parameter spec_1: [one] = 1.
- Parameter spec_2: [two] = 2.
- Parameter spec_m1: [minus_one] = -1.
- Parameter spec_succ: forall n, [succ n] = [n] + 1.
- Parameter spec_add: forall x y, [add x y] = [x] + [y].
- Parameter spec_pred: forall x, [pred x] = [x] - 1.
- Parameter spec_sub: forall x y, [sub x y] = [x] - [y].
- Parameter spec_opp: forall x, [opp x] = - [x].
- Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
- Parameter spec_square: forall x, [square x] = [x] * [x].
- Parameter spec_pow_pos: forall x n, [pow_pos x n] = [x] ^ Zpos n.
- Parameter spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z.of_N n.
- Parameter spec_pow: forall x n, [pow x n] = [x] ^ [n].
- Parameter spec_sqrt: forall x, [sqrt x] = Z.sqrt [x].
- Parameter spec_log2: forall x, [log2 x] = Z.log2 [x].
- Parameter spec_div_eucl: forall x y,
- let (q,r) := div_eucl x y in ([q], [r]) = Z.div_eucl [x] [y].
- Parameter spec_div: forall x y, [div x y] = [x] / [y].
- Parameter spec_modulo: forall x y, [modulo x y] = [x] mod [y].
- Parameter spec_quot: forall x y, [quot x y] = [x] ÷ [y].
- Parameter spec_rem: forall x y, [rem x y] = Z.rem [x] [y].
- Parameter spec_gcd: forall a b, [gcd a b] = Z.gcd [a] [b].
- Parameter spec_sgn : forall x, [sgn x] = Z.sgn [x].
- Parameter spec_abs : forall x, [abs x] = Z.abs [x].
- Parameter spec_even : forall x, even x = Z.even [x].
- Parameter spec_odd : forall x, odd x = Z.odd [x].
- Parameter spec_testbit: forall x p, testbit x p = Z.testbit [x] [p].
- Parameter spec_shiftr: forall x p, [shiftr x p] = Z.shiftr [x] [p].
- Parameter spec_shiftl: forall x p, [shiftl x p] = Z.shiftl [x] [p].
- Parameter spec_land: forall x y, [land x y] = Z.land [x] [y].
- Parameter spec_lor: forall x y, [lor x y] = Z.lor [x] [y].
- Parameter spec_ldiff: forall x y, [ldiff x y] = Z.ldiff [x] [y].
- Parameter spec_lxor: forall x y, [lxor x y] = Z.lxor [x] [y].
- Parameter spec_div2: forall x, [div2 x] = Z.div2 [x].
-
-End ZType.
-
-Module Type ZType_Notation (Import Z:ZType).
- Notation "[ x ]" := (to_Z x).
- Infix "==" := eq (at level 70).
- Notation "0" := zero.
- Notation "1" := one.
- Notation "2" := two.
- Infix "+" := add.
- Infix "-" := sub.
- Infix "*" := mul.
- Infix "^" := pow.
- Notation "- x" := (opp x).
- Infix "<=" := le.
- Infix "<" := lt.
-End ZType_Notation.
-
-Module Type ZType' := ZType <+ ZType_Notation.
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
deleted file mode 100644
index 32410d1d0..000000000
--- a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
+++ /dev/null
@@ -1,527 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-Require Import Bool ZArith OrdersFacts Nnat ZAxioms ZSig.
-
-(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig] *)
-
-Module ZTypeIsZAxioms (Import ZZ : ZType').
-
-Hint Rewrite
- spec_0 spec_1 spec_2 spec_add spec_sub spec_pred spec_succ
- spec_mul spec_opp spec_of_Z spec_div spec_modulo spec_square spec_sqrt
- spec_compare spec_eqb spec_ltb spec_leb spec_max spec_min
- spec_abs spec_sgn spec_pow spec_log2 spec_even spec_odd spec_gcd
- spec_quot spec_rem spec_testbit spec_shiftl spec_shiftr
- spec_land spec_lor spec_ldiff spec_lxor spec_div2
- : zsimpl.
-
-Ltac zsimpl := autorewrite with zsimpl.
-Ltac zcongruence := repeat red; intros; zsimpl; congruence.
-Ltac zify := unfold eq, lt, le in *; zsimpl.
-
-Instance eq_equiv : Equivalence eq.
-Proof. unfold eq. firstorder. Qed.
-
-Local Obligation Tactic := zcongruence.
-
-Program Instance succ_wd : Proper (eq ==> eq) succ.
-Program Instance pred_wd : Proper (eq ==> eq) pred.
-Program Instance add_wd : Proper (eq ==> eq ==> eq) add.
-Program Instance sub_wd : Proper (eq ==> eq ==> eq) sub.
-Program Instance mul_wd : Proper (eq ==> eq ==> eq) mul.
-
-Theorem pred_succ : forall n, pred (succ n) == n.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem one_succ : 1 == succ 0.
-Proof.
-now zify.
-Qed.
-
-Theorem two_succ : 2 == succ 1.
-Proof.
-now zify.
-Qed.
-
-Section Induction.
-
-Variable A : ZZ.t -> Prop.
-Hypothesis A_wd : Proper (eq==>iff) A.
-Hypothesis A0 : A 0.
-Hypothesis AS : forall n, A n <-> A (succ n).
-
-Let B (z : Z) := A (of_Z z).
-
-Lemma B0 : B 0.
-Proof.
-unfold B; simpl.
-rewrite <- (A_wd 0); auto.
-zify. auto.
-Qed.
-
-Lemma BS : forall z : Z, B z -> B (z + 1).
-Proof.
-intros z H.
-unfold B in *. apply -> AS in H.
-setoid_replace (of_Z (z + 1)) with (succ (of_Z z)); auto.
-zify. auto.
-Qed.
-
-Lemma BP : forall z : Z, B z -> B (z - 1).
-Proof.
-intros z H.
-unfold B in *. rewrite AS.
-setoid_replace (succ (of_Z (z - 1))) with (of_Z z); auto.
-zify. auto with zarith.
-Qed.
-
-Lemma B_holds : forall z : Z, B z.
-Proof.
-intros; destruct (Z_lt_le_dec 0 z).
-apply natlike_ind; auto with zarith.
-apply B0.
-intros; apply BS; auto.
-replace z with (-(-z))%Z in * by (auto with zarith).
-remember (-z)%Z as z'.
-pattern z'; apply natlike_ind.
-apply B0.
-intros; rewrite Z.opp_succ; unfold Z.pred; apply BP; auto.
-subst z'; auto with zarith.
-Qed.
-
-Theorem bi_induction : forall n, A n.
-Proof.
-intro n. setoid_replace n with (of_Z (to_Z n)).
-apply B_holds.
-zify. auto.
-Qed.
-
-End Induction.
-
-Theorem add_0_l : forall n, 0 + n == n.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem add_succ_l : forall n m, (succ n) + m == succ (n + m).
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem sub_0_r : forall n, n - 0 == n.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem sub_succ_r : forall n m, n - (succ m) == pred (n - m).
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem mul_0_l : forall n, 0 * n == 0.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem mul_succ_l : forall n m, (succ n) * m == n * m + m.
-Proof.
-intros. zify. ring.
-Qed.
-
-(** Order *)
-
-Lemma eqb_eq x y : eqb x y = true <-> x == y.
-Proof.
- zify. apply Z.eqb_eq.
-Qed.
-
-Lemma leb_le x y : leb x y = true <-> x <= y.
-Proof.
- zify. apply Z.leb_le.
-Qed.
-
-Lemma ltb_lt x y : ltb x y = true <-> x < y.
-Proof.
- zify. apply Z.ltb_lt.
-Qed.
-
-Lemma compare_eq_iff n m : compare n m = Eq <-> n == m.
-Proof.
- intros. zify. apply Z.compare_eq_iff.
-Qed.
-
-Lemma compare_lt_iff n m : compare n m = Lt <-> n < m.
-Proof.
- intros. zify. reflexivity.
-Qed.
-
-Lemma compare_le_iff n m : compare n m <> Gt <-> n <= m.
-Proof.
- intros. zify. reflexivity.
-Qed.
-
-Lemma compare_antisym n m : compare m n = CompOpp (compare n m).
-Proof.
- intros. zify. apply Z.compare_antisym.
-Qed.
-
-Include BoolOrderFacts ZZ ZZ ZZ [no inline].
-
-Instance compare_wd : Proper (eq ==> eq ==> Logic.eq) compare.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance eqb_wd : Proper (eq ==> eq ==> Logic.eq) eqb.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance ltb_wd : Proper (eq ==> eq ==> Logic.eq) ltb.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance leb_wd : Proper (eq ==> eq ==> Logic.eq) leb.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
-Proof.
-intros x x' Hx y y' Hy; unfold lt; rewrite Hx, Hy; intuition.
-Qed.
-
-Theorem lt_succ_r : forall n m, n < (succ m) <-> n <= m.
-Proof.
-intros. zify. omega.
-Qed.
-
-Theorem min_l : forall n m, n <= m -> min n m == n.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-Theorem min_r : forall n m, m <= n -> min n m == m.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-Theorem max_l : forall n m, m <= n -> max n m == n.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-Theorem max_r : forall n m, n <= m -> max n m == m.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-(** Part specific to integers, not natural numbers *)
-
-Theorem succ_pred : forall n, succ (pred n) == n.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-(** Opp *)
-
-Program Instance opp_wd : Proper (eq ==> eq) opp.
-
-Theorem opp_0 : - 0 == 0.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem opp_succ : forall n, - (succ n) == pred (- n).
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-(** Abs / Sgn *)
-
-Theorem abs_eq : forall n, 0 <= n -> abs n == n.
-Proof.
-intros n. zify. omega with *.
-Qed.
-
-Theorem abs_neq : forall n, n <= 0 -> abs n == -n.
-Proof.
-intros n. zify. omega with *.
-Qed.
-
-Theorem sgn_null : forall n, n==0 -> sgn n == 0.
-Proof.
-intros n. zify. omega with *.
-Qed.
-
-Theorem sgn_pos : forall n, 0<n -> sgn n == 1.
-Proof.
-intros n. zify. omega with *.
-Qed.
-
-Theorem sgn_neg : forall n, n<0 -> sgn n == opp 1.
-Proof.
-intros n. zify. omega with *.
-Qed.
-
-(** Power *)
-
-Program Instance pow_wd : Proper (eq==>eq==>eq) pow.
-
-Lemma pow_0_r : forall a, a^0 == 1.
-Proof.
- intros. now zify.
-Qed.
-
-Lemma pow_succ_r : forall a b, 0<=b -> a^(succ b) == a * a^b.
-Proof.
- intros a b. zify. intros. now rewrite Z.add_1_r, Z.pow_succ_r.
-Qed.
-
-Lemma pow_neg_r : forall a b, b<0 -> a^b == 0.
-Proof.
- intros a b. zify. intros Hb.
- destruct [b]; reflexivity || discriminate.
-Qed.
-
-Lemma pow_pow_N : forall a b, 0<=b -> a^b == pow_N a (Z.to_N (to_Z b)).
-Proof.
- intros a b. zify. intros Hb. now rewrite spec_pow_N, Z2N.id.
-Qed.
-
-Lemma pow_pos_N : forall a p, pow_pos a p == pow_N a (Npos p).
-Proof.
- intros a b. red. now rewrite spec_pow_N, spec_pow_pos.
-Qed.
-
-(** Square *)
-
-Lemma square_spec n : square n == n * n.
-Proof.
- now zify.
-Qed.
-
-(** Sqrt *)
-
-Lemma sqrt_spec : forall n, 0<=n ->
- (sqrt n)*(sqrt n) <= n /\ n < (succ (sqrt n))*(succ (sqrt n)).
-Proof.
- intros n. zify. apply Z.sqrt_spec.
-Qed.
-
-Lemma sqrt_neg : forall n, n<0 -> sqrt n == 0.
-Proof.
- intros n. zify. apply Z.sqrt_neg.
-Qed.
-
-(** Log2 *)
-
-Lemma log2_spec : forall n, 0<n ->
- 2^(log2 n) <= n /\ n < 2^(succ (log2 n)).
-Proof.
- intros n. zify. apply Z.log2_spec.
-Qed.
-
-Lemma log2_nonpos : forall n, n<=0 -> log2 n == 0.
-Proof.
- intros n. zify. apply Z.log2_nonpos.
-Qed.
-
-(** Even / Odd *)
-
-Definition Even n := exists m, n == 2*m.
-Definition Odd n := exists m, n == 2*m+1.
-
-Lemma even_spec n : even n = true <-> Even n.
-Proof.
- unfold Even. zify. rewrite Z.even_spec.
- split; intros (m,Hm).
- - exists (of_Z m). now zify.
- - exists [m]. revert Hm. now zify.
-Qed.
-
-Lemma odd_spec n : odd n = true <-> Odd n.
-Proof.
- unfold Odd. zify. rewrite Z.odd_spec.
- split; intros (m,Hm).
- - exists (of_Z m). now zify.
- - exists [m]. revert Hm. now zify.
-Qed.
-
-(** Div / Mod *)
-
-Program Instance div_wd : Proper (eq==>eq==>eq) div.
-Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
-
-Theorem div_mod : forall a b, ~b==0 -> a == b*(div a b) + (modulo a b).
-Proof.
-intros a b. zify. intros. apply Z.div_mod; auto.
-Qed.
-
-Theorem mod_pos_bound :
- forall a b, 0 < b -> 0 <= modulo a b /\ modulo a b < b.
-Proof.
-intros a b. zify. intros. apply Z_mod_lt; auto with zarith.
-Qed.
-
-Theorem mod_neg_bound :
- forall a b, b < 0 -> b < modulo a b /\ modulo a b <= 0.
-Proof.
-intros a b. zify. intros. apply Z_mod_neg; auto with zarith.
-Qed.
-
-Definition mod_bound_pos :
- forall a b, 0<=a -> 0<b -> 0 <= modulo a b /\ modulo a b < b :=
- fun a b _ H => mod_pos_bound a b H.
-
-(** Quot / Rem *)
-
-Program Instance quot_wd : Proper (eq==>eq==>eq) quot.
-Program Instance rem_wd : Proper (eq==>eq==>eq) rem.
-
-Theorem quot_rem : forall a b, ~b==0 -> a == b*(quot a b) + rem a b.
-Proof.
-intros a b. zify. apply Z.quot_rem.
-Qed.
-
-Theorem rem_bound_pos :
- forall a b, 0<=a -> 0<b -> 0 <= rem a b /\ rem a b < b.
-Proof.
-intros a b. zify. apply Z.rem_bound_pos.
-Qed.
-
-Theorem rem_opp_l : forall a b, ~b==0 -> rem (-a) b == -(rem a b).
-Proof.
-intros a b. zify. apply Z.rem_opp_l.
-Qed.
-
-Theorem rem_opp_r : forall a b, ~b==0 -> rem a (-b) == rem a b.
-Proof.
-intros a b. zify. apply Z.rem_opp_r.
-Qed.
-
-(** Gcd *)
-
-Definition divide n m := exists p, m == p*n.
-Local Notation "( x | y )" := (divide x y) (at level 0).
-
-Lemma spec_divide : forall n m, (n|m) <-> Z.divide [n] [m].
-Proof.
- intros n m. split.
- - intros (p,H). exists [p]. revert H; now zify.
- - intros (z,H). exists (of_Z z). now zify.
-Qed.
-
-Lemma gcd_divide_l : forall n m, (gcd n m | n).
-Proof.
- intros n m. apply spec_divide. zify. apply Z.gcd_divide_l.
-Qed.
-
-Lemma gcd_divide_r : forall n m, (gcd n m | m).
-Proof.
- intros n m. apply spec_divide. zify. apply Z.gcd_divide_r.
-Qed.
-
-Lemma gcd_greatest : forall n m p, (p|n) -> (p|m) -> (p|gcd n m).
-Proof.
- intros n m p. rewrite !spec_divide. zify. apply Z.gcd_greatest.
-Qed.
-
-Lemma gcd_nonneg : forall n m, 0 <= gcd n m.
-Proof.
- intros. zify. apply Z.gcd_nonneg.
-Qed.
-
-(** Bitwise operations *)
-
-Program Instance testbit_wd : Proper (eq==>eq==>Logic.eq) testbit.
-
-Lemma testbit_odd_0 : forall a, testbit (2*a+1) 0 = true.
-Proof.
- intros. zify. apply Z.testbit_odd_0.
-Qed.
-
-Lemma testbit_even_0 : forall a, testbit (2*a) 0 = false.
-Proof.
- intros. zify. apply Z.testbit_even_0.
-Qed.
-
-Lemma testbit_odd_succ : forall a n, 0<=n ->
- testbit (2*a+1) (succ n) = testbit a n.
-Proof.
- intros a n. zify. apply Z.testbit_odd_succ.
-Qed.
-
-Lemma testbit_even_succ : forall a n, 0<=n ->
- testbit (2*a) (succ n) = testbit a n.
-Proof.
- intros a n. zify. apply Z.testbit_even_succ.
-Qed.
-
-Lemma testbit_neg_r : forall a n, n<0 -> testbit a n = false.
-Proof.
- intros a n. zify. apply Z.testbit_neg_r.
-Qed.
-
-Lemma shiftr_spec : forall a n m, 0<=m ->
- testbit (shiftr a n) m = testbit a (m+n).
-Proof.
- intros a n m. zify. apply Z.shiftr_spec.
-Qed.
-
-Lemma shiftl_spec_high : forall a n m, 0<=m -> n<=m ->
- testbit (shiftl a n) m = testbit a (m-n).
-Proof.
- intros a n m. zify. intros Hn H.
- now apply Z.shiftl_spec_high.
-Qed.
-
-Lemma shiftl_spec_low : forall a n m, m<n ->
- testbit (shiftl a n) m = false.
-Proof.
- intros a n m. zify. intros H. now apply Z.shiftl_spec_low.
-Qed.
-
-Lemma land_spec : forall a b n,
- testbit (land a b) n = testbit a n && testbit b n.
-Proof.
- intros a n m. zify. now apply Z.land_spec.
-Qed.
-
-Lemma lor_spec : forall a b n,
- testbit (lor a b) n = testbit a n || testbit b n.
-Proof.
- intros a n m. zify. now apply Z.lor_spec.
-Qed.
-
-Lemma ldiff_spec : forall a b n,
- testbit (ldiff a b) n = testbit a n && negb (testbit b n).
-Proof.
- intros a n m. zify. now apply Z.ldiff_spec.
-Qed.
-
-Lemma lxor_spec : forall a b n,
- testbit (lxor a b) n = xorb (testbit a n) (testbit b n).
-Proof.
- intros a n m. zify. now apply Z.lxor_spec.
-Qed.
-
-Lemma div2_spec : forall a, div2 a == shiftr a 1.
-Proof.
- intros a. zify. now apply Z.div2_spec.
-Qed.
-
-End ZTypeIsZAxioms.
-
-Module ZType_ZAxioms (ZZ : ZType)
- <: ZAxiomsSig <: OrderFunctions ZZ <: HasMinMax ZZ
- := ZZ <+ ZTypeIsZAxioms.
diff --git a/theories/Numbers/Natural/BigN/BigN.v b/theories/Numbers/Natural/BigN/BigN.v
deleted file mode 100644
index e8ff516f3..000000000
--- a/theories/Numbers/Natural/BigN/BigN.v
+++ /dev/null
@@ -1,198 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** * Efficient arbitrary large natural numbers in base 2^31 *)
-
-(** Initial Author: Arnaud Spiwack *)
-
-Require Export Int31.
-Require Import CyclicAxioms Cyclic31 Ring31 NSig NSigNAxioms NMake
- NProperties GenericMinMax.
-
-(** The following [BigN] module regroups both the operations and
- all the abstract properties:
-
- - [NMake.Make Int31Cyclic] provides the operations and basic specs
- w.r.t. ZArith
- - [NTypeIsNAxioms] shows (mainly) that these operations implement
- the interface [NAxioms]
- - [NProp] adds all generic properties derived from [NAxioms]
- - [MinMax*Properties] provides properties of [min] and [max].
-
-*)
-
-Delimit Scope bigN_scope with bigN.
-
-Module BigN <: NType <: OrderedTypeFull <: TotalOrder :=
- NMake.Make Int31Cyclic
- <+ NTypeIsNAxioms
- <+ NBasicProp [no inline] <+ NExtraProp [no inline]
- <+ HasEqBool2Dec [no inline]
- <+ MinMaxLogicalProperties [no inline]
- <+ MinMaxDecProperties [no inline].
-
-(** Notations about [BigN] *)
-
-Local Open Scope bigN_scope.
-
-Notation bigN := BigN.t.
-Bind Scope bigN_scope with bigN BigN.t BigN.t'.
-Arguments BigN.N0 _%int31.
-Local Notation "0" := BigN.zero : bigN_scope. (* temporary notation *)
-Local Notation "1" := BigN.one : bigN_scope. (* temporary notation *)
-Local Notation "2" := BigN.two : bigN_scope. (* temporary notation *)
-Infix "+" := BigN.add : bigN_scope.
-Infix "-" := BigN.sub : bigN_scope.
-Infix "*" := BigN.mul : bigN_scope.
-Infix "/" := BigN.div : bigN_scope.
-Infix "^" := BigN.pow : bigN_scope.
-Infix "?=" := BigN.compare : bigN_scope.
-Infix "=?" := BigN.eqb (at level 70, no associativity) : bigN_scope.
-Infix "<=?" := BigN.leb (at level 70, no associativity) : bigN_scope.
-Infix "<?" := BigN.ltb (at level 70, no associativity) : bigN_scope.
-Infix "==" := BigN.eq (at level 70, no associativity) : bigN_scope.
-Notation "x != y" := (~x==y) (at level 70, no associativity) : bigN_scope.
-Infix "<" := BigN.lt : bigN_scope.
-Infix "<=" := BigN.le : bigN_scope.
-Notation "x > y" := (y < x) (only parsing) : bigN_scope.
-Notation "x >= y" := (y <= x) (only parsing) : bigN_scope.
-Notation "x < y < z" := (x<y /\ y<z) : bigN_scope.
-Notation "x < y <= z" := (x<y /\ y<=z) : bigN_scope.
-Notation "x <= y < z" := (x<=y /\ y<z) : bigN_scope.
-Notation "x <= y <= z" := (x<=y /\ y<=z) : bigN_scope.
-Notation "[ i ]" := (BigN.to_Z i) : bigN_scope.
-Infix "mod" := BigN.modulo (at level 40, no associativity) : bigN_scope.
-
-(** Example of reasoning about [BigN] *)
-
-Theorem succ_pred: forall q : bigN,
- 0 < q -> BigN.succ (BigN.pred q) == q.
-Proof.
-intros; apply BigN.succ_pred.
-intro H'; rewrite H' in H; discriminate.
-Qed.
-
-(** [BigN] is a semi-ring *)
-
-Lemma BigNring : semi_ring_theory 0 1 BigN.add BigN.mul BigN.eq.
-Proof.
-constructor.
-exact BigN.add_0_l. exact BigN.add_comm. exact BigN.add_assoc.
-exact BigN.mul_1_l. exact BigN.mul_0_l. exact BigN.mul_comm.
-exact BigN.mul_assoc. exact BigN.mul_add_distr_r.
-Qed.
-
-Lemma BigNeqb_correct : forall x y, (x =? y) = true -> x==y.
-Proof. now apply BigN.eqb_eq. Qed.
-
-Lemma BigNpower : power_theory 1 BigN.mul BigN.eq BigN.of_N BigN.pow.
-Proof.
-constructor.
-intros. red. rewrite BigN.spec_pow, BigN.spec_of_N.
-rewrite Zpower_theory.(rpow_pow_N).
-destruct n; simpl. reflexivity.
-induction p; simpl; intros; BigN.zify; rewrite ?IHp; auto.
-Qed.
-
-Lemma BigNdiv : div_theory BigN.eq BigN.add BigN.mul (@id _)
- (fun a b => if b =? 0 then (0,a) else BigN.div_eucl a b).
-Proof.
-constructor. unfold id. intros a b.
-BigN.zify.
-case Z.eqb_spec.
-BigN.zify. auto with zarith.
-intros NEQ.
-generalize (BigN.spec_div_eucl a b).
-generalize (Z_div_mod_full [a] [b] NEQ).
-destruct BigN.div_eucl as (q,r), Z.div_eucl as (q',r').
-intros (EQ,_). injection 1 as EQr EQq.
-BigN.zify. rewrite EQr, EQq; auto.
-Qed.
-
-
-(** Detection of constants *)
-
-Ltac isStaticWordCst t :=
- match t with
- | W0 => constr:(true)
- | WW ?t1 ?t2 =>
- match isStaticWordCst t1 with
- | false => constr:(false)
- | true => isStaticWordCst t2
- end
- | _ => isInt31cst t
- end.
-
-Ltac isBigNcst t :=
- match t with
- | BigN.N0 ?t => isStaticWordCst t
- | BigN.N1 ?t => isStaticWordCst t
- | BigN.N2 ?t => isStaticWordCst t
- | BigN.N3 ?t => isStaticWordCst t
- | BigN.N4 ?t => isStaticWordCst t
- | BigN.N5 ?t => isStaticWordCst t
- | BigN.N6 ?t => isStaticWordCst t
- | BigN.Nn ?n ?t => match isnatcst n with
- | true => isStaticWordCst t
- | false => constr:(false)
- end
- | BigN.zero => constr:(true)
- | BigN.one => constr:(true)
- | BigN.two => constr:(true)
- | _ => constr:(false)
- end.
-
-Ltac BigNcst t :=
- match isBigNcst t with
- | true => constr:(t)
- | false => constr:(NotConstant)
- end.
-
-Ltac BigN_to_N t :=
- match isBigNcst t with
- | true => eval vm_compute in (BigN.to_N t)
- | false => constr:(NotConstant)
- end.
-
-Ltac Ncst t :=
- match isNcst t with
- | true => constr:(t)
- | false => constr:(NotConstant)
- end.
-
-(** Registration for the "ring" tactic *)
-
-Add Ring BigNr : BigNring
- (decidable BigNeqb_correct,
- constants [BigNcst],
- power_tac BigNpower [BigN_to_N],
- div BigNdiv).
-
-Section TestRing.
-Let test : forall x y, 1 + x*y^1 + x^2 + 1 == 1*1 + 1 + y*x + 1*x*x.
-intros. ring_simplify. reflexivity.
-Qed.
-End TestRing.
-
-(** We benefit also from an "order" tactic *)
-
-Ltac bigN_order := BigN.order.
-
-Section TestOrder.
-Let test : forall x y : bigN, x<=y -> y<=x -> x==y.
-Proof. bigN_order. Qed.
-End TestOrder.
-
-(** We can use at least a bit of (r)omega by translating to [Z]. *)
-
-Section TestOmega.
-Let test : forall x y : bigN, x<=y -> y<=x -> x==y.
-Proof. intros x y. BigN.zify. omega. Qed.
-End TestOmega.
-
-(** Todo: micromega *)
diff --git a/theories/Numbers/Natural/BigN/NMake.v b/theories/Numbers/Natural/BigN/NMake.v
deleted file mode 100644
index 1425041a1..000000000
--- a/theories/Numbers/Natural/BigN/NMake.v
+++ /dev/null
@@ -1,1706 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-(** * NMake *)
-
-(** From a cyclic Z/nZ representation to arbitrary precision natural numbers.*)
-
-(** NB: This file contain the part which is independent from the underlying
- representation. The representation-dependent (and macro-generated) part
- is now in [NMake_gen]. *)
-
-Require Import Bool BigNumPrelude ZArith Nnat Ndigits CyclicAxioms DoubleType
- Nbasic Wf_nat StreamMemo NSig NMake_gen.
-
-Module Make (W0:CyclicType) <: NType.
-
- (** Let's include the macro-generated part. Even if we can't functorize
- things (due to Eval red_t below), the rest of the module only uses
- elements mentionned in interface [NAbstract]. *)
-
- Include NMake_gen.Make W0.
-
- Open Scope Z_scope.
-
- Local Notation "[ x ]" := (to_Z x).
-
- Definition eq (x y : t) := [x] = [y].
-
- Declare Reduction red_t :=
- lazy beta iota delta
- [iter_t reduce same_level mk_t mk_t_S succ_t dom_t dom_op].
-
- Ltac red_t :=
- match goal with |- ?u => let v := (eval red_t in u) in change v end.
-
- (** * Generic results *)
-
- Tactic Notation "destr_t" constr(x) "as" simple_intropattern(pat) :=
- destruct (destr_t x) as pat; cbv zeta;
- rewrite ?iter_mk_t, ?spec_mk_t, ?spec_reduce.
-
- Lemma spec_same_level : forall A (P:Z->Z->A->Prop)
- (f : forall n, dom_t n -> dom_t n -> A),
- (forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y)) ->
- forall x y, P [x] [y] (same_level f x y).
- Proof.
- intros. apply spec_same_level_dep with (P:=fun _ => P); auto.
- Qed.
-
- Theorem spec_pos: forall x, 0 <= [x].
- Proof.
- intros x. destr_t x as (n,x). now case (ZnZ.spec_to_Z x).
- Qed.
-
- Lemma digits_dom_op_incr : forall n m, (n<=m)%nat ->
- (ZnZ.digits (dom_op n) <= ZnZ.digits (dom_op m))%positive.
- Proof.
- intros.
- change (Zpos (ZnZ.digits (dom_op n)) <= Zpos (ZnZ.digits (dom_op m))).
- rewrite !digits_dom_op, !Pshiftl_nat_Zpower.
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- apply Z.pow_le_mono_r; auto with zarith.
- Qed.
-
- Definition to_N (x : t) := Z.to_N (to_Z x).
-
- (** * Zero, One *)
-
- Definition zero := mk_t O ZnZ.zero.
- Definition one := mk_t O ZnZ.one.
-
- Theorem spec_0: [zero] = 0.
- Proof.
- unfold zero. rewrite spec_mk_t. exact ZnZ.spec_0.
- Qed.
-
- Theorem spec_1: [one] = 1.
- Proof.
- unfold one. rewrite spec_mk_t. exact ZnZ.spec_1.
- Qed.
-
- (** * Successor *)
-
- (** NB: it is crucial here and for the rest of this file to preserve
- the let-in's. They allow to pre-compute once and for all the
- field access to Z/nZ initial structures (when n=0..6). *)
-
- Local Notation succn := (fun n =>
- let op := dom_op n in
- let succ_c := ZnZ.succ_c in
- let one := ZnZ.one in
- fun x => match succ_c x with
- | C0 r => mk_t n r
- | C1 r => mk_t_S n (WW one r)
- end).
-
- Definition succ : t -> t := Eval red_t in iter_t succn.
-
- Lemma succ_fold : succ = iter_t succn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_succ: forall n, [succ n] = [n] + 1.
- Proof.
- intros x. rewrite succ_fold. destr_t x as (n,x).
- generalize (ZnZ.spec_succ_c x); case ZnZ.succ_c.
- intros. rewrite spec_mk_t. assumption.
- intros. unfold interp_carry in *.
- rewrite spec_mk_t_S. simpl. rewrite ZnZ.spec_1. assumption.
- Qed.
-
- (** Two *)
-
- (** Not really pretty, but since W0 might be Z/2Z, we're not sure
- there's a proper 2 there. *)
-
- Definition two := succ one.
-
- Lemma spec_2 : [two] = 2.
- Proof.
- unfold two. now rewrite spec_succ, spec_1.
- Qed.
-
- (** * Addition *)
-
- Local Notation addn := (fun n =>
- let op := dom_op n in
- let add_c := ZnZ.add_c in
- let one := ZnZ.one in
- fun x y =>match add_c x y with
- | C0 r => mk_t n r
- | C1 r => mk_t_S n (WW one r)
- end).
-
- Definition add : t -> t -> t := Eval red_t in same_level addn.
-
- Lemma add_fold : add = same_level addn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_add: forall x y, [add x y] = [x] + [y].
- Proof.
- intros x y. rewrite add_fold. apply spec_same_level; clear x y.
- intros n x y. cbv beta iota zeta.
- generalize (ZnZ.spec_add_c x y); case ZnZ.add_c; intros z H.
- rewrite spec_mk_t. assumption.
- rewrite spec_mk_t_S. unfold interp_carry in H.
- simpl. rewrite ZnZ.spec_1. assumption.
- Qed.
-
- (** * Predecessor *)
-
- Local Notation predn := (fun n =>
- let pred_c := ZnZ.pred_c in
- fun x => match pred_c x with
- | C0 r => reduce n r
- | C1 _ => zero
- end).
-
- Definition pred : t -> t := Eval red_t in iter_t predn.
-
- Lemma pred_fold : pred = iter_t predn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_pred_pos : forall x, 0 < [x] -> [pred x] = [x] - 1.
- Proof.
- intros x. rewrite pred_fold. destr_t x as (n,x). intros H.
- generalize (ZnZ.spec_pred_c x); case ZnZ.pred_c; intros y H'.
- rewrite spec_reduce. assumption.
- exfalso. unfold interp_carry in *.
- generalize (ZnZ.spec_to_Z x) (ZnZ.spec_to_Z y); auto with zarith.
- Qed.
-
- Theorem spec_pred0 : forall x, [x] = 0 -> [pred x] = 0.
- Proof.
- intros x. rewrite pred_fold. destr_t x as (n,x). intros H.
- generalize (ZnZ.spec_pred_c x); case ZnZ.pred_c; intros y H'.
- rewrite spec_reduce.
- unfold interp_carry in H'.
- generalize (ZnZ.spec_to_Z y); auto with zarith.
- exact spec_0.
- Qed.
-
- Lemma spec_pred x : [pred x] = Z.max 0 ([x]-1).
- Proof.
- rewrite Z.max_comm.
- destruct (Z.max_spec ([x]-1) 0) as [(H,->)|(H,->)].
- - apply spec_pred0; generalize (spec_pos x); auto with zarith.
- - apply spec_pred_pos; auto with zarith.
- Qed.
-
- (** * Subtraction *)
-
- Local Notation subn := (fun n =>
- let sub_c := ZnZ.sub_c in
- fun x y => match sub_c x y with
- | C0 r => reduce n r
- | C1 r => zero
- end).
-
- Definition sub : t -> t -> t := Eval red_t in same_level subn.
-
- Lemma sub_fold : sub = same_level subn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_sub_pos : forall x y, [y] <= [x] -> [sub x y] = [x] - [y].
- Proof.
- intros x y. rewrite sub_fold. apply spec_same_level. clear x y.
- intros n x y. simpl.
- generalize (ZnZ.spec_sub_c x y); case ZnZ.sub_c; intros z H LE.
- rewrite spec_reduce. assumption.
- unfold interp_carry in H.
- exfalso.
- generalize (ZnZ.spec_to_Z z); auto with zarith.
- Qed.
-
- Theorem spec_sub0 : forall x y, [x] < [y] -> [sub x y] = 0.
- Proof.
- intros x y. rewrite sub_fold. apply spec_same_level. clear x y.
- intros n x y. simpl.
- generalize (ZnZ.spec_sub_c x y); case ZnZ.sub_c; intros z H LE.
- rewrite spec_reduce.
- unfold interp_carry in H.
- generalize (ZnZ.spec_to_Z z); auto with zarith.
- exact spec_0.
- Qed.
-
- Lemma spec_sub : forall x y, [sub x y] = Z.max 0 ([x]-[y]).
- Proof.
- intros. destruct (Z.le_gt_cases [y] [x]).
- rewrite Z.max_r; auto with zarith. apply spec_sub_pos; auto.
- rewrite Z.max_l; auto with zarith. apply spec_sub0; auto.
- Qed.
-
- (** * Comparison *)
-
- Definition comparen_m n :
- forall m, word (dom_t n) (S m) -> dom_t n -> comparison :=
- let op := dom_op n in
- let zero := ZnZ.zero (Ops:=op) in
- let compare := ZnZ.compare (Ops:=op) in
- let compare0 := compare zero in
- fun m => compare_mn_1 (dom_t n) (dom_t n) zero compare compare0 compare (S m).
-
- Let spec_comparen_m:
- forall n m (x : word (dom_t n) (S m)) (y : dom_t n),
- comparen_m n m x y = Z.compare (eval n (S m) x) (ZnZ.to_Z y).
- Proof.
- intros n m x y.
- unfold comparen_m, eval.
- rewrite nmake_double.
- apply spec_compare_mn_1.
- exact ZnZ.spec_0.
- intros. apply ZnZ.spec_compare.
- exact ZnZ.spec_to_Z.
- exact ZnZ.spec_compare.
- exact ZnZ.spec_compare.
- exact ZnZ.spec_to_Z.
- Qed.
-
- Definition comparenm n m wx wy :=
- let mn := Max.max n m in
- let d := diff n m in
- let op := make_op mn in
- ZnZ.compare
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d))).
-
- Local Notation compare_folded :=
- (iter_sym _
- (fun n => ZnZ.compare (Ops:=dom_op n))
- comparen_m
- comparenm
- CompOpp).
-
- Definition compare : t -> t -> comparison :=
- Eval lazy beta iota delta [iter_sym dom_op dom_t comparen_m] in
- compare_folded.
-
- Lemma compare_fold : compare = compare_folded.
- Proof.
- lazy beta iota delta [iter_sym dom_op dom_t comparen_m]. reflexivity.
- Qed.
-
- Theorem spec_compare : forall x y,
- compare x y = Z.compare [x] [y].
- Proof.
- intros x y. rewrite compare_fold. apply spec_iter_sym; clear x y.
- intros. apply ZnZ.spec_compare.
- intros. cbv beta zeta. apply spec_comparen_m.
- intros n m x y; unfold comparenm.
- rewrite (spec_cast_l n m x), (spec_cast_r n m y).
- unfold to_Z; apply ZnZ.spec_compare.
- intros. subst. now rewrite <- Z.compare_antisym.
- Qed.
-
- Definition eqb (x y : t) : bool :=
- match compare x y with
- | Eq => true
- | _ => false
- end.
-
- Theorem spec_eqb x y : eqb x y = Z.eqb [x] [y].
- Proof.
- apply eq_iff_eq_true.
- unfold eqb. rewrite Z.eqb_eq, <- Z.compare_eq_iff, spec_compare.
- split; [now destruct Z.compare | now intros ->].
- Qed.
-
- Definition lt (n m : t) := [n] < [m].
- Definition le (n m : t) := [n] <= [m].
-
- Definition ltb (x y : t) : bool :=
- match compare x y with
- | Lt => true
- | _ => false
- end.
-
- Theorem spec_ltb x y : ltb x y = Z.ltb [x] [y].
- Proof.
- apply eq_iff_eq_true.
- rewrite Z.ltb_lt. unfold Z.lt, ltb. rewrite spec_compare.
- split; [now destruct Z.compare | now intros ->].
- Qed.
-
- Definition leb (x y : t) : bool :=
- match compare x y with
- | Gt => false
- | _ => true
- end.
-
- Theorem spec_leb x y : leb x y = Z.leb [x] [y].
- Proof.
- apply eq_iff_eq_true.
- rewrite Z.leb_le. unfold Z.le, leb. rewrite spec_compare.
- now destruct Z.compare; split.
- Qed.
-
- Definition min (n m : t) : t := match compare n m with Gt => m | _ => n end.
- Definition max (n m : t) : t := match compare n m with Lt => m | _ => n end.
-
- Theorem spec_max : forall n m, [max n m] = Z.max [n] [m].
- Proof.
- intros. unfold max, Z.max. rewrite spec_compare; destruct Z.compare; reflexivity.
- Qed.
-
- Theorem spec_min : forall n m, [min n m] = Z.min [n] [m].
- Proof.
- intros. unfold min, Z.min. rewrite spec_compare; destruct Z.compare; reflexivity.
- Qed.
-
- (** * Multiplication *)
-
- Definition wn_mul n : forall m, word (dom_t n) (S m) -> dom_t n -> t :=
- let op := dom_op n in
- let zero := ZnZ.zero in
- let succ := ZnZ.succ (Ops:=op) in
- let add_c := ZnZ.add_c (Ops:=op) in
- let mul_c := ZnZ.mul_c (Ops:=op) in
- let ww := @ZnZ.WW _ op in
- let ow := @ZnZ.OW _ op in
- let eq0 := ZnZ.eq0 in
- let mul_add := @DoubleMul.w_mul_add _ zero succ add_c mul_c in
- let mul_add_n1 := @DoubleMul.double_mul_add_n1 _ zero ww ow mul_add in
- fun m x y =>
- let (w,r) := mul_add_n1 (S m) x y zero in
- if eq0 w then mk_t_w' n m r
- else mk_t_w' n (S m) (WW (extend n m w) r).
-
- Definition mulnm n m x y :=
- let mn := Max.max n m in
- let d := diff n m in
- let op := make_op mn in
- reduce_n (S mn) (ZnZ.mul_c
- (castm (diff_r n m) (extend_tr x (snd d)))
- (castm (diff_l n m) (extend_tr y (fst d)))).
-
- Local Notation mul_folded :=
- (iter_sym _
- (fun n => let mul_c := ZnZ.mul_c in
- fun x y => reduce (S n) (succ_t _ (mul_c x y)))
- wn_mul
- mulnm
- (fun x => x)).
-
- Definition mul : t -> t -> t :=
- Eval lazy beta iota delta
- [iter_sym dom_op dom_t reduce succ_t extend zeron
- wn_mul DoubleMul.w_mul_add mk_t_w'] in
- mul_folded.
-
- Lemma mul_fold : mul = mul_folded.
- Proof.
- lazy beta iota delta
- [iter_sym dom_op dom_t reduce succ_t extend zeron
- wn_mul DoubleMul.w_mul_add mk_t_w']. reflexivity.
- Qed.
-
- Lemma spec_muln:
- forall n (x: word _ (S n)) y,
- [Nn (S n) (ZnZ.mul_c (Ops:=make_op n) x y)] = [Nn n x] * [Nn n y].
- Proof.
- intros n x y; unfold to_Z.
- rewrite <- ZnZ.spec_mul_c.
- rewrite make_op_S.
- case ZnZ.mul_c; auto.
- Qed.
-
- Lemma spec_mul_add_n1: forall n m x y z,
- let (q,r) := DoubleMul.double_mul_add_n1 ZnZ.zero ZnZ.WW ZnZ.OW
- (DoubleMul.w_mul_add ZnZ.zero ZnZ.succ ZnZ.add_c ZnZ.mul_c)
- (S m) x y z in
- ZnZ.to_Z q * (base (ZnZ.digits (nmake_op _ (dom_op n) (S m))))
- + eval n (S m) r =
- eval n (S m) x * ZnZ.to_Z y + ZnZ.to_Z z.
- Proof.
- intros n m x y z.
- rewrite digits_nmake.
- unfold eval. rewrite nmake_double.
- apply DoubleMul.spec_double_mul_add_n1.
- apply ZnZ.spec_0.
- exact ZnZ.spec_WW.
- exact ZnZ.spec_OW.
- apply DoubleCyclic.spec_mul_add.
- Qed.
-
- Lemma spec_wn_mul : forall n m x y,
- [wn_mul n m x y] = (eval n (S m) x) * ZnZ.to_Z y.
- Proof.
- intros; unfold wn_mul.
- generalize (spec_mul_add_n1 n m x y ZnZ.zero).
- case DoubleMul.double_mul_add_n1; intros q r Hqr.
- rewrite ZnZ.spec_0, Z.add_0_r in Hqr. rewrite <- Hqr.
- generalize (ZnZ.spec_eq0 q); case ZnZ.eq0; intros HH.
- rewrite HH; auto. simpl. apply spec_mk_t_w'.
- clear.
- rewrite spec_mk_t_w'.
- set (m' := S m) in *.
- unfold eval.
- rewrite nmake_WW. f_equal. f_equal.
- rewrite <- spec_mk_t.
- symmetry. apply spec_extend.
- Qed.
-
- Theorem spec_mul : forall x y, [mul x y] = [x] * [y].
- Proof.
- intros x y. rewrite mul_fold. apply spec_iter_sym; clear x y.
- intros n x y. cbv zeta beta.
- rewrite spec_reduce, spec_succ_t, <- ZnZ.spec_mul_c; auto.
- apply spec_wn_mul.
- intros n m x y; unfold mulnm. rewrite spec_reduce_n.
- rewrite (spec_cast_l n m x), (spec_cast_r n m y).
- apply spec_muln.
- intros. rewrite Z.mul_comm; auto.
- Qed.
-
- (** * Division by a smaller number *)
-
- Definition wn_divn1 n :=
- let op := dom_op n in
- let zd := ZnZ.zdigits op in
- let zero := ZnZ.zero in
- let ww := ZnZ.WW in
- let head0 := ZnZ.head0 in
- let add_mul_div := ZnZ.add_mul_div in
- let div21 := ZnZ.div21 in
- let compare := ZnZ.compare in
- let sub := ZnZ.sub in
- let ddivn1 :=
- DoubleDivn1.double_divn1 zd zero ww head0 add_mul_div div21 compare sub in
- fun m x y => let (u,v) := ddivn1 (S m) x y in (mk_t_w' n m u, mk_t n v).
-
- Definition div_gtnm n m wx wy :=
- let mn := Max.max n m in
- let d := diff n m in
- let op := make_op mn in
- let (q, r):= ZnZ.div_gt
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d))) in
- (reduce_n mn q, reduce_n mn r).
-
- Local Notation div_gt_folded :=
- (iter _
- (fun n => let div_gt := ZnZ.div_gt in
- fun x y => let (u,v) := div_gt x y in (reduce n u, reduce n v))
- (fun n =>
- let div_gt := ZnZ.div_gt in
- fun m x y =>
- let y' := DoubleBase.get_low (zeron n) (S m) y in
- let (u,v) := div_gt x y' in (reduce n u, reduce n v))
- wn_divn1
- div_gtnm).
-
- Definition div_gt :=
- Eval lazy beta iota delta
- [iter dom_op dom_t reduce zeron wn_divn1 mk_t_w' mk_t] in
- div_gt_folded.
-
- Lemma div_gt_fold : div_gt = div_gt_folded.
- Proof.
- lazy beta iota delta [iter dom_op dom_t reduce zeron wn_divn1 mk_t_w' mk_t].
- reflexivity.
- Qed.
-
- Lemma spec_get_endn: forall n m x y,
- eval n m x <= [mk_t n y] ->
- [mk_t n (DoubleBase.get_low (zeron n) m x)] = eval n m x.
- Proof.
- intros n m x y H.
- unfold eval. rewrite nmake_double.
- rewrite spec_mk_t in *.
- apply DoubleBase.spec_get_low.
- apply spec_zeron.
- exact ZnZ.spec_to_Z.
- apply Z.le_lt_trans with (ZnZ.to_Z y); auto.
- rewrite <- nmake_double; auto.
- case (ZnZ.spec_to_Z y); auto.
- Qed.
-
- Definition spec_divn1 n :=
- DoubleDivn1.spec_double_divn1
- (ZnZ.zdigits (dom_op n)) (ZnZ.zero:dom_t n)
- ZnZ.WW ZnZ.head0
- ZnZ.add_mul_div ZnZ.div21
- ZnZ.compare ZnZ.sub ZnZ.to_Z
- ZnZ.spec_to_Z
- ZnZ.spec_zdigits
- ZnZ.spec_0 ZnZ.spec_WW ZnZ.spec_head0
- ZnZ.spec_add_mul_div ZnZ.spec_div21
- ZnZ.spec_compare ZnZ.spec_sub.
-
- Lemma spec_div_gt_aux : forall x y, [x] > [y] -> 0 < [y] ->
- let (q,r) := div_gt x y in
- [x] = [q] * [y] + [r] /\ 0 <= [r] < [y].
- Proof.
- intros x y. rewrite div_gt_fold. apply spec_iter; clear x y.
- intros n x y H1 H2. simpl.
- generalize (ZnZ.spec_div_gt x y H1 H2); case ZnZ.div_gt.
- intros u v. rewrite 2 spec_reduce. auto.
- intros n m x y H1 H2. cbv zeta beta.
- generalize (ZnZ.spec_div_gt x
- (DoubleBase.get_low (zeron n) (S m) y)).
- case ZnZ.div_gt.
- intros u v H3; repeat rewrite spec_reduce.
- generalize (spec_get_endn n (S m) y x). rewrite !spec_mk_t. intros H4.
- rewrite H4 in H3; auto with zarith.
- intros n m x y H1 H2.
- generalize (spec_divn1 n (S m) x y H2).
- unfold wn_divn1; case DoubleDivn1.double_divn1.
- intros u v H3.
- rewrite spec_mk_t_w', spec_mk_t.
- rewrite <- !nmake_double in H3; auto.
- intros n m x y H1 H2; unfold div_gtnm.
- generalize (ZnZ.spec_div_gt
- (castm (diff_r n m)
- (extend_tr x (snd (diff n m))))
- (castm (diff_l n m)
- (extend_tr y (fst (diff n m))))).
- case ZnZ.div_gt.
- intros xx yy HH.
- repeat rewrite spec_reduce_n.
- rewrite (spec_cast_l n m x), (spec_cast_r n m y).
- unfold to_Z; apply HH.
- rewrite (spec_cast_l n m x) in H1; auto.
- rewrite (spec_cast_r n m y) in H1; auto.
- rewrite (spec_cast_r n m y) in H2; auto.
- Qed.
-
- Theorem spec_div_gt: forall x y, [x] > [y] -> 0 < [y] ->
- let (q,r) := div_gt x y in
- [q] = [x] / [y] /\ [r] = [x] mod [y].
- Proof.
- intros x y H1 H2; generalize (spec_div_gt_aux x y H1 H2); case div_gt.
- intros q r (H3, H4); split.
- apply (Zdiv_unique [x] [y] [q] [r]); auto.
- rewrite Z.mul_comm; auto.
- apply (Zmod_unique [x] [y] [q] [r]); auto.
- rewrite Z.mul_comm; auto.
- Qed.
-
- (** * General Division *)
-
- Definition div_eucl (x y : t) : t * t :=
- if eqb y zero then (zero,zero) else
- match compare x y with
- | Eq => (one, zero)
- | Lt => (zero, x)
- | Gt => div_gt x y
- end.
-
- Theorem spec_div_eucl: forall x y,
- let (q,r) := div_eucl x y in
- ([q], [r]) = Z.div_eucl [x] [y].
- Proof.
- intros x y. unfold div_eucl.
- rewrite spec_eqb, spec_compare, spec_0.
- case Z.eqb_spec.
- intros ->. rewrite spec_0. destruct [x]; auto.
- intros H'.
- assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
- clear H'.
- case Z.compare_spec; intros Cmp;
- rewrite ?spec_0, ?spec_1; intros; auto with zarith.
- rewrite Cmp; generalize (Z_div_same [y] (Z.lt_gt _ _ H))
- (Z_mod_same [y] (Z.lt_gt _ _ H));
- unfold Z.div, Z.modulo; case Z.div_eucl; intros; subst; auto.
- assert (LeLt: 0 <= [x] < [y]) by (generalize (spec_pos x); auto).
- generalize (Zdiv_small _ _ LeLt) (Zmod_small _ _ LeLt);
- unfold Z.div, Z.modulo; case Z.div_eucl; intros; subst; auto.
- generalize (spec_div_gt _ _ (Z.lt_gt _ _ Cmp) H); auto.
- unfold Z.div, Z.modulo; case Z.div_eucl; case div_gt.
- intros a b c d (H1, H2); subst; auto.
- Qed.
-
- Definition div (x y : t) : t := fst (div_eucl x y).
-
- Theorem spec_div:
- forall x y, [div x y] = [x] / [y].
- Proof.
- intros x y; unfold div; generalize (spec_div_eucl x y);
- case div_eucl; simpl fst.
- intros xx yy; unfold Z.div; case Z.div_eucl; intros qq rr H;
- injection H; auto.
- Qed.
-
- (** * Modulo by a smaller number *)
-
- Definition wn_modn1 n :=
- let op := dom_op n in
- let zd := ZnZ.zdigits op in
- let zero := ZnZ.zero in
- let head0 := ZnZ.head0 in
- let add_mul_div := ZnZ.add_mul_div in
- let div21 := ZnZ.div21 in
- let compare := ZnZ.compare in
- let sub := ZnZ.sub in
- let dmodn1 :=
- DoubleDivn1.double_modn1 zd zero head0 add_mul_div div21 compare sub in
- fun m x y => reduce n (dmodn1 (S m) x y).
-
- Definition mod_gtnm n m wx wy :=
- let mn := Max.max n m in
- let d := diff n m in
- let op := make_op mn in
- reduce_n mn (ZnZ.modulo_gt
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d)))).
-
- Local Notation mod_gt_folded :=
- (iter _
- (fun n => let modulo_gt := ZnZ.modulo_gt in
- fun x y => reduce n (modulo_gt x y))
- (fun n => let modulo_gt := ZnZ.modulo_gt in
- fun m x y =>
- reduce n (modulo_gt x (DoubleBase.get_low (zeron n) (S m) y)))
- wn_modn1
- mod_gtnm).
-
- Definition mod_gt :=
- Eval lazy beta iota delta [iter dom_op dom_t reduce wn_modn1 zeron] in
- mod_gt_folded.
-
- Lemma mod_gt_fold : mod_gt = mod_gt_folded.
- Proof.
- lazy beta iota delta [iter dom_op dom_t reduce wn_modn1 zeron].
- reflexivity.
- Qed.
-
- Definition spec_modn1 n :=
- DoubleDivn1.spec_double_modn1
- (ZnZ.zdigits (dom_op n)) (ZnZ.zero:dom_t n)
- ZnZ.WW ZnZ.head0
- ZnZ.add_mul_div ZnZ.div21
- ZnZ.compare ZnZ.sub ZnZ.to_Z
- ZnZ.spec_to_Z
- ZnZ.spec_zdigits
- ZnZ.spec_0 ZnZ.spec_WW ZnZ.spec_head0
- ZnZ.spec_add_mul_div ZnZ.spec_div21
- ZnZ.spec_compare ZnZ.spec_sub.
-
- Theorem spec_mod_gt:
- forall x y, [x] > [y] -> 0 < [y] -> [mod_gt x y] = [x] mod [y].
- Proof.
- intros x y. rewrite mod_gt_fold. apply spec_iter; clear x y.
- intros n x y H1 H2. simpl. rewrite spec_reduce.
- exact (ZnZ.spec_modulo_gt x y H1 H2).
- intros n m x y H1 H2. cbv zeta beta. rewrite spec_reduce.
- rewrite <- spec_mk_t in H1.
- rewrite <- (spec_get_endn n (S m) y x); auto with zarith.
- rewrite spec_mk_t.
- apply ZnZ.spec_modulo_gt; auto.
- rewrite <- (spec_get_endn n (S m) y x), !spec_mk_t in H1; auto with zarith.
- rewrite <- (spec_get_endn n (S m) y x), !spec_mk_t in H2; auto with zarith.
- intros n m x y H1 H2. unfold wn_modn1. rewrite spec_reduce.
- unfold eval; rewrite nmake_double.
- apply (spec_modn1 n); auto.
- intros n m x y H1 H2; unfold mod_gtnm.
- repeat rewrite spec_reduce_n.
- rewrite (spec_cast_l n m x), (spec_cast_r n m y).
- unfold to_Z; apply ZnZ.spec_modulo_gt.
- rewrite (spec_cast_l n m x) in H1; auto.
- rewrite (spec_cast_r n m y) in H1; auto.
- rewrite (spec_cast_r n m y) in H2; auto.
- Qed.
-
- (** * General Modulo *)
-
- Definition modulo (x y : t) : t :=
- if eqb y zero then zero else
- match compare x y with
- | Eq => zero
- | Lt => x
- | Gt => mod_gt x y
- end.
-
- Theorem spec_modulo:
- forall x y, [modulo x y] = [x] mod [y].
- Proof.
- intros x y. unfold modulo.
- rewrite spec_eqb, spec_compare, spec_0.
- case Z.eqb_spec.
- intros ->; rewrite spec_0. destruct [x]; auto.
- intro H'.
- assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
- clear H'.
- case Z.compare_spec;
- rewrite ?spec_0, ?spec_1; intros; try split; auto with zarith.
- rewrite H0; symmetry; apply Z_mod_same; auto with zarith.
- symmetry; apply Zmod_small; auto with zarith.
- generalize (spec_pos x); auto with zarith.
- apply spec_mod_gt; auto with zarith.
- Qed.
-
- (** * Square *)
-
- Local Notation squaren := (fun n =>
- let square_c := ZnZ.square_c in
- fun x => reduce (S n) (succ_t _ (square_c x))).
-
- Definition square : t -> t := Eval red_t in iter_t squaren.
-
- Lemma square_fold : square = iter_t squaren.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_square: forall x, [square x] = [x] * [x].
- Proof.
- intros x. rewrite square_fold. destr_t x as (n,x).
- rewrite spec_succ_t. exact (ZnZ.spec_square_c x).
- Qed.
-
- (** * Square Root *)
-
- Local Notation sqrtn := (fun n =>
- let sqrt := ZnZ.sqrt in
- fun x => reduce n (sqrt x)).
-
- Definition sqrt : t -> t := Eval red_t in iter_t sqrtn.
-
- Lemma sqrt_fold : sqrt = iter_t sqrtn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_sqrt_aux: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
- Proof.
- intros x. rewrite sqrt_fold. destr_t x as (n,x). exact (ZnZ.spec_sqrt x).
- Qed.
-
- Theorem spec_sqrt: forall x, [sqrt x] = Z.sqrt [x].
- Proof.
- intros x.
- symmetry. apply Z.sqrt_unique.
- rewrite <- ! Z.pow_2_r. apply spec_sqrt_aux.
- Qed.
-
- (** * Power *)
-
- Fixpoint pow_pos (x:t)(p:positive) : t :=
- match p with
- | xH => x
- | xO p => square (pow_pos x p)
- | xI p => mul (square (pow_pos x p)) x
- end.
-
- Theorem spec_pow_pos: forall x n, [pow_pos x n] = [x] ^ Zpos n.
- Proof.
- intros x n; generalize x; elim n; clear n x; simpl pow_pos.
- intros; rewrite spec_mul; rewrite spec_square; rewrite H.
- rewrite Pos2Z.inj_xI; rewrite Zpower_exp; auto with zarith.
- rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r; auto with zarith.
- rewrite Z.pow_2_r; rewrite Z.pow_1_r; auto.
- intros; rewrite spec_square; rewrite H.
- rewrite Pos2Z.inj_xO; auto with zarith.
- rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r; auto with zarith.
- rewrite Z.pow_2_r; auto.
- intros; rewrite Z.pow_1_r; auto.
- Qed.
-
- Definition pow_N (x:t)(n:N) : t := match n with
- | BinNat.N0 => one
- | BinNat.Npos p => pow_pos x p
- end.
-
- Theorem spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z.of_N n.
- Proof.
- destruct n; simpl. apply spec_1.
- apply spec_pow_pos.
- Qed.
-
- Definition pow (x y:t) : t := pow_N x (to_N y).
-
- Theorem spec_pow : forall x y, [pow x y] = [x] ^ [y].
- Proof.
- intros. unfold pow, to_N.
- now rewrite spec_pow_N, Z2N.id by apply spec_pos.
- Qed.
-
-
- (** * digits
-
- Number of digits in the representation of a numbers
- (including head zero's).
- NB: This function isn't a morphism for setoid [eq].
- *)
-
- Local Notation digitsn := (fun n =>
- let digits := ZnZ.digits (dom_op n) in
- fun _ => digits).
-
- Definition digits : t -> positive := Eval red_t in iter_t digitsn.
-
- Lemma digits_fold : digits = iter_t digitsn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_digits: forall x, 0 <= [x] < 2 ^ Zpos (digits x).
- Proof.
- intros x. rewrite digits_fold. destr_t x as (n,x). exact (ZnZ.spec_to_Z x).
- Qed.
-
- Lemma digits_level : forall x, digits x = ZnZ.digits (dom_op (level x)).
- Proof.
- intros x. rewrite digits_fold. unfold level. destr_t x as (n,x). reflexivity.
- Qed.
-
- (** * Gcd *)
-
- Definition gcd_gt_body a b cont :=
- match compare b zero with
- | Gt =>
- let r := mod_gt a b in
- match compare r zero with
- | Gt => cont r (mod_gt b r)
- | _ => b
- end
- | _ => a
- end.
-
- Theorem Zspec_gcd_gt_body: forall a b cont p,
- [a] > [b] -> [a] < 2 ^ p ->
- (forall a1 b1, [a1] < 2 ^ (p - 1) -> [a1] > [b1] ->
- Zis_gcd [a1] [b1] [cont a1 b1]) ->
- Zis_gcd [a] [b] [gcd_gt_body a b cont].
- Proof.
- intros a b cont p H2 H3 H4; unfold gcd_gt_body.
- rewrite ! spec_compare, spec_0. case Z.compare_spec.
- intros ->; apply Zis_gcd_0.
- intros HH; absurd (0 <= [b]); auto with zarith.
- case (spec_digits b); auto with zarith.
- intros H5; case Z.compare_spec.
- intros H6; rewrite <- (Z.mul_1_r [b]).
- rewrite (Z_div_mod_eq [a] [b]); auto with zarith.
- rewrite <- spec_mod_gt; auto with zarith.
- rewrite H6; rewrite Z.add_0_r.
- apply Zis_gcd_mult; apply Zis_gcd_1.
- intros; apply False_ind.
- case (spec_digits (mod_gt a b)); auto with zarith.
- intros H6; apply DoubleDiv.Zis_gcd_mod; auto with zarith.
- apply DoubleDiv.Zis_gcd_mod; auto with zarith.
- rewrite <- spec_mod_gt; auto with zarith.
- assert (F2: [b] > [mod_gt a b]).
- case (Z_mod_lt [a] [b]); auto with zarith.
- repeat rewrite <- spec_mod_gt; auto with zarith.
- assert (F3: [mod_gt a b] > [mod_gt b (mod_gt a b)]).
- case (Z_mod_lt [b] [mod_gt a b]); auto with zarith.
- rewrite <- spec_mod_gt; auto with zarith.
- repeat rewrite <- spec_mod_gt; auto with zarith.
- apply H4; auto with zarith.
- apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
- apply Z.le_lt_trans with ([b] + [mod_gt a b]); auto with zarith.
- apply Z.le_lt_trans with (([a]/[b]) * [b] + [mod_gt a b]); auto with zarith.
- - apply Z.add_le_mono_r.
- rewrite <- (Z.mul_1_l [b]) at 1.
- apply Z.mul_le_mono_nonneg_r; auto with zarith.
- change 1 with (Z.succ 0). apply Z.le_succ_l.
- apply Z.div_str_pos; auto with zarith.
- - rewrite Z.mul_comm; rewrite spec_mod_gt; auto with zarith.
- rewrite <- Z_div_mod_eq; auto with zarith.
- rewrite Z.mul_comm, <- Z.pow_succ_r, Z.sub_1_r, Z.succ_pred; auto.
- apply Z.le_0_sub. change 1 with (Z.succ 0). apply Z.le_succ_l.
- destruct p; simpl in H3; auto with zarith.
- Qed.
-
- Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) : t :=
- gcd_gt_body a b
- (fun a b =>
- match p with
- | xH => cont a b
- | xO p => gcd_gt_aux p (gcd_gt_aux p cont) a b
- | xI p => gcd_gt_aux p (gcd_gt_aux p cont) a b
- end).
-
- Theorem Zspec_gcd_gt_aux: forall p n a b cont,
- [a] > [b] -> [a] < 2 ^ (Zpos p + n) ->
- (forall a1 b1, [a1] < 2 ^ n -> [a1] > [b1] ->
- Zis_gcd [a1] [b1] [cont a1 b1]) ->
- Zis_gcd [a] [b] [gcd_gt_aux p cont a b].
- intros p; elim p; clear p.
- intros p Hrec n a b cont H2 H3 H4.
- unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xI p) + n); auto.
- intros a1 b1 H6 H7.
- apply Hrec with (Zpos p + n); auto.
- replace (Zpos p + (Zpos p + n)) with
- (Zpos (xI p) + n - 1); auto.
- rewrite Pos2Z.inj_xI; ring.
- intros a2 b2 H9 H10.
- apply Hrec with n; auto.
- intros p Hrec n a b cont H2 H3 H4.
- unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xO p) + n); auto.
- intros a1 b1 H6 H7.
- apply Hrec with (Zpos p + n - 1); auto.
- replace (Zpos p + (Zpos p + n - 1)) with
- (Zpos (xO p) + n - 1); auto.
- rewrite Pos2Z.inj_xO; ring.
- intros a2 b2 H9 H10.
- apply Hrec with (n - 1); auto.
- replace (Zpos p + (n - 1)) with
- (Zpos p + n - 1); auto with zarith.
- intros a3 b3 H12 H13; apply H4; auto with zarith.
- apply Z.lt_le_trans with (1 := H12).
- apply Z.pow_le_mono_r; auto with zarith.
- intros n a b cont H H2 H3.
- simpl gcd_gt_aux.
- apply Zspec_gcd_gt_body with (n + 1); auto with zarith.
- rewrite Z.add_comm; auto.
- intros a1 b1 H5 H6; apply H3; auto.
- replace n with (n + 1 - 1); auto; try ring.
- Qed.
-
- Definition gcd_cont a b :=
- match compare one b with
- | Eq => one
- | _ => a
- end.
-
- Definition gcd_gt a b := gcd_gt_aux (digits a) gcd_cont a b.
-
- Theorem spec_gcd_gt: forall a b,
- [a] > [b] -> [gcd_gt a b] = Z.gcd [a] [b].
- Proof.
- intros a b H2.
- case (spec_digits (gcd_gt a b)); intros H3 H4.
- case (spec_digits a); intros H5 H6.
- symmetry; apply Zis_gcd_gcd; auto with zarith.
- unfold gcd_gt; apply Zspec_gcd_gt_aux with 0; auto with zarith.
- intros a1 a2; rewrite Z.pow_0_r.
- case (spec_digits a2); intros H7 H8;
- intros; apply False_ind; auto with zarith.
- Qed.
-
- Definition gcd (a b : t) : t :=
- match compare a b with
- | Eq => a
- | Lt => gcd_gt b a
- | Gt => gcd_gt a b
- end.
-
- Theorem spec_gcd: forall a b, [gcd a b] = Z.gcd [a] [b].
- Proof.
- intros a b.
- case (spec_digits a); intros H1 H2.
- case (spec_digits b); intros H3 H4.
- unfold gcd. rewrite spec_compare. case Z.compare_spec.
- intros HH; rewrite HH; symmetry; apply Zis_gcd_gcd; auto.
- apply Zis_gcd_refl.
- intros; transitivity (Z.gcd [b] [a]).
- apply spec_gcd_gt; auto with zarith.
- apply Zis_gcd_gcd; auto with zarith.
- apply Z.gcd_nonneg.
- apply Zis_gcd_sym; apply Zgcd_is_gcd.
- intros; apply spec_gcd_gt; auto with zarith.
- Qed.
-
- (** * Parity test *)
-
- Definition even : t -> bool := Eval red_t in
- iter_t (fun n x => ZnZ.is_even x).
-
- Definition odd x := negb (even x).
-
- Lemma even_fold : even = iter_t (fun n x => ZnZ.is_even x).
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_even_aux: forall x,
- if even x then [x] mod 2 = 0 else [x] mod 2 = 1.
- Proof.
- intros x. rewrite even_fold. destr_t x as (n,x).
- exact (ZnZ.spec_is_even x).
- Qed.
-
- Theorem spec_even: forall x, even x = Z.even [x].
- Proof.
- intros x. assert (H := spec_even_aux x). symmetry.
- rewrite (Z.div_mod [x] 2); auto with zarith.
- destruct (even x); rewrite H, ?Z.add_0_r.
- rewrite Zeven_bool_iff. apply Zeven_2p.
- apply not_true_is_false. rewrite Zeven_bool_iff.
- apply Zodd_not_Zeven. apply Zodd_2p_plus_1.
- Qed.
-
- Theorem spec_odd: forall x, odd x = Z.odd [x].
- Proof.
- intros x. unfold odd.
- assert (H := spec_even_aux x). symmetry.
- rewrite (Z.div_mod [x] 2); auto with zarith.
- destruct (even x); rewrite H, ?Z.add_0_r; simpl negb.
- apply not_true_is_false. rewrite Zodd_bool_iff.
- apply Zeven_not_Zodd. apply Zeven_2p.
- apply Zodd_bool_iff. apply Zodd_2p_plus_1.
- Qed.
-
- (** * Conversion *)
-
- Definition pheight p :=
- Peano.pred (Pos.to_nat (get_height (ZnZ.digits (dom_op 0)) (plength p))).
-
- Theorem pheight_correct: forall p,
- Zpos p < 2 ^ (Zpos (ZnZ.digits (dom_op 0)) * 2 ^ (Z.of_nat (pheight p))).
- Proof.
- intros p; unfold pheight.
- rewrite Nat2Z.inj_pred by apply Pos2Nat.is_pos.
- rewrite positive_nat_Z.
- rewrite <- Z.sub_1_r.
- assert (F2:= (get_height_correct (ZnZ.digits (dom_op 0)) (plength p))).
- apply Z.lt_le_trans with (Zpos (Pos.succ p)).
- rewrite Pos2Z.inj_succ; auto with zarith.
- apply Z.le_trans with (1 := plength_pred_correct (Pos.succ p)).
- rewrite Pos.pred_succ.
- apply Z.pow_le_mono_r; auto with zarith.
- Qed.
-
- Definition of_pos (x:positive) : t :=
- let n := pheight x in
- reduce n (snd (ZnZ.of_pos x)).
-
- Theorem spec_of_pos: forall x,
- [of_pos x] = Zpos x.
- Proof.
- intros x; unfold of_pos.
- rewrite spec_reduce.
- simpl.
- apply ZnZ.of_pos_correct.
- unfold base.
- apply Z.lt_le_trans with (1 := pheight_correct x).
- apply Z.pow_le_mono_r; auto with zarith.
- rewrite (digits_dom_op (_ _)), Pshiftl_nat_Zpower. auto with zarith.
- Qed.
-
- Definition of_N (x:N) : t :=
- match x with
- | BinNat.N0 => zero
- | Npos p => of_pos p
- end.
-
- Theorem spec_of_N: forall x,
- [of_N x] = Z.of_N x.
- Proof.
- intros x; case x.
- simpl of_N. exact spec_0.
- intros p; exact (spec_of_pos p).
- Qed.
-
- (** * [head0] and [tail0]
-
- Number of zero at the beginning and at the end of
- the representation of the number.
- NB: these functions are not morphism for setoid [eq].
- *)
-
- Local Notation head0n := (fun n =>
- let head0 := ZnZ.head0 in
- fun x => reduce n (head0 x)).
-
- Definition head0 : t -> t := Eval red_t in iter_t head0n.
-
- Lemma head0_fold : head0 = iter_t head0n.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_head00: forall x, [x] = 0 -> [head0 x] = Zpos (digits x).
- Proof.
- intros x. rewrite head0_fold, digits_fold. destr_t x as (n,x).
- exact (ZnZ.spec_head00 x).
- Qed.
-
- Lemma pow2_pos_minus_1 : forall z, 0<z -> 2^(z-1) = 2^z / 2.
- Proof.
- intros. apply Zdiv_unique with 0; auto with zarith.
- change 2 with (2^1) at 2.
- rewrite <- Zpower_exp; auto with zarith.
- rewrite Z.add_0_r. f_equal. auto with zarith.
- Qed.
-
- Theorem spec_head0: forall x, 0 < [x] ->
- 2 ^ (Zpos (digits x) - 1) <= 2 ^ [head0 x] * [x] < 2 ^ Zpos (digits x).
- Proof.
- intros x. rewrite pow2_pos_minus_1 by (red; auto).
- rewrite head0_fold, digits_fold. destr_t x as (n,x). exact (ZnZ.spec_head0 x).
- Qed.
-
- Local Notation tail0n := (fun n =>
- let tail0 := ZnZ.tail0 in
- fun x => reduce n (tail0 x)).
-
- Definition tail0 : t -> t := Eval red_t in iter_t tail0n.
-
- Lemma tail0_fold : tail0 = iter_t tail0n.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_tail00: forall x, [x] = 0 -> [tail0 x] = Zpos (digits x).
- Proof.
- intros x. rewrite tail0_fold, digits_fold. destr_t x as (n,x).
- exact (ZnZ.spec_tail00 x).
- Qed.
-
- Theorem spec_tail0: forall x,
- 0 < [x] -> exists y, 0 <= y /\ [x] = (2 * y + 1) * 2 ^ [tail0 x].
- Proof.
- intros x. rewrite tail0_fold. destr_t x as (n,x). exact (ZnZ.spec_tail0 x).
- Qed.
-
- (** * [Ndigits]
-
- Same as [digits] but encoded using large integers
- NB: this function is not a morphism for setoid [eq].
- *)
-
- Local Notation Ndigitsn := (fun n =>
- let d := reduce n (ZnZ.zdigits (dom_op n)) in
- fun _ => d).
-
- Definition Ndigits : t -> t := Eval red_t in iter_t Ndigitsn.
-
- Lemma Ndigits_fold : Ndigits = iter_t Ndigitsn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_Ndigits: forall x, [Ndigits x] = Zpos (digits x).
- Proof.
- intros x. rewrite Ndigits_fold, digits_fold. destr_t x as (n,x).
- apply ZnZ.spec_zdigits.
- Qed.
-
- (** * Binary logarithm *)
-
- Local Notation log2n := (fun n =>
- let op := dom_op n in
- let zdigits := ZnZ.zdigits op in
- let head0 := ZnZ.head0 in
- let sub_carry := ZnZ.sub_carry in
- fun x => reduce n (sub_carry zdigits (head0 x))).
-
- Definition log2 : t -> t := Eval red_t in
- let log2 := iter_t log2n in
- fun x => if eqb x zero then zero else log2 x.
-
- Lemma log2_fold :
- log2 = fun x => if eqb x zero then zero else iter_t log2n x.
- Proof. red_t; reflexivity. Qed.
-
- Lemma spec_log2_0 : forall x, [x] = 0 -> [log2 x] = 0.
- Proof.
- intros x H. rewrite log2_fold.
- rewrite spec_eqb, H. rewrite spec_0. simpl. exact spec_0.
- Qed.
-
- Lemma head0_zdigits : forall n (x : dom_t n),
- 0 < ZnZ.to_Z x ->
- ZnZ.to_Z (ZnZ.head0 x) < ZnZ.to_Z (ZnZ.zdigits (dom_op n)).
- Proof.
- intros n x H.
- destruct (ZnZ.spec_head0 x H) as (_,H0).
- intros.
- assert (H1 := ZnZ.spec_to_Z (ZnZ.head0 x)).
- assert (H2 := ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
- unfold base in *.
- rewrite ZnZ.spec_zdigits in H2 |- *.
- set (h := ZnZ.to_Z (ZnZ.head0 x)) in *; clearbody h.
- set (d := ZnZ.digits (dom_op n)) in *; clearbody d.
- destruct (Z_lt_le_dec h (Zpos d)); auto. exfalso.
- assert (1 * 2^Zpos d <= ZnZ.to_Z x * 2^h).
- apply Z.mul_le_mono_nonneg; auto with zarith.
- apply Z.pow_le_mono_r; auto with zarith.
- rewrite Z.mul_comm in H0. auto with zarith.
- Qed.
-
- Lemma spec_log2_pos : forall x, [x]<>0 ->
- 2^[log2 x] <= [x] < 2^([log2 x]+1).
- Proof.
- intros x H. rewrite log2_fold.
- rewrite spec_eqb. rewrite spec_0.
- case Z.eqb_spec.
- auto with zarith.
- clear H.
- destr_t x as (n,x). intros H.
- rewrite ZnZ.spec_sub_carry.
- assert (H0 := ZnZ.spec_to_Z x).
- assert (H1 := ZnZ.spec_to_Z (ZnZ.head0 x)).
- assert (H2 := ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
- assert (H3 := head0_zdigits n x).
- rewrite Zmod_small by auto with zarith.
- rewrite Z.sub_simpl_r.
- rewrite (Z.mul_lt_mono_pos_l (2^(ZnZ.to_Z (ZnZ.head0 x))));
- auto with zarith.
- rewrite (Z.mul_le_mono_pos_l _ _ (2^(ZnZ.to_Z (ZnZ.head0 x))));
- auto with zarith.
- rewrite <- 2 Zpower_exp; auto with zarith.
- rewrite !Z.add_sub_assoc, !Z.add_simpl_l.
- rewrite ZnZ.spec_zdigits.
- rewrite pow2_pos_minus_1 by (red; auto).
- apply ZnZ.spec_head0; auto with zarith.
- Qed.
-
- Lemma spec_log2 : forall x, [log2 x] = Z.log2 [x].
- Proof.
- intros. destruct (Z_lt_ge_dec 0 [x]).
- symmetry. apply Z.log2_unique. apply spec_pos.
- apply spec_log2_pos. intro EQ; rewrite EQ in *; auto with zarith.
- rewrite spec_log2_0. rewrite Z.log2_nonpos; auto with zarith.
- generalize (spec_pos x); auto with zarith.
- Qed.
-
- Lemma log2_digits_head0 : forall x, 0 < [x] ->
- [log2 x] = Zpos (digits x) - [head0 x] - 1.
- Proof.
- intros. rewrite log2_fold.
- rewrite spec_eqb. rewrite spec_0.
- case Z.eqb_spec.
- auto with zarith.
- intros _. revert H. rewrite digits_fold, head0_fold. destr_t x as (n,x).
- rewrite ZnZ.spec_sub_carry.
- intros.
- generalize (head0_zdigits n x H).
- generalize (ZnZ.spec_to_Z (ZnZ.head0 x)).
- generalize (ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
- rewrite ZnZ.spec_zdigits. intros. apply Zmod_small.
- auto with zarith.
- Qed.
-
- (** * Right shift *)
-
- Local Notation shiftrn := (fun n =>
- let op := dom_op n in
- let zdigits := ZnZ.zdigits op in
- let sub_c := ZnZ.sub_c in
- let add_mul_div := ZnZ.add_mul_div in
- let zzero := ZnZ.zero in
- fun x p => match sub_c zdigits p with
- | C0 d => reduce n (add_mul_div d zzero x)
- | C1 _ => zero
- end).
-
- Definition shiftr : t -> t -> t := Eval red_t in
- same_level shiftrn.
-
- Lemma shiftr_fold : shiftr = same_level shiftrn.
- Proof. red_t; reflexivity. Qed.
-
- Lemma div_pow2_bound :forall x y z,
- 0 <= x -> 0 <= y -> x < z -> 0 <= x / 2 ^ y < z.
- Proof.
- intros x y z HH HH1 HH2.
- split; auto with zarith.
- apply Z.le_lt_trans with (2 := HH2); auto with zarith.
- apply Zdiv_le_upper_bound; auto with zarith.
- pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith.
- apply Z.mul_le_mono_nonneg_l; auto.
- apply Z.pow_le_mono_r; auto with zarith.
- rewrite Z.pow_0_r; ring.
- Qed.
-
- Theorem spec_shiftr_pow2 : forall x n,
- [shiftr x n] = [x] / 2 ^ [n].
- Proof.
- intros x y. rewrite shiftr_fold. apply spec_same_level. clear x y.
- intros n x p. simpl.
- assert (Hx := ZnZ.spec_to_Z x).
- assert (Hy := ZnZ.spec_to_Z p).
- generalize (ZnZ.spec_sub_c (ZnZ.zdigits (dom_op n)) p).
- case ZnZ.sub_c; intros d H; unfold interp_carry in *; simpl.
- (** Subtraction without underflow : [ p <= digits ] *)
- rewrite spec_reduce.
- rewrite ZnZ.spec_zdigits in H.
- rewrite ZnZ.spec_add_mul_div by auto with zarith.
- rewrite ZnZ.spec_0, Z.mul_0_l, Z.add_0_l.
- rewrite Zmod_small.
- f_equal. f_equal. auto with zarith.
- split. auto with zarith.
- apply div_pow2_bound; auto with zarith.
- (** Subtraction with underflow : [ digits < p ] *)
- rewrite ZnZ.spec_0. symmetry.
- apply Zdiv_small.
- split; auto with zarith.
- apply Z.lt_le_trans with (base (ZnZ.digits (dom_op n))); auto with zarith.
- unfold base. apply Z.pow_le_mono_r; auto with zarith.
- rewrite ZnZ.spec_zdigits in H.
- generalize (ZnZ.spec_to_Z d); auto with zarith.
- Qed.
-
- Lemma spec_shiftr: forall x p, [shiftr x p] = Z.shiftr [x] [p].
- Proof.
- intros.
- now rewrite spec_shiftr_pow2, Z.shiftr_div_pow2 by apply spec_pos.
- Qed.
-
- (** * Left shift *)
-
- (** First an unsafe version, working correctly only if
- the representation is large enough *)
-
- Local Notation unsafe_shiftln := (fun n =>
- let op := dom_op n in
- let add_mul_div := ZnZ.add_mul_div in
- let zero := ZnZ.zero in
- fun x p => reduce n (add_mul_div p x zero)).
-
- Definition unsafe_shiftl : t -> t -> t := Eval red_t in
- same_level unsafe_shiftln.
-
- Lemma unsafe_shiftl_fold : unsafe_shiftl = same_level unsafe_shiftln.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_unsafe_shiftl_aux : forall x p K,
- 0 <= K ->
- [x] < 2^K ->
- [p] + K <= Zpos (digits x) ->
- [unsafe_shiftl x p] = [x] * 2 ^ [p].
- Proof.
- intros x p.
- rewrite unsafe_shiftl_fold. rewrite digits_level.
- apply spec_same_level_dep.
- intros n m z z' r LE H K HK H1 H2. apply (H K); auto.
- transitivity (Zpos (ZnZ.digits (dom_op n))); auto.
- apply digits_dom_op_incr; auto.
- clear x p.
- intros n x p K HK Hx Hp. simpl. rewrite spec_reduce.
- destruct (ZnZ.spec_to_Z x).
- destruct (ZnZ.spec_to_Z p).
- rewrite ZnZ.spec_add_mul_div by (omega with *).
- rewrite ZnZ.spec_0, Zdiv_0_l, Z.add_0_r.
- apply Zmod_small. unfold base.
- split; auto with zarith.
- rewrite Z.mul_comm.
- apply Z.lt_le_trans with (2^(ZnZ.to_Z p + K)).
- rewrite Zpower_exp; auto with zarith.
- apply Z.mul_lt_mono_pos_l; auto with zarith.
- apply Z.pow_le_mono_r; auto with zarith.
- Qed.
-
- Theorem spec_unsafe_shiftl: forall x p,
- [p] <= [head0 x] -> [unsafe_shiftl x p] = [x] * 2 ^ [p].
- Proof.
- intros.
- destruct (Z.eq_dec [x] 0) as [EQ|NEQ].
- (* [x] = 0 *)
- apply spec_unsafe_shiftl_aux with 0; auto with zarith.
- now rewrite EQ.
- rewrite spec_head00 in *; auto with zarith.
- (* [x] <> 0 *)
- apply spec_unsafe_shiftl_aux with ([log2 x] + 1); auto with zarith.
- generalize (spec_pos (log2 x)); auto with zarith.
- destruct (spec_log2_pos x); auto with zarith.
- rewrite log2_digits_head0; auto with zarith.
- generalize (spec_pos x); auto with zarith.
- Qed.
-
- (** Then we define a function doubling the size of the representation
- but without changing the value of the number. *)
-
- Local Notation double_size_n := (fun n =>
- let zero := ZnZ.zero in
- fun x => mk_t_S n (WW zero x)).
-
- Definition double_size : t -> t := Eval red_t in
- iter_t double_size_n.
-
- Lemma double_size_fold : double_size = iter_t double_size_n.
- Proof. red_t; reflexivity. Qed.
-
- Lemma double_size_level : forall x, level (double_size x) = S (level x).
- Proof.
- intros x. rewrite double_size_fold; unfold level at 2. destr_t x as (n,x).
- apply mk_t_S_level.
- Qed.
-
- Theorem spec_double_size_digits:
- forall x, Zpos (digits (double_size x)) = 2 * (Zpos (digits x)).
- Proof.
- intros x. rewrite ! digits_level, double_size_level.
- rewrite 2 digits_dom_op, 2 Pshiftl_nat_Zpower,
- Nat2Z.inj_succ, Z.pow_succ_r; auto with zarith.
- ring.
- Qed.
-
- Theorem spec_double_size: forall x, [double_size x] = [x].
- Proof.
- intros x. rewrite double_size_fold. destr_t x as (n,x).
- rewrite spec_mk_t_S. simpl. rewrite ZnZ.spec_0. auto with zarith.
- Qed.
-
- Theorem spec_double_size_head0:
- forall x, 2 * [head0 x] <= [head0 (double_size x)].
- Proof.
- intros x.
- assert (F1:= spec_pos (head0 x)).
- assert (F2: 0 < Zpos (digits x)).
- red; auto.
- assert (HH := spec_pos x). Z.le_elim HH.
- generalize HH; rewrite <- (spec_double_size x); intros HH1.
- case (spec_head0 x HH); intros _ HH2.
- case (spec_head0 _ HH1).
- rewrite (spec_double_size x); rewrite (spec_double_size_digits x).
- intros HH3 _.
- case (Z.le_gt_cases ([head0 (double_size x)]) (2 * [head0 x])); auto; intros HH4.
- absurd (2 ^ (2 * [head0 x] )* [x] < 2 ^ [head0 (double_size x)] * [x]); auto.
- apply Z.le_ngt.
- apply Z.mul_le_mono_nonneg_r; auto with zarith.
- apply Z.pow_le_mono_r; auto; auto with zarith.
- assert (HH5: 2 ^[head0 x] <= 2 ^(Zpos (digits x) - 1)).
- { apply Z.le_succ_l in HH. change (1 <= [x]) in HH.
- Z.le_elim HH.
- - apply Z.mul_le_mono_pos_r with (2 ^ 1); auto with zarith.
- rewrite <- (fun x y z => Z.pow_add_r x (y - z)); auto with zarith.
- rewrite Z.sub_add.
- apply Z.le_trans with (2 := Z.lt_le_incl _ _ HH2).
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- rewrite Z.pow_1_r; auto with zarith.
- - apply Z.pow_le_mono_r; auto with zarith.
- case (Z.le_gt_cases (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6.
- absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith.
- rewrite <- HH; rewrite Z.mul_1_r.
- apply Z.pow_le_mono_r; auto with zarith. }
- rewrite (Z.mul_comm 2).
- rewrite Z.pow_mul_r; auto with zarith.
- rewrite Z.pow_2_r.
- apply Z.lt_le_trans with (2 := HH3).
- rewrite <- Z.mul_assoc.
- replace (2 * Zpos (digits x) - 1) with
- ((Zpos (digits x) - 1) + (Zpos (digits x))).
- rewrite Zpower_exp; auto with zarith.
- apply Zmult_lt_compat2; auto with zarith.
- split; auto with zarith.
- apply Z.mul_pos_pos; auto with zarith.
- rewrite Pos2Z.inj_xO; ring.
- apply Z.lt_le_incl; auto.
- repeat rewrite spec_head00; auto.
- rewrite spec_double_size_digits.
- rewrite Pos2Z.inj_xO; auto with zarith.
- rewrite spec_double_size; auto.
- Qed.
-
- Theorem spec_double_size_head0_pos:
- forall x, 0 < [head0 (double_size x)].
- Proof.
- intros x.
- assert (F := Pos2Z.is_pos (digits x)).
- assert (F0 := spec_pos (head0 (double_size x))).
- Z.le_elim F0; auto.
- assert (F1 := spec_pos (head0 x)).
- Z.le_elim F1.
- apply Z.lt_le_trans with (2 := (spec_double_size_head0 x)); auto with zarith.
- assert (F3 := spec_pos x).
- Z.le_elim F3.
- generalize F3; rewrite <- (spec_double_size x); intros F4.
- absurd (2 ^ (Zpos (xO (digits x)) - 1) < 2 ^ (Zpos (digits x))).
- { apply Z.le_ngt.
- apply Z.pow_le_mono_r; auto with zarith.
- rewrite Pos2Z.inj_xO; auto with zarith. }
- case (spec_head0 x F3).
- rewrite <- F1; rewrite Z.pow_0_r; rewrite Z.mul_1_l; intros _ HH.
- apply Z.le_lt_trans with (2 := HH).
- case (spec_head0 _ F4).
- rewrite (spec_double_size x); rewrite (spec_double_size_digits x).
- rewrite <- F0; rewrite Z.pow_0_r; rewrite Z.mul_1_l; auto.
- generalize F1; rewrite (spec_head00 _ (eq_sym F3)); auto with zarith.
- Qed.
-
- (** Finally we iterate [double_size] enough before [unsafe_shiftl]
- in order to get a fully correct [shiftl]. *)
-
- Definition shiftl_aux_body cont x n :=
- match compare n (head0 x) with
- Gt => cont (double_size x) n
- | _ => unsafe_shiftl x n
- end.
-
- Theorem spec_shiftl_aux_body: forall n x p cont,
- 2^ Zpos p <= [head0 x] ->
- (forall x, 2 ^ (Zpos p + 1) <= [head0 x]->
- [cont x n] = [x] * 2 ^ [n]) ->
- [shiftl_aux_body cont x n] = [x] * 2 ^ [n].
- Proof.
- intros n x p cont H1 H2; unfold shiftl_aux_body.
- rewrite spec_compare; case Z.compare_spec; intros H.
- apply spec_unsafe_shiftl; auto with zarith.
- apply spec_unsafe_shiftl; auto with zarith.
- rewrite H2.
- rewrite spec_double_size; auto.
- rewrite Z.add_comm; rewrite Zpower_exp; auto with zarith.
- apply Z.le_trans with (2 := spec_double_size_head0 x).
- rewrite Z.pow_1_r; apply Z.mul_le_mono_nonneg_l; auto with zarith.
- Qed.
-
- Fixpoint shiftl_aux p cont x n :=
- shiftl_aux_body
- (fun x n => match p with
- | xH => cont x n
- | xO p => shiftl_aux p (shiftl_aux p cont) x n
- | xI p => shiftl_aux p (shiftl_aux p cont) x n
- end) x n.
-
- Theorem spec_shiftl_aux: forall p q x n cont,
- 2 ^ (Zpos q) <= [head0 x] ->
- (forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->
- [cont x n] = [x] * 2 ^ [n]) ->
- [shiftl_aux p cont x n] = [x] * 2 ^ [n].
- Proof.
- intros p; elim p; unfold shiftl_aux; fold shiftl_aux; clear p.
- intros p Hrec q x n cont H1 H2.
- apply spec_shiftl_aux_body with (q); auto.
- intros x1 H3; apply Hrec with (q + 1)%positive; auto.
- intros x2 H4; apply Hrec with (p + q + 1)%positive; auto.
- rewrite <- Pos.add_assoc.
- rewrite Pos2Z.inj_add; auto.
- intros x3 H5; apply H2.
- rewrite Pos2Z.inj_xI.
- replace (2 * Zpos p + 1 + Zpos q) with (Zpos p + Zpos (p + q + 1));
- auto.
- rewrite !Pos2Z.inj_add; ring.
- intros p Hrec q n x cont H1 H2.
- apply spec_shiftl_aux_body with (q); auto.
- intros x1 H3; apply Hrec with (q); auto.
- apply Z.le_trans with (2 := H3); auto with zarith.
- apply Z.pow_le_mono_r; auto with zarith.
- intros x2 H4; apply Hrec with (p + q)%positive; auto.
- intros x3 H5; apply H2.
- rewrite (Pos2Z.inj_xO p).
- replace (2 * Zpos p + Zpos q) with (Zpos p + Zpos (p + q));
- auto.
- rewrite Pos2Z.inj_add; ring.
- intros q n x cont H1 H2.
- apply spec_shiftl_aux_body with (q); auto.
- rewrite Z.add_comm; auto.
- Qed.
-
- Definition shiftl x n :=
- shiftl_aux_body
- (shiftl_aux_body
- (shiftl_aux (digits n) unsafe_shiftl)) x n.
-
- Theorem spec_shiftl_pow2 : forall x n,
- [shiftl x n] = [x] * 2 ^ [n].
- Proof.
- intros x n; unfold shiftl, shiftl_aux_body.
- rewrite spec_compare; case Z.compare_spec; intros H.
- apply spec_unsafe_shiftl; auto with zarith.
- apply spec_unsafe_shiftl; auto with zarith.
- rewrite <- (spec_double_size x).
- rewrite spec_compare; case Z.compare_spec; intros H1.
- apply spec_unsafe_shiftl; auto with zarith.
- apply spec_unsafe_shiftl; auto with zarith.
- rewrite <- (spec_double_size (double_size x)).
- apply spec_shiftl_aux with 1%positive.
- apply Z.le_trans with (2 := spec_double_size_head0 (double_size x)).
- replace (2 ^ 1) with (2 * 1).
- apply Z.mul_le_mono_nonneg_l; auto with zarith.
- generalize (spec_double_size_head0_pos x); auto with zarith.
- rewrite Z.pow_1_r; ring.
- intros x1 H2; apply spec_unsafe_shiftl.
- apply Z.le_trans with (2 := H2).
- apply Z.le_trans with (2 ^ Zpos (digits n)); auto with zarith.
- case (spec_digits n); auto with zarith.
- apply Z.pow_le_mono_r; auto with zarith.
- Qed.
-
- Lemma spec_shiftl: forall x p, [shiftl x p] = Z.shiftl [x] [p].
- Proof.
- intros.
- now rewrite spec_shiftl_pow2, Z.shiftl_mul_pow2 by apply spec_pos.
- Qed.
-
- (** Other bitwise operations *)
-
- Definition testbit x n := odd (shiftr x n).
-
- Lemma spec_testbit: forall x p, testbit x p = Z.testbit [x] [p].
- Proof.
- intros. unfold testbit. symmetry.
- rewrite spec_odd, spec_shiftr. apply Z.testbit_odd.
- Qed.
-
- Definition div2 x := shiftr x one.
-
- Lemma spec_div2: forall x, [div2 x] = Z.div2 [x].
- Proof.
- intros. unfold div2. symmetry.
- rewrite spec_shiftr, spec_1. apply Z.div2_spec.
- Qed.
-
- Local Notation lorn := (fun n =>
- let op := dom_op n in
- let lor := ZnZ.lor in
- fun x y => reduce n (lor x y)).
-
- Definition lor : t -> t -> t := Eval red_t in same_level lorn.
-
- Lemma lor_fold : lor = same_level lorn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_lor x y : [lor x y] = Z.lor [x] [y].
- Proof.
- rewrite lor_fold. apply spec_same_level; clear x y.
- intros n x y. simpl. rewrite spec_reduce. apply ZnZ.spec_lor.
- Qed.
-
- Local Notation landn := (fun n =>
- let op := dom_op n in
- let land := ZnZ.land in
- fun x y => reduce n (land x y)).
-
- Definition land : t -> t -> t := Eval red_t in same_level landn.
-
- Lemma land_fold : land = same_level landn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_land x y : [land x y] = Z.land [x] [y].
- Proof.
- rewrite land_fold. apply spec_same_level; clear x y.
- intros n x y. simpl. rewrite spec_reduce. apply ZnZ.spec_land.
- Qed.
-
- Local Notation lxorn := (fun n =>
- let op := dom_op n in
- let lxor := ZnZ.lxor in
- fun x y => reduce n (lxor x y)).
-
- Definition lxor : t -> t -> t := Eval red_t in same_level lxorn.
-
- Lemma lxor_fold : lxor = same_level lxorn.
- Proof. red_t; reflexivity. Qed.
-
- Theorem spec_lxor x y : [lxor x y] = Z.lxor [x] [y].
- Proof.
- rewrite lxor_fold. apply spec_same_level; clear x y.
- intros n x y. simpl. rewrite spec_reduce. apply ZnZ.spec_lxor.
- Qed.
-
- Local Notation ldiffn := (fun n =>
- let op := dom_op n in
- let lxor := ZnZ.lxor in
- let land := ZnZ.land in
- let m1 := ZnZ.minus_one in
- fun x y => reduce n (land x (lxor y m1))).
-
- Definition ldiff : t -> t -> t := Eval red_t in same_level ldiffn.
-
- Lemma ldiff_fold : ldiff = same_level ldiffn.
- Proof. red_t; reflexivity. Qed.
-
- Lemma ldiff_alt x y p :
- 0 <= x < 2^p -> 0 <= y < 2^p ->
- Z.ldiff x y = Z.land x (Z.lxor y (2^p - 1)).
- Proof.
- intros (Hx,Hx') (Hy,Hy').
- destruct p as [|p|p].
- - simpl in *; replace x with 0; replace y with 0; auto with zarith.
- - rewrite <- Z.shiftl_1_l. change (_ - 1) with (Z.ones (Z.pos p)).
- rewrite <- Z.ldiff_ones_l_low; trivial.
- rewrite !Z.ldiff_land, Z.land_assoc. f_equal.
- rewrite Z.land_ones; try easy.
- symmetry. apply Z.mod_small; now split.
- Z.le_elim Hy.
- + now apply Z.log2_lt_pow2.
- + now subst.
- - simpl in *; omega.
- Qed.
-
- Theorem spec_ldiff x y : [ldiff x y] = Z.ldiff [x] [y].
- Proof.
- rewrite ldiff_fold. apply spec_same_level; clear x y.
- intros n x y. simpl. rewrite spec_reduce.
- rewrite ZnZ.spec_land, ZnZ.spec_lxor, ZnZ.spec_m1.
- symmetry. apply ldiff_alt; apply ZnZ.spec_to_Z.
- Qed.
-
-End Make.
diff --git a/theories/Numbers/Natural/BigN/NMake_gen.ml b/theories/Numbers/Natural/BigN/NMake_gen.ml
deleted file mode 100644
index 5177fae65..000000000
--- a/theories/Numbers/Natural/BigN/NMake_gen.ml
+++ /dev/null
@@ -1,1017 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-(*S NMake_gen.ml : this file generates NMake_gen.v *)
-
-
-(*s The parameter that control the generation: *)
-
-let size = 6 (* how many times should we repeat the Z/nZ --> Z/2nZ
- process before relying on a generic construct *)
-
-(*s Some utilities *)
-
-let rec iter_str n s = if n = 0 then "" else (iter_str (n-1) s) ^ s
-
-let rec iter_str_gen n f = if n < 0 then "" else (iter_str_gen (n-1) f) ^ (f n)
-
-let rec iter_name i j base sep =
- if i >= j then base^(string_of_int i)
- else (iter_name i (j-1) base sep)^sep^" "^base^(string_of_int j)
-
-let pr s = Printf.printf (s^^"\n")
-
-(*s The actual printing *)
-
-let _ =
-
-pr
-"(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
-(* \\VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-(** * NMake_gen *)
-
-(** From a cyclic Z/nZ representation to arbitrary precision natural numbers.*)
-
-(** Remark: File automatically generated by NMake_gen.ml, DO NOT EDIT ! *)
-
-Require Import BigNumPrelude ZArith Ndigits CyclicAxioms
- DoubleType DoubleMul DoubleDivn1 DoubleCyclic Nbasic
- Wf_nat StreamMemo.
-
-Module Make (W0:CyclicType) <: NAbstract.
-
- (** * The word types *)
-";
-
-pr " Local Notation w0 := W0.t.";
-for i = 1 to size do
- pr " Definition w%i := zn2z w%i." i (i-1)
-done;
-pr "";
-
-pr " (** * The operation type classes for the word types *)
-";
-
-pr " Local Notation w0_op := W0.ops.";
-for i = 1 to min 3 size do
- pr " Instance w%i_op : ZnZ.Ops w%i := mk_zn2z_ops w%i_op." i i (i-1)
-done;
-for i = 4 to size do
- pr " Instance w%i_op : ZnZ.Ops w%i := mk_zn2z_ops_karatsuba w%i_op." i i (i-1)
-done;
-for i = size+1 to size+3 do
- pr " Instance w%i_op : ZnZ.Ops (word w%i %i) := mk_zn2z_ops_karatsuba w%i_op." i size (i-size) (i-1)
-done;
-pr "";
-
- pr " Section Make_op.";
- pr " Variable mk : forall w', ZnZ.Ops w' -> ZnZ.Ops (zn2z w').";
- pr "";
- pr " Fixpoint make_op_aux (n:nat) : ZnZ.Ops (word w%i (S n)):=" size;
- pr " match n return ZnZ.Ops (word w%i (S n)) with" size;
- pr " | O => w%i_op" (size+1);
- pr " | S n1 =>";
- pr " match n1 return ZnZ.Ops (word w%i (S (S n1))) with" size;
- pr " | O => w%i_op" (size+2);
- pr " | S n2 =>";
- pr " match n2 return ZnZ.Ops (word w%i (S (S (S n2)))) with" size;
- pr " | O => w%i_op" (size+3);
- pr " | S n3 => mk _ (mk _ (mk _ (make_op_aux n3)))";
- pr " end";
- pr " end";
- pr " end.";
- pr "";
- pr " End Make_op.";
- pr "";
- pr " Definition omake_op := make_op_aux mk_zn2z_ops_karatsuba.";
- pr "";
- pr "";
- pr " Definition make_op_list := dmemo_list _ omake_op.";
- pr "";
- pr " Instance make_op n : ZnZ.Ops (word w%i (S n))" size;
- pr " := dmemo_get _ omake_op n make_op_list.";
- pr "";
-
-pr " Ltac unfold_ops := unfold omake_op, make_op_aux, w%i_op, w%i_op." (size+3) (size+2);
-
-pr
-"
- Lemma make_op_omake: forall n, make_op n = omake_op n.
- Proof.
- intros n; unfold make_op, make_op_list.
- refine (dmemo_get_correct _ _ _).
- Qed.
-
- Theorem make_op_S: forall n,
- make_op (S n) = mk_zn2z_ops_karatsuba (make_op n).
- Proof.
- intros n. do 2 rewrite make_op_omake.
- revert n. fix IHn 1.
- do 3 (destruct n; [unfold_ops; reflexivity|]).
- simpl mk_zn2z_ops_karatsuba. simpl word in *.
- rewrite <- (IHn n). auto.
- Qed.
-
- (** * The main type [t], isomorphic with [exists n, word w0 n] *)
-";
-
- pr " Inductive t' :=";
- for i = 0 to size do
- pr " | N%i : w%i -> t'" i i
- done;
- pr " | Nn : forall n, word w%i (S n) -> t'." size;
- pr "";
- pr " Definition t := t'.";
- pr "";
-
- pr " (** * A generic toolbox for building and deconstructing [t] *)";
- pr "";
-
- pr " Local Notation SizePlus n := %sn%s."
- (iter_str size "(S ") (iter_str size ")");
- pr " Local Notation Size := (SizePlus O).";
- pr "";
-
- pr " Tactic Notation (at level 3) \"do_size\" tactic3(t) := do %i t." (size+1);
- pr "";
-
- pr " Definition dom_t n := match n with";
- for i = 0 to size do
- pr " | %i => w%i" i i;
- done;
- pr " | %sn => word w%i n" (if size=0 then "" else "SizePlus ") size;
- pr " end.";
- pr "";
-
-pr
-" Instance dom_op n : ZnZ.Ops (dom_t n) | 10.
- Proof.
- do_size (destruct n; [simpl;auto with *|]).
- unfold dom_t. auto with *.
- Defined.
-";
-
- pr " Definition iter_t {A:Type}(f : forall n, dom_t n -> A) : t -> A :=";
- for i = 0 to size do
- pr " let f%i := f %i in" i i;
- done;
- pr " let fn n := f (SizePlus (S n)) in";
- pr " fun x => match x with";
- for i = 0 to size do
- pr " | N%i wx => f%i wx" i i;
- done;
- pr " | Nn n wx => fn n wx";
- pr " end.";
- pr "";
-
- pr " Definition mk_t (n:nat) : dom_t n -> t :=";
- pr " match n as n' return dom_t n' -> t with";
- for i = 0 to size do
- pr " | %i => N%i" i i;
- done;
- pr " | %s(S n) => Nn n" (if size=0 then "" else "SizePlus ");
- pr " end.";
- pr "";
-
-pr
-" Definition level := iter_t (fun n _ => n).
-
- Inductive View_t : t -> Prop :=
- Mk_t : forall n (x : dom_t n), View_t (mk_t n x).
-
- Lemma destr_t : forall x, View_t x.
- Proof.
- intros x. generalize (Mk_t (level x)). destruct x; simpl; auto.
- Defined.
-
- Lemma iter_mk_t : forall A (f:forall n, dom_t n -> A),
- forall n x, iter_t f (mk_t n x) = f n x.
- Proof.
- do_size (destruct n; try reflexivity).
- Qed.
-
- (** * Projection to ZArith *)
-
- Definition to_Z : t -> Z :=
- Eval lazy beta iota delta [iter_t dom_t dom_op] in
- iter_t (fun _ x => ZnZ.to_Z x).
-
- Notation \"[ x ]\" := (to_Z x).
-
- Theorem spec_mk_t : forall n (x:dom_t n), [mk_t n x] = ZnZ.to_Z x.
- Proof.
- intros. change to_Z with (iter_t (fun _ x => ZnZ.to_Z x)).
- rewrite iter_mk_t; auto.
- Qed.
-
- (** * Regular make op, without memoization or karatsuba
-
- This will normally never be used for actual computations,
- but only for specification purpose when using
- [word (dom_t n) m] intermediate values. *)
-
- Fixpoint nmake_op (ww:Type) (ww_op: ZnZ.Ops ww) (n: nat) :
- ZnZ.Ops (word ww n) :=
- match n return ZnZ.Ops (word ww n) with
- O => ww_op
- | S n1 => mk_zn2z_ops (nmake_op ww ww_op n1)
- end.
-
- Definition eval n m := ZnZ.to_Z (Ops:=nmake_op _ (dom_op n) m).
-
- Theorem nmake_op_S: forall ww (w_op: ZnZ.Ops ww) x,
- nmake_op _ w_op (S x) = mk_zn2z_ops (nmake_op _ w_op x).
- Proof.
- auto.
- Qed.
-
- Theorem digits_nmake_S :forall n ww (w_op: ZnZ.Ops ww),
- ZnZ.digits (nmake_op _ w_op (S n)) =
- xO (ZnZ.digits (nmake_op _ w_op n)).
- Proof.
- auto.
- Qed.
-
- Theorem digits_nmake : forall n ww (w_op: ZnZ.Ops ww),
- ZnZ.digits (nmake_op _ w_op n) = Pos.shiftl_nat (ZnZ.digits w_op) n.
- Proof.
- induction n. auto.
- intros ww ww_op. rewrite Pshiftl_nat_S, <- IHn; auto.
- Qed.
-
- Theorem nmake_double: forall n ww (w_op: ZnZ.Ops ww),
- ZnZ.to_Z (Ops:=nmake_op _ w_op n) =
- @DoubleBase.double_to_Z _ (ZnZ.digits w_op) (ZnZ.to_Z (Ops:=w_op)) n.
- Proof.
- intros n; elim n; auto; clear n.
- intros n Hrec ww ww_op; simpl DoubleBase.double_to_Z; unfold zn2z_to_Z.
- rewrite <- Hrec; auto.
- unfold DoubleBase.double_wB; rewrite <- digits_nmake; auto.
- Qed.
-
- Theorem nmake_WW: forall ww ww_op n xh xl,
- (ZnZ.to_Z (Ops:=nmake_op ww ww_op (S n)) (WW xh xl) =
- ZnZ.to_Z (Ops:=nmake_op ww ww_op n) xh *
- base (ZnZ.digits (nmake_op ww ww_op n)) +
- ZnZ.to_Z (Ops:=nmake_op ww ww_op n) xl)%%Z.
- Proof.
- auto.
- Qed.
-
- (** * The specification proofs for the word operators *)
-";
-
- if size <> 0 then
- pr " Typeclasses Opaque %s." (iter_name 1 size "w" "");
- pr "";
-
- pr " Instance w0_spec: ZnZ.Specs w0_op := W0.specs.";
- for i = 1 to min 3 size do
- pr " Instance w%i_spec: ZnZ.Specs w%i_op := mk_zn2z_specs w%i_spec." i i (i-1)
- done;
- for i = 4 to size do
- pr " Instance w%i_spec: ZnZ.Specs w%i_op := mk_zn2z_specs_karatsuba w%i_spec." i i (i-1)
- done;
- pr " Instance w%i_spec: ZnZ.Specs w%i_op := mk_zn2z_specs_karatsuba w%i_spec." (size+1) (size+1) size;
-
-
-pr "
- Instance wn_spec (n:nat) : ZnZ.Specs (make_op n).
- Proof.
- induction n.
- rewrite make_op_omake; simpl; auto with *.
- rewrite make_op_S. exact (mk_zn2z_specs_karatsuba IHn).
- Qed.
-
- Instance dom_spec n : ZnZ.Specs (dom_op n) | 10.
- Proof.
- do_size (destruct n; auto with *). apply wn_spec.
- Qed.
-
- Let make_op_WW : forall n x y,
- (ZnZ.to_Z (Ops:=make_op (S n)) (WW x y) =
- ZnZ.to_Z (Ops:=make_op n) x * base (ZnZ.digits (make_op n))
- + ZnZ.to_Z (Ops:=make_op n) y)%%Z.
- Proof.
- intros n x y; rewrite make_op_S; auto.
- Qed.
-
- (** * Zero *)
-
- Definition zero0 : w0 := ZnZ.zero.
-
- Definition zeron n : dom_t n :=
- match n with
- | O => zero0
- | SizePlus (S n) => W0
- | _ => W0
- end.
-
- Lemma spec_zeron : forall n, ZnZ.to_Z (zeron n) = 0%%Z.
- Proof.
- do_size (destruct n;
- [match goal with
- |- @eq Z (_ (zeron ?n)) _ =>
- apply (ZnZ.spec_0 (Specs:=dom_spec n))
- end|]).
- destruct n; auto. simpl. rewrite make_op_S. fold word.
- apply (ZnZ.spec_0 (Specs:=wn_spec (SizePlus 0))).
- Qed.
-
- (** * Digits *)
-
- Lemma digits_make_op_0 : forall n,
- ZnZ.digits (make_op n) = Pos.shiftl_nat (ZnZ.digits (dom_op Size)) (S n).
- Proof.
- induction n.
- auto.
- replace (ZnZ.digits (make_op (S n))) with (xO (ZnZ.digits (make_op n))).
- rewrite IHn; auto.
- rewrite make_op_S; auto.
- Qed.
-
- Lemma digits_make_op : forall n,
- ZnZ.digits (make_op n) = Pos.shiftl_nat (ZnZ.digits w0_op) (SizePlus (S n)).
- Proof.
- intros. rewrite digits_make_op_0.
- replace (SizePlus (S n)) with (S n + Size) by (rewrite <- plus_comm; auto).
- rewrite Pshiftl_nat_plus. auto.
- Qed.
-
- Lemma digits_dom_op : forall n,
- ZnZ.digits (dom_op n) = Pos.shiftl_nat (ZnZ.digits w0_op) n.
- Proof.
- do_size (destruct n; try reflexivity).
- exact (digits_make_op n).
- Qed.
-
- Lemma digits_dom_op_nmake : forall n m,
- ZnZ.digits (dom_op (m+n)) = ZnZ.digits (nmake_op _ (dom_op n) m).
- Proof.
- intros. rewrite digits_nmake, 2 digits_dom_op. apply Pshiftl_nat_plus.
- Qed.
-
- (** * Conversion between [zn2z (dom_t n)] and [dom_t (S n)].
-
- These two types are provably equal, but not convertible,
- hence we need some work. We now avoid using generic casts
- (i.e. rewrite via proof of equalities in types), since
- proving things with them is a mess.
- *)
-
- Definition succ_t n : zn2z (dom_t n) -> dom_t (S n) :=
- match n with
- | SizePlus (S _) => fun x => x
- | _ => fun x => x
- end.
-
- Lemma spec_succ_t : forall n x,
- ZnZ.to_Z (succ_t n x) =
- zn2z_to_Z (base (ZnZ.digits (dom_op n))) ZnZ.to_Z x.
- Proof.
- do_size (destruct n ; [reflexivity|]).
- intros. simpl. rewrite make_op_S. simpl. auto.
- Qed.
-
- Definition pred_t n : dom_t (S n) -> zn2z (dom_t n) :=
- match n with
- | SizePlus (S _) => fun x => x
- | _ => fun x => x
- end.
-
- Lemma succ_pred_t : forall n x, succ_t n (pred_t n x) = x.
- Proof.
- do_size (destruct n ; [reflexivity|]). reflexivity.
- Qed.
-
- (** We can hence project from [zn2z (dom_t n)] to [t] : *)
-
- Definition mk_t_S n (x : zn2z (dom_t n)) : t :=
- mk_t (S n) (succ_t n x).
-
- Lemma spec_mk_t_S : forall n x,
- [mk_t_S n x] = zn2z_to_Z (base (ZnZ.digits (dom_op n))) ZnZ.to_Z x.
- Proof.
- intros. unfold mk_t_S. rewrite spec_mk_t. apply spec_succ_t.
- Qed.
-
- Lemma mk_t_S_level : forall n x, level (mk_t_S n x) = S n.
- Proof.
- intros. unfold mk_t_S, level. rewrite iter_mk_t; auto.
- Qed.
-
- (** * Conversion from [word (dom_t n) m] to [dom_t (m+n)].
-
- Things are more complex here. We start with a naive version
- that breaks zn2z-trees and reconstruct them. Doing this is
- quite unfortunate, but I don't know how to fully avoid that.
- (cast someday ?). Then we build an optimized version where
- all basic cases (n<=6 or m<=7) are nicely handled.
- *)
-
- Definition zn2z_map {A} {B} (f:A->B) (x:zn2z A) : zn2z B :=
- match x with
- | W0 => W0
- | WW h l => WW (f h) (f l)
- end.
-
- Lemma zn2z_map_id : forall A f (x:zn2z A), (forall u, f u = u) ->
- zn2z_map f x = x.
- Proof.
- destruct x; auto; intros.
- simpl; f_equal; auto.
- Qed.
-
- (** The naive version *)
-
- Fixpoint plus_t n m : word (dom_t n) m -> dom_t (m+n) :=
- match m as m' return word (dom_t n) m' -> dom_t (m'+n) with
- | O => fun x => x
- | S m => fun x => succ_t _ (zn2z_map (plus_t n m) x)
- end.
-
- Theorem spec_plus_t : forall n m (x:word (dom_t n) m),
- ZnZ.to_Z (plus_t n m x) = eval n m x.
- Proof.
- unfold eval.
- induction m.
- simpl; auto.
- intros.
- simpl plus_t; simpl plus. rewrite spec_succ_t.
- destruct x.
- simpl; auto.
- fold word in w, w0.
- simpl. rewrite 2 IHm. f_equal. f_equal. f_equal.
- apply digits_dom_op_nmake.
- Qed.
-
- Definition mk_t_w n m (x:word (dom_t n) m) : t :=
- mk_t (m+n) (plus_t n m x).
-
- Theorem spec_mk_t_w : forall n m (x:word (dom_t n) m),
- [mk_t_w n m x] = eval n m x.
- Proof.
- intros. unfold mk_t_w. rewrite spec_mk_t. apply spec_plus_t.
- Qed.
-
- (** The optimized version.
-
- NB: the last particular case for m could depend on n,
- but it's simplier to just expand everywhere up to m=7
- (cf [mk_t_w'] later).
- *)
-
- Definition plus_t' n : forall m, word (dom_t n) m -> dom_t (m+n) :=
- match n return (forall m, word (dom_t n) m -> dom_t (m+n)) with
- | SizePlus (S n') as n => plus_t n
- | _ as n =>
- fun m => match m return (word (dom_t n) m -> dom_t (m+n)) with
- | SizePlus (S (S m')) as m => plus_t n m
- | _ => fun x => x
- end
- end.
-
- Lemma plus_t_equiv : forall n m x,
- plus_t' n m x = plus_t n m x.
- Proof.
- (do_size try destruct n); try reflexivity;
- (do_size try destruct m); try destruct m; try reflexivity;
- simpl; symmetry; repeat (intros; apply zn2z_map_id; trivial).
- Qed.
-
- Lemma spec_plus_t' : forall n m x,
- ZnZ.to_Z (plus_t' n m x) = eval n m x.
- Proof.
- intros; rewrite plus_t_equiv. apply spec_plus_t.
- Qed.
-
- (** Particular cases [Nk x] = eval i j x with specific k,i,j
- can be solved by the following tactic *)
-
- Ltac solve_eval :=
- intros; rewrite <- spec_plus_t'; unfold to_Z; simpl dom_op; reflexivity.
-
- (** The last particular case that remains useful *)
-
- Lemma spec_eval_size : forall n x, [Nn n x] = eval Size (S n) x.
- Proof.
- induction n.
- solve_eval.
- destruct x as [ | xh xl ].
- simpl. unfold eval. rewrite make_op_S. rewrite nmake_op_S. auto.
- simpl word in xh, xl |- *.
- unfold to_Z in *. rewrite make_op_WW.
- unfold eval in *. rewrite nmake_WW.
- f_equal; auto.
- f_equal; auto.
- f_equal.
- rewrite <- digits_dom_op_nmake. rewrite plus_comm; auto.
- Qed.
-
- (** An optimized [mk_t_w].
-
- We could say mk_t_w' := mk_t _ (plus_t' n m x)
- (TODO: WHY NOT, BTW ??).
- Instead we directly define functions for all intersting [n],
- reverting to naive [mk_t_w] at places that should normally
- never be used (see [mul] and [div_gt]).
- *)
-";
-
-for i = 0 to size-1 do
-let pattern = (iter_str (size+1-i) "(S ") ^ "_" ^ (iter_str (size+1-i) ")") in
-pr
-" Definition mk_t_%iw m := Eval cbv beta zeta iota delta [ mk_t plus ] in
- match m return word w%i (S m) -> t with
- | %s as p => mk_t_w %i (S p)
- | p => mk_t (%i+p)
- end.
-" i i pattern i (i+1)
-done;
-
-pr
-" Definition mk_t_w' n : forall m, word (dom_t n) (S m) -> t :=
- match n return (forall m, word (dom_t n) (S m) -> t) with";
-for i = 0 to size-1 do pr " | %i => mk_t_%iw" i i done;
-pr
-" | Size => Nn
- | _ as n' => fun m => mk_t_w n' (S m)
- end.
-";
-
-pr
-" Ltac solve_spec_mk_t_w' :=
- rewrite <- spec_plus_t';
- match goal with _ : word (dom_t ?n) ?m |- _ => apply (spec_mk_t (n+m)) end.
-
- Theorem spec_mk_t_w' :
- forall n m x, [mk_t_w' n m x] = eval n (S m) x.
- Proof.
- intros.
- repeat (apply spec_mk_t_w || (destruct n;
- [repeat (apply spec_mk_t_w || (destruct m; [solve_spec_mk_t_w'|]))|])).
- apply spec_eval_size.
- Qed.
-
- (** * Extend : injecting [dom_t n] into [word (dom_t n) (S m)] *)
-
- Definition extend n m (x:dom_t n) : word (dom_t n) (S m) :=
- DoubleBase.extend_aux m (WW (zeron n) x).
-
- Lemma spec_extend : forall n m x,
- [mk_t n x] = eval n (S m) (extend n m x).
- Proof.
- intros. unfold eval, extend.
- rewrite spec_mk_t.
- assert (H : forall (x:dom_t n),
- (ZnZ.to_Z (zeron n) * base (ZnZ.digits (dom_op n)) + ZnZ.to_Z x =
- ZnZ.to_Z x)%%Z).
- clear; intros; rewrite spec_zeron; auto.
- rewrite <- (@DoubleBase.spec_extend _
- (WW (zeron n)) (ZnZ.digits (dom_op n)) ZnZ.to_Z H m x).
- simpl. rewrite digits_nmake, <- nmake_double. auto.
- Qed.
-
- (** A particular case of extend, used in [same_level]:
- [extend_size] is [extend Size] *)
-
- Definition extend_size := DoubleBase.extend (WW (W0:dom_t Size)).
-
- Lemma spec_extend_size : forall n x, [mk_t Size x] = [Nn n (extend_size n x)].
- Proof.
- intros. rewrite spec_eval_size. apply (spec_extend Size n).
- Qed.
-
- (** Misc results about extensions *)
-
- Let spec_extend_WW : forall n x,
- [Nn (S n) (WW W0 x)] = [Nn n x].
- Proof.
- intros n x.
- set (N:=SizePlus (S n)).
- change ([Nn (S n) (extend N 0 x)]=[mk_t N x]).
- rewrite (spec_extend N 0).
- solve_eval.
- Qed.
-
- Let spec_extend_tr: forall m n w,
- [Nn (m + n) (extend_tr w m)] = [Nn n w].
- Proof.
- induction m; auto.
- intros n x; simpl extend_tr.
- simpl plus; rewrite spec_extend_WW; auto.
- Qed.
-
- Let spec_cast_l: forall n m x1,
- [Nn n x1] =
- [Nn (Max.max n m) (castm (diff_r n m) (extend_tr x1 (snd (diff n m))))].
- Proof.
- intros n m x1; case (diff_r n m); simpl castm.
- rewrite spec_extend_tr; auto.
- Qed.
-
- Let spec_cast_r: forall n m x1,
- [Nn m x1] =
- [Nn (Max.max n m) (castm (diff_l n m) (extend_tr x1 (fst (diff n m))))].
- Proof.
- intros n m x1; case (diff_l n m); simpl castm.
- rewrite spec_extend_tr; auto.
- Qed.
-
- Ltac unfold_lets :=
- match goal with
- | h : _ |- _ => unfold h; clear h; unfold_lets
- | _ => idtac
- end.
-
- (** * [same_level]
-
- Generic binary operator construction, by extending the smaller
- argument to the level of the other.
- *)
-
- Section SameLevel.
-
- Variable res: Type.
- Variable P : Z -> Z -> res -> Prop.
- Variable f : forall n, dom_t n -> dom_t n -> res.
- Variable Pf : forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y).
-";
-
-for i = 0 to size do
-pr " Let f%i : w%i -> w%i -> res := f %i." i i i i
-done;
-pr
-" Let fn n := f (SizePlus (S n)).
-
- Let Pf' :
- forall n x y u v, u = [mk_t n x] -> v = [mk_t n y] -> P u v (f n x y).
- Proof.
- intros. subst. rewrite 2 spec_mk_t. apply Pf.
- Qed.
-";
-
-let ext i j s =
- if j <= i then s else Printf.sprintf "(extend %i %i %s)" i (j-i-1) s
-in
-
-pr " Notation same_level_folded := (fun x y => match x, y with";
-for i = 0 to size do
- for j = 0 to size do
- pr " | N%i wx, N%i wy => f%i %s %s" i j (max i j) (ext i j "wx") (ext j i "wy")
- done;
- pr " | N%i wx, Nn m wy => fn m (extend_size m %s) wy" i (ext i size "wx")
-done;
-for i = 0 to size do
- pr " | Nn n wx, N%i wy => fn n wx (extend_size n %s)" i (ext i size "wy")
-done;
-pr
-" | Nn n wx, Nn m wy =>
- let mn := Max.max n m in
- let d := diff n m in
- fn mn
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d)))
- end).
-";
-
-pr
-" Definition same_level := Eval lazy beta iota delta
- [ DoubleBase.extend DoubleBase.extend_aux extend zeron ]
- in same_level_folded.
-
- Lemma spec_same_level_0: forall x y, P [x] [y] (same_level x y).
- Proof.
- change same_level with same_level_folded. unfold_lets.
- destruct x, y; apply Pf'; simpl mk_t; rewrite <- ?spec_extend_size;
- match goal with
- | |- context [ extend ?n ?m _ ] => apply (spec_extend n m)
- | |- context [ castm _ _ ] => apply spec_cast_l || apply spec_cast_r
- | _ => reflexivity
- end.
- Qed.
-
- End SameLevel.
-
- Arguments same_level [res] f x y.
-
- Theorem spec_same_level_dep :
- forall res
- (P : nat -> Z -> Z -> res -> Prop)
- (Pantimon : forall n m z z' r, n <= m -> P m z z' r -> P n z z' r)
- (f : forall n, dom_t n -> dom_t n -> res)
- (Pf: forall n x y, P n (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y)),
- forall x y, P (level x) [x] [y] (same_level f x y).
- Proof.
- intros res P Pantimon f Pf.
- set (f' := fun n x y => (n, f n x y)).
- set (P' := fun z z' r => P (fst r) z z' (snd r)).
- assert (FST : forall x y, level x <= fst (same_level f' x y))
- by (destruct x, y; simpl; omega with * ).
- assert (SND : forall x y, same_level f x y = snd (same_level f' x y))
- by (destruct x, y; reflexivity).
- intros. eapply Pantimon; [eapply FST|].
- rewrite SND. eapply (@spec_same_level_0 _ P' f'); eauto.
- Qed.
-
- (** * [iter]
-
- Generic binary operator construction, by splitting the larger
- argument in blocks and applying the smaller argument to them.
- *)
-
- Section Iter.
-
- Variable res: Type.
- Variable P: Z -> Z -> res -> Prop.
-
- Variable f : forall n, dom_t n -> dom_t n -> res.
- Variable Pf : forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y).
-
- Variable fd : forall n m, dom_t n -> word (dom_t n) (S m) -> res.
- Variable fg : forall n m, word (dom_t n) (S m) -> dom_t n -> res.
- Variable Pfd : forall n m x y, P (ZnZ.to_Z x) (eval n (S m) y) (fd n m x y).
- Variable Pfg : forall n m x y, P (eval n (S m) x) (ZnZ.to_Z y) (fg n m x y).
-
- Variable fnm: forall n m, word (dom_t Size) (S n) -> word (dom_t Size) (S m) -> res.
- Variable Pfnm: forall n m x y, P [Nn n x] [Nn m y] (fnm n m x y).
-
- Let Pf' :
- forall n x y u v, u = [mk_t n x] -> v = [mk_t n y] -> P u v (f n x y).
- Proof.
- intros. subst. rewrite 2 spec_mk_t. apply Pf.
- Qed.
-
- Let Pfd' : forall n m x y u v, u = [mk_t n x] -> v = eval n (S m) y ->
- P u v (fd n m x y).
- Proof.
- intros. subst. rewrite spec_mk_t. apply Pfd.
- Qed.
-
- Let Pfg' : forall n m x y u v, u = eval n (S m) x -> v = [mk_t n y] ->
- P u v (fg n m x y).
- Proof.
- intros. subst. rewrite spec_mk_t. apply Pfg.
- Qed.
-";
-
-for i = 0 to size do
-pr " Let f%i := f %i." i i
-done;
-
-for i = 0 to size do
-pr " Let f%in := fd %i." i i;
-pr " Let fn%i := fg %i." i i;
-done;
-
-pr " Notation iter_folded := (fun x y => match x, y with";
-for i = 0 to size do
- for j = 0 to size do
- pr " | N%i wx, N%i wy => f%s wx wy" i j
- (if i = j then string_of_int i
- else if i < j then string_of_int i ^ "n " ^ string_of_int (j-i-1)
- else "n" ^ string_of_int j ^ " " ^ string_of_int (i-j-1))
- done;
- pr " | N%i wx, Nn m wy => f%in m %s wy" i size (ext i size "wx")
-done;
-for i = 0 to size do
- pr " | Nn n wx, N%i wy => fn%i n wx %s" i size (ext i size "wy")
-done;
-pr
-" | Nn n wx, Nn m wy => fnm n m wx wy
- end).
-";
-
-pr
-" Definition iter := Eval lazy beta iota delta
- [extend DoubleBase.extend DoubleBase.extend_aux zeron]
- in iter_folded.
-
- Lemma spec_iter: forall x y, P [x] [y] (iter x y).
- Proof.
- change iter with iter_folded; unfold_lets.
- destruct x; destruct y; apply Pf' || apply Pfd' || apply Pfg' || apply Pfnm;
- simpl mk_t;
- match goal with
- | |- ?x = ?x => reflexivity
- | |- [Nn _ _] = _ => apply spec_eval_size
- | |- context [extend ?n ?m _] => apply (spec_extend n m)
- | _ => idtac
- end;
- unfold to_Z; rewrite <- spec_plus_t'; simpl dom_op; reflexivity.
- Qed.
-
- End Iter.
-";
-
-pr
-" Definition switch
- (P:nat->Type)%s
- (fn:forall n, P n) n :=
- match n return P n with"
- (iter_str_gen size (fun i -> Printf.sprintf "(f%i:P %i)" i i));
-for i = 0 to size do pr " | %i => f%i" i i done;
-pr
-" | n => fn n
- end.
-";
-
-pr
-" Lemma spec_switch : forall P (f:forall n, P n) n,
- switch P %sf n = f n.
- Proof.
- repeat (destruct n; try reflexivity).
- Qed.
-" (iter_str_gen size (fun i -> Printf.sprintf "(f %i) " i));
-
-pr
-" (** * [iter_sym]
-
- A variant of [iter] for symmetric functions, or pseudo-symmetric
- functions (when f y x can be deduced from f x y).
- *)
-
- Section IterSym.
-
- Variable res: Type.
- Variable P: Z -> Z -> res -> Prop.
-
- Variable f : forall n, dom_t n -> dom_t n -> res.
- Variable Pf : forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y).
-
- Variable fg : forall n m, word (dom_t n) (S m) -> dom_t n -> res.
- Variable Pfg : forall n m x y, P (eval n (S m) x) (ZnZ.to_Z y) (fg n m x y).
-
- Variable fnm: forall n m, word (dom_t Size) (S n) -> word (dom_t Size) (S m) -> res.
- Variable Pfnm: forall n m x y, P [Nn n x] [Nn m y] (fnm n m x y).
-
- Variable opp: res -> res.
- Variable Popp : forall u v r, P u v r -> P v u (opp r).
-";
-
-for i = 0 to size do
-pr " Let f%i := f %i." i i
-done;
-
-for i = 0 to size do
-pr " Let fn%i := fg %i." i i;
-done;
-
-pr " Let f' := switch _ %s f." (iter_name 0 size "f" "");
-pr " Let fg' := switch _ %s fg." (iter_name 0 size "fn" "");
-
-pr
-" Local Notation iter_sym_folded :=
- (iter res f' (fun n m x y => opp (fg' n m y x)) fg' fnm).
-
- Definition iter_sym :=
- Eval lazy beta zeta iota delta [iter f' fg' switch] in iter_sym_folded.
-
- Lemma spec_iter_sym: forall x y, P [x] [y] (iter_sym x y).
- Proof.
- intros. change iter_sym with iter_sym_folded. apply spec_iter; clear x y.
- unfold_lets.
- intros. rewrite spec_switch. auto.
- intros. apply Popp. unfold_lets. rewrite spec_switch; auto.
- intros. unfold_lets. rewrite spec_switch; auto.
- auto.
- Qed.
-
- End IterSym.
-
- (** * Reduction
-
- [reduce] can be used instead of [mk_t], it will choose the
- lowest possible level. NB: We only search and remove leftmost
- W0's via ZnZ.eq0, any non-W0 block ends the process, even
- if its value is 0.
- *)
-
- (** First, a direct version ... *)
-
- Fixpoint red_t n : dom_t n -> t :=
- match n return dom_t n -> t with
- | O => N0
- | S n => fun x =>
- let x' := pred_t n x in
- reduce_n1 _ _ (N0 zero0) ZnZ.eq0 (red_t n) (mk_t_S n) x'
- end.
-
- Lemma spec_red_t : forall n x, [red_t n x] = [mk_t n x].
- Proof.
- induction n.
- reflexivity.
- intros.
- simpl red_t. unfold reduce_n1.
- rewrite <- (succ_pred_t n x) at 2.
- remember (pred_t n x) as x'.
- rewrite spec_mk_t, spec_succ_t.
- destruct x' as [ | xh xl]. simpl. apply ZnZ.spec_0.
- generalize (ZnZ.spec_eq0 xh); case ZnZ.eq0; intros H.
- rewrite IHn, spec_mk_t. simpl. rewrite H; auto.
- apply spec_mk_t_S.
- Qed.
-
- (** ... then a specialized one *)
-";
-
-for i = 0 to size do
-pr " Definition eq0%i := @ZnZ.eq0 _ w%i_op." i i;
-done;
-
-pr "
- Definition reduce_0 := N0.";
-for i = 1 to size do
- pr " Definition reduce_%i :=" i;
- pr " Eval lazy beta iota delta [reduce_n1] in";
- pr " reduce_n1 _ _ (N0 zero0) eq0%i reduce_%i N%i." (i-1) (i-1) i
-done;
-
- pr " Definition reduce_%i :=" (size+1);
- pr " Eval lazy beta iota delta [reduce_n1] in";
- pr " reduce_n1 _ _ (N0 zero0) eq0%i reduce_%i (Nn 0)." size size;
-
- pr " Definition reduce_n n :=";
- pr " Eval lazy beta iota delta [reduce_n] in";
- pr " reduce_n _ _ (N0 zero0) reduce_%i Nn n." (size + 1);
- pr "";
-
-pr " Definition reduce n : dom_t n -> t :=";
-pr " match n with";
-for i = 0 to size do
-pr " | %i => reduce_%i" i i;
-done;
-pr " | %s(S n) => reduce_n n" (if size=0 then "" else "SizePlus ");
-pr " end.";
-pr "";
-
-pr " Ltac unfold_red := unfold reduce, %s." (iter_name 1 size "reduce_" ",");
-pr "";
-for i = 0 to size do
-pr " Declare Equivalent Keys reduce reduce_%i." i;
-done;
-pr " Declare Equivalent Keys reduce_n reduce_%i." (size + 1);
-
-pr "
- Ltac solve_red :=
- let H := fresh in let G := fresh in
- match goal with
- | |- ?P (S ?n) => assert (H:P n) by solve_red
- | _ => idtac
- end;
- intros n G x; destruct (le_lt_eq_dec _ _ G) as [LT|EQ];
- solve [
- apply (H _ (lt_n_Sm_le _ _ LT)) |
- inversion LT |
- subst; change (reduce 0 x = red_t 0 x); reflexivity |
- specialize (H (pred n)); subst; destruct x;
- [|unfold_red; rewrite H; auto]; reflexivity
- ].
-
- Lemma reduce_equiv : forall n x, n <= Size -> reduce n x = red_t n x.
- Proof.
- set (P N := forall n, n <= N -> forall x, reduce n x = red_t n x).
- intros n x H. revert n H x. change (P Size). solve_red.
- Qed.
-
- Lemma spec_reduce_n : forall n x, [reduce_n n x] = [Nn n x].
- Proof.
- assert (H : forall x, reduce_%i x = red_t (SizePlus 1) x).
- destruct x; [|unfold reduce_%i; rewrite (reduce_equiv Size)]; auto.
- induction n.
- intros. rewrite H. apply spec_red_t.
- destruct x as [|xh xl].
- simpl. rewrite make_op_S. exact ZnZ.spec_0.
- fold word in *.
- destruct xh; auto.
- simpl reduce_n.
- rewrite IHn.
- rewrite spec_extend_WW; auto.
- Qed.
-" (size+1) (size+1);
-
-pr
-" Lemma spec_reduce : forall n x, [reduce n x] = ZnZ.to_Z x.
- Proof.
- do_size (destruct n;
- [intros; rewrite reduce_equiv;[apply spec_red_t|auto with arith]|]).
- apply spec_reduce_n.
- Qed.
-
-End Make.
-";
diff --git a/theories/Numbers/Natural/BigN/Nbasic.v b/theories/Numbers/Natural/BigN/Nbasic.v
deleted file mode 100644
index 18d0262c9..000000000
--- a/theories/Numbers/Natural/BigN/Nbasic.v
+++ /dev/null
@@ -1,569 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Require Import ZArith Ndigits.
-Require Import BigNumPrelude.
-Require Import Max.
-Require Import DoubleType.
-Require Import DoubleBase.
-Require Import CyclicAxioms.
-Require Import DoubleCyclic.
-
-Arguments mk_zn2z_ops [t] ops.
-Arguments mk_zn2z_ops_karatsuba [t] ops.
-Arguments mk_zn2z_specs [t ops] specs.
-Arguments mk_zn2z_specs_karatsuba [t ops] specs.
-Arguments ZnZ.digits [t] Ops.
-Arguments ZnZ.zdigits [t] Ops.
-
-Lemma Pshiftl_nat_Zpower : forall n p,
- Zpos (Pos.shiftl_nat p n) = Zpos p * 2 ^ Z.of_nat n.
-Proof.
- intros.
- rewrite Z.mul_comm.
- induction n. simpl; auto.
- transitivity (2 * (2 ^ Z.of_nat n * Zpos p)).
- rewrite <- IHn. auto.
- rewrite Z.mul_assoc.
- rewrite Nat2Z.inj_succ.
- rewrite <- Z.pow_succ_r; auto with zarith.
-Qed.
-
-(* To compute the necessary height *)
-
-Fixpoint plength (p: positive) : positive :=
- match p with
- xH => xH
- | xO p1 => Pos.succ (plength p1)
- | xI p1 => Pos.succ (plength p1)
- end.
-
-Theorem plength_correct: forall p, (Zpos p < 2 ^ Zpos (plength p))%Z.
-assert (F: (forall p, 2 ^ (Zpos (Pos.succ p)) = 2 * 2 ^ Zpos p)%Z).
-intros p; replace (Zpos (Pos.succ p)) with (1 + Zpos p)%Z.
-rewrite Zpower_exp; auto with zarith.
-rewrite Pos2Z.inj_succ; unfold Z.succ; auto with zarith.
-intros p; elim p; simpl plength; auto.
-intros p1 Hp1; rewrite F; repeat rewrite Pos2Z.inj_xI.
-assert (tmp: (forall p, 2 * p = p + p)%Z);
- try repeat rewrite tmp; auto with zarith.
-intros p1 Hp1; rewrite F; rewrite (Pos2Z.inj_xO p1).
-assert (tmp: (forall p, 2 * p = p + p)%Z);
- try repeat rewrite tmp; auto with zarith.
-rewrite Z.pow_1_r; auto with zarith.
-Qed.
-
-Theorem plength_pred_correct: forall p, (Zpos p <= 2 ^ Zpos (plength (Pos.pred p)))%Z.
-intros p; case (Pos.succ_pred_or p); intros H1.
-subst; simpl plength.
-rewrite Z.pow_1_r; auto with zarith.
-pattern p at 1; rewrite <- H1.
-rewrite Pos2Z.inj_succ; unfold Z.succ; auto with zarith.
-generalize (plength_correct (Pos.pred p)); auto with zarith.
-Qed.
-
-Definition Pdiv p q :=
- match Z.div (Zpos p) (Zpos q) with
- Zpos q1 => match (Zpos p) - (Zpos q) * (Zpos q1) with
- Z0 => q1
- | _ => (Pos.succ q1)
- end
- | _ => xH
- end.
-
-Theorem Pdiv_le: forall p q,
- Zpos p <= Zpos q * Zpos (Pdiv p q).
-intros p q.
-unfold Pdiv.
-assert (H1: Zpos q > 0); auto with zarith.
-assert (H1b: Zpos p >= 0); auto with zarith.
-generalize (Z_div_ge0 (Zpos p) (Zpos q) H1 H1b).
-generalize (Z_div_mod_eq (Zpos p) (Zpos q) H1); case Z.div.
- intros HH _; rewrite HH; rewrite Z.mul_0_r; rewrite Z.mul_1_r; simpl.
-case (Z_mod_lt (Zpos p) (Zpos q) H1); auto with zarith.
-intros q1 H2.
-replace (Zpos p - Zpos q * Zpos q1) with (Zpos p mod Zpos q).
- 2: pattern (Zpos p) at 2; rewrite H2; auto with zarith.
-generalize H2 (Z_mod_lt (Zpos p) (Zpos q) H1); clear H2;
- case Z.modulo.
- intros HH _; rewrite HH; auto with zarith.
- intros r1 HH (_,HH1); rewrite HH; rewrite Pos2Z.inj_succ.
- unfold Z.succ; rewrite Z.mul_add_distr_l; auto with zarith.
- intros r1 _ (HH,_); case HH; auto.
-intros q1 HH; rewrite HH.
-unfold Z.ge; simpl Z.compare; intros HH1; case HH1; auto.
-Qed.
-
-Definition is_one p := match p with xH => true | _ => false end.
-
-Theorem is_one_one: forall p, is_one p = true -> p = xH.
-intros p; case p; auto; intros p1 H1; discriminate H1.
-Qed.
-
-Definition get_height digits p :=
- let r := Pdiv p digits in
- if is_one r then xH else Pos.succ (plength (Pos.pred r)).
-
-Theorem get_height_correct:
- forall digits N,
- Zpos N <= Zpos digits * (2 ^ (Zpos (get_height digits N) -1)).
-intros digits N.
-unfold get_height.
-assert (H1 := Pdiv_le N digits).
-case_eq (is_one (Pdiv N digits)); intros H2.
-rewrite (is_one_one _ H2) in H1.
-rewrite Z.mul_1_r in H1.
-change (2^(1-1))%Z with 1; rewrite Z.mul_1_r; auto.
-clear H2.
-apply Z.le_trans with (1 := H1).
-apply Z.mul_le_mono_nonneg_l; auto with zarith.
-rewrite Pos2Z.inj_succ; unfold Z.succ.
-rewrite Z.add_comm; rewrite Z.add_simpl_l.
-apply plength_pred_correct.
-Qed.
-
-Definition zn2z_word_comm : forall w n, zn2z (word w n) = word (zn2z w) n.
- fix zn2z_word_comm 2.
- intros w n; case n.
- reflexivity.
- intros n0;simpl.
- case (zn2z_word_comm w n0).
- reflexivity.
-Defined.
-
-Fixpoint extend (n:nat) {struct n} : forall w:Type, zn2z w -> word w (S n) :=
- match n return forall w:Type, zn2z w -> word w (S n) with
- | O => fun w x => x
- | S m =>
- let aux := extend m in
- fun w x => WW W0 (aux w x)
- end.
-
-Section ExtendMax.
-
-Open Scope nat_scope.
-
-Fixpoint plusnS (n m: nat) {struct n} : (n + S m = S (n + m))%nat :=
- match n return (n + S m = S (n + m))%nat with
- | 0 => eq_refl (S m)
- | S n1 =>
- let v := S (S n1 + m) in
- eq_ind_r (fun n => S n = v) (eq_refl v) (plusnS n1 m)
- end.
-
-Fixpoint plusn0 n : n + 0 = n :=
- match n return (n + 0 = n) with
- | 0 => eq_refl 0
- | S n1 =>
- let v := S n1 in
- eq_ind_r (fun n : nat => S n = v) (eq_refl v) (plusn0 n1)
- end.
-
- Fixpoint diff (m n: nat) {struct m}: nat * nat :=
- match m, n with
- O, n => (O, n)
- | m, O => (m, O)
- | S m1, S n1 => diff m1 n1
- end.
-
-Fixpoint diff_l (m n : nat) {struct m} : fst (diff m n) + n = max m n :=
- match m return fst (diff m n) + n = max m n with
- | 0 =>
- match n return (n = max 0 n) with
- | 0 => eq_refl _
- | S n0 => eq_refl _
- end
- | S m1 =>
- match n return (fst (diff (S m1) n) + n = max (S m1) n)
- with
- | 0 => plusn0 _
- | S n1 =>
- let v := fst (diff m1 n1) + n1 in
- let v1 := fst (diff m1 n1) + S n1 in
- eq_ind v (fun n => v1 = S n)
- (eq_ind v1 (fun n => v1 = n) (eq_refl v1) (S v) (plusnS _ _))
- _ (diff_l _ _)
- end
- end.
-
-Fixpoint diff_r (m n: nat) {struct m}: snd (diff m n) + m = max m n :=
- match m return (snd (diff m n) + m = max m n) with
- | 0 =>
- match n return (snd (diff 0 n) + 0 = max 0 n) with
- | 0 => eq_refl _
- | S _ => plusn0 _
- end
- | S m =>
- match n return (snd (diff (S m) n) + S m = max (S m) n) with
- | 0 => eq_refl (snd (diff (S m) 0) + S m)
- | S n1 =>
- let v := S (max m n1) in
- eq_ind_r (fun n => n = v)
- (eq_ind_r (fun n => S n = v)
- (eq_refl v) (diff_r _ _)) (plusnS _ _)
- end
- end.
-
- Variable w: Type.
-
- Definition castm (m n: nat) (H: m = n) (x: word w (S m)):
- (word w (S n)) :=
- match H in (_ = y) return (word w (S y)) with
- | eq_refl => x
- end.
-
-Variable m: nat.
-Variable v: (word w (S m)).
-
-Fixpoint extend_tr (n : nat) {struct n}: (word w (S (n + m))) :=
- match n return (word w (S (n + m))) with
- | O => v
- | S n1 => WW W0 (extend_tr n1)
- end.
-
-End ExtendMax.
-
-Arguments extend_tr [w m] v n.
-Arguments castm [w m n] H x.
-
-
-
-Section Reduce.
-
- Variable w : Type.
- Variable nT : Type.
- Variable N0 : nT.
- Variable eq0 : w -> bool.
- Variable reduce_n : w -> nT.
- Variable zn2z_to_Nt : zn2z w -> nT.
-
- Definition reduce_n1 (x:zn2z w) :=
- match x with
- | W0 => N0
- | WW xh xl =>
- if eq0 xh then reduce_n xl
- else zn2z_to_Nt x
- end.
-
-End Reduce.
-
-Section ReduceRec.
-
- Variable w : Type.
- Variable nT : Type.
- Variable N0 : nT.
- Variable reduce_1n : zn2z w -> nT.
- Variable c : forall n, word w (S n) -> nT.
-
- Fixpoint reduce_n (n:nat) : word w (S n) -> nT :=
- match n return word w (S n) -> nT with
- | O => reduce_1n
- | S m => fun x =>
- match x with
- | W0 => N0
- | WW xh xl =>
- match xh with
- | W0 => @reduce_n m xl
- | _ => @c (S m) x
- end
- end
- end.
-
-End ReduceRec.
-
-Section CompareRec.
-
- Variable wm w : Type.
- Variable w_0 : w.
- Variable compare : w -> w -> comparison.
- Variable compare0_m : wm -> comparison.
- Variable compare_m : wm -> w -> comparison.
-
- Fixpoint compare0_mn (n:nat) : word wm n -> comparison :=
- match n return word wm n -> comparison with
- | O => compare0_m
- | S m => fun x =>
- match x with
- | W0 => Eq
- | WW xh xl =>
- match compare0_mn m xh with
- | Eq => compare0_mn m xl
- | r => Lt
- end
- end
- end.
-
- Variable wm_base: positive.
- Variable wm_to_Z: wm -> Z.
- Variable w_to_Z: w -> Z.
- Variable w_to_Z_0: w_to_Z w_0 = 0.
- Variable spec_compare0_m: forall x,
- compare0_m x = (w_to_Z w_0 ?= wm_to_Z x).
- Variable wm_to_Z_pos: forall x, 0 <= wm_to_Z x < base wm_base.
-
- Let double_to_Z := double_to_Z wm_base wm_to_Z.
- Let double_wB := double_wB wm_base.
-
- Lemma base_xO: forall n, base (xO n) = (base n)^2.
- Proof.
- intros n1; unfold base.
- rewrite (Pos2Z.inj_xO n1); rewrite Z.mul_comm; rewrite Z.pow_mul_r; auto with zarith.
- Qed.
-
- Let double_to_Z_pos: forall n x, 0 <= double_to_Z n x < double_wB n :=
- (spec_double_to_Z wm_base wm_to_Z wm_to_Z_pos).
-
- Declare Equivalent Keys compare0_mn compare0_m.
-
- Lemma spec_compare0_mn: forall n x,
- compare0_mn n x = (0 ?= double_to_Z n x).
- Proof.
- intros n; elim n; clear n; auto.
- intros x; rewrite spec_compare0_m; rewrite w_to_Z_0; auto.
- intros n Hrec x; case x; unfold compare0_mn; fold compare0_mn; auto.
- fold word in *.
- intros xh xl.
- rewrite 2 Hrec.
- simpl double_to_Z.
- set (wB := DoubleBase.double_wB wm_base n).
- case Z.compare_spec; intros Cmp.
- rewrite <- Cmp. reflexivity.
- symmetry. apply Z.gt_lt, Z.lt_gt. (* ;-) *)
- assert (0 < wB).
- unfold wB, DoubleBase.double_wB, base; auto with zarith.
- change 0 with (0 + 0); apply Z.add_lt_le_mono; auto with zarith.
- apply Z.mul_pos_pos; auto with zarith.
- case (double_to_Z_pos n xl); auto with zarith.
- case (double_to_Z_pos n xh); intros; exfalso; omega.
- Qed.
-
- Fixpoint compare_mn_1 (n:nat) : word wm n -> w -> comparison :=
- match n return word wm n -> w -> comparison with
- | O => compare_m
- | S m => fun x y =>
- match x with
- | W0 => compare w_0 y
- | WW xh xl =>
- match compare0_mn m xh with
- | Eq => compare_mn_1 m xl y
- | r => Gt
- end
- end
- end.
-
- Variable spec_compare: forall x y,
- compare x y = Z.compare (w_to_Z x) (w_to_Z y).
- Variable spec_compare_m: forall x y,
- compare_m x y = Z.compare (wm_to_Z x) (w_to_Z y).
- Variable wm_base_lt: forall x,
- 0 <= w_to_Z x < base (wm_base).
-
- Let double_wB_lt: forall n x,
- 0 <= w_to_Z x < (double_wB n).
- Proof.
- intros n x; elim n; simpl; auto; clear n.
- intros n (H0, H); split; auto.
- apply Z.lt_le_trans with (1:= H).
- unfold double_wB, DoubleBase.double_wB; simpl.
- rewrite base_xO.
- set (u := base (Pos.shiftl_nat wm_base n)).
- assert (0 < u).
- unfold u, base; auto with zarith.
- replace (u^2) with (u * u); simpl; auto with zarith.
- apply Z.le_trans with (1 * u); auto with zarith.
- unfold Z.pow_pos; simpl; ring.
- Qed.
-
-
- Lemma spec_compare_mn_1: forall n x y,
- compare_mn_1 n x y = Z.compare (double_to_Z n x) (w_to_Z y).
- Proof.
- intros n; elim n; simpl; auto; clear n.
- intros n Hrec x; case x; clear x; auto.
- intros y; rewrite spec_compare; rewrite w_to_Z_0. reflexivity.
- intros xh xl y; simpl;
- rewrite spec_compare0_mn, Hrec. case Z.compare_spec.
- intros H1b.
- rewrite <- H1b; rewrite Z.mul_0_l; rewrite Z.add_0_l; auto.
- symmetry. apply Z.lt_gt.
- case (double_wB_lt n y); intros _ H0.
- apply Z.lt_le_trans with (1:= H0).
- fold double_wB.
- case (double_to_Z_pos n xl); intros H1 H2.
- apply Z.le_trans with (double_to_Z n xh * double_wB n); auto with zarith.
- apply Z.le_trans with (1 * double_wB n); auto with zarith.
- case (double_to_Z_pos n xh); intros; exfalso; omega.
- Qed.
-
-End CompareRec.
-
-
-Section AddS.
-
- Variable w wm : Type.
- Variable incr : wm -> carry wm.
- Variable addr : w -> wm -> carry wm.
- Variable injr : w -> zn2z wm.
-
- Variable w_0 u: w.
- Fixpoint injs (n:nat): word w (S n) :=
- match n return (word w (S n)) with
- O => WW w_0 u
- | S n1 => (WW W0 (injs n1))
- end.
-
- Definition adds x y :=
- match y with
- W0 => C0 (injr x)
- | WW hy ly => match addr x ly with
- C0 z => C0 (WW hy z)
- | C1 z => match incr hy with
- C0 z1 => C0 (WW z1 z)
- | C1 z1 => C1 (WW z1 z)
- end
- end
- end.
-
-End AddS.
-
- Fixpoint length_pos x :=
- match x with xH => O | xO x1 => S (length_pos x1) | xI x1 => S (length_pos x1) end.
-
- Theorem length_pos_lt: forall x y,
- (length_pos x < length_pos y)%nat -> Zpos x < Zpos y.
- Proof.
- intros x; elim x; clear x; [intros x1 Hrec | intros x1 Hrec | idtac];
- intros y; case y; clear y; intros y1 H || intros H; simpl length_pos;
- try (rewrite (Pos2Z.inj_xI x1) || rewrite (Pos2Z.inj_xO x1));
- try (rewrite (Pos2Z.inj_xI y1) || rewrite (Pos2Z.inj_xO y1));
- try (inversion H; fail);
- try (assert (Zpos x1 < Zpos y1); [apply Hrec; apply lt_S_n | idtac]; auto with zarith);
- assert (0 < Zpos y1); auto with zarith; red; auto.
- Qed.
-
- Theorem cancel_app: forall A B (f g: A -> B) x, f = g -> f x = g x.
- Proof.
- intros A B f g x H; rewrite H; auto.
- Qed.
-
-
- Section SimplOp.
-
- Variable w: Type.
-
- Theorem digits_zop: forall t (ops : ZnZ.Ops t),
- ZnZ.digits (mk_zn2z_ops ops) = xO (ZnZ.digits ops).
- Proof.
- intros ww x; auto.
- Qed.
-
- Theorem digits_kzop: forall t (ops : ZnZ.Ops t),
- ZnZ.digits (mk_zn2z_ops_karatsuba ops) = xO (ZnZ.digits ops).
- Proof.
- intros ww x; auto.
- Qed.
-
- Theorem make_zop: forall t (ops : ZnZ.Ops t),
- @ZnZ.to_Z _ (mk_zn2z_ops ops) =
- fun z => match z with
- | W0 => 0
- | WW xh xl => ZnZ.to_Z xh * base (ZnZ.digits ops)
- + ZnZ.to_Z xl
- end.
- Proof.
- intros ww x; auto.
- Qed.
-
- Theorem make_kzop: forall t (ops: ZnZ.Ops t),
- @ZnZ.to_Z _ (mk_zn2z_ops_karatsuba ops) =
- fun z => match z with
- | W0 => 0
- | WW xh xl => ZnZ.to_Z xh * base (ZnZ.digits ops)
- + ZnZ.to_Z xl
- end.
- Proof.
- intros ww x; auto.
- Qed.
-
- End SimplOp.
-
-(** Abstract vision of a datatype of arbitrary-large numbers.
- Concrete operations can be derived from these generic
- fonctions, in particular from [iter_t] and [same_level].
-*)
-
-Module Type NAbstract.
-
-(** The domains: a sequence of [Z/nZ] structures *)
-
-Parameter dom_t : nat -> Type.
-Declare Instance dom_op n : ZnZ.Ops (dom_t n).
-Declare Instance dom_spec n : ZnZ.Specs (dom_op n).
-
-Axiom digits_dom_op : forall n,
- ZnZ.digits (dom_op n) = Pos.shiftl_nat (ZnZ.digits (dom_op 0)) n.
-
-(** The type [t] of arbitrary-large numbers, with abstract constructor [mk_t]
- and destructor [destr_t] and iterator [iter_t] *)
-
-Parameter t : Type.
-
-Parameter mk_t : forall (n:nat), dom_t n -> t.
-
-Inductive View_t : t -> Prop :=
- Mk_t : forall n (x : dom_t n), View_t (mk_t n x).
-
-Axiom destr_t : forall x, View_t x. (* i.e. every x is a (mk_t n xw) *)
-
-Parameter iter_t : forall {A:Type}(f : forall n, dom_t n -> A), t -> A.
-
-Axiom iter_mk_t : forall A (f:forall n, dom_t n -> A),
- forall n x, iter_t f (mk_t n x) = f n x.
-
-(** Conversion to [ZArith] *)
-
-Parameter to_Z : t -> Z.
-Local Notation "[ x ]" := (to_Z x).
-
-Axiom spec_mk_t : forall n x, [mk_t n x] = ZnZ.to_Z x.
-
-(** [reduce] is like [mk_t], but try to minimise the level of the number *)
-
-Parameter reduce : forall (n:nat), dom_t n -> t.
-Axiom spec_reduce : forall n x, [reduce n x] = ZnZ.to_Z x.
-
-(** Number of level in the tree representation of a number.
- NB: This function isn't a morphism for setoid [eq]. *)
-
-Definition level := iter_t (fun n _ => n).
-
-(** [same_level] and its rich specification, indexed by [level] *)
-
-Parameter same_level : forall {A:Type}
- (f : forall n, dom_t n -> dom_t n -> A), t -> t -> A.
-
-Axiom spec_same_level_dep :
- forall res
- (P : nat -> Z -> Z -> res -> Prop)
- (Pantimon : forall n m z z' r, (n <= m)%nat -> P m z z' r -> P n z z' r)
- (f : forall n, dom_t n -> dom_t n -> res)
- (Pf: forall n x y, P n (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y)),
- forall x y, P (level x) [x] [y] (same_level f x y).
-
-(** [mk_t_S] : building a number of the next level *)
-
-Parameter mk_t_S : forall (n:nat), zn2z (dom_t n) -> t.
-
-Axiom spec_mk_t_S : forall n (x:zn2z (dom_t n)),
- [mk_t_S n x] = zn2z_to_Z (base (ZnZ.digits (dom_op n))) ZnZ.to_Z x.
-
-Axiom mk_t_S_level : forall n x, level (mk_t_S n x) = S n.
-
-End NAbstract.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSig.v b/theories/Numbers/Natural/SpecViaZ/NSig.v
deleted file mode 100644
index 258e03159..000000000
--- a/theories/Numbers/Natural/SpecViaZ/NSig.v
+++ /dev/null
@@ -1,124 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-
-Require Import BinInt.
-
-Open Scope Z_scope.
-
-(** * NSig *)
-
-(** Interface of a rich structure about natural numbers.
- Specifications are written via translation to Z.
-*)
-
-Module Type NType.
-
- Parameter t : Type.
-
- Parameter to_Z : t -> Z.
- Local Notation "[ x ]" := (to_Z x).
- Parameter spec_pos: forall x, 0 <= [x].
-
- Parameter of_N : N -> t.
- Parameter spec_of_N: forall x, to_Z (of_N x) = Z.of_N x.
- Definition to_N n := Z.to_N (to_Z n).
-
- Definition eq n m := [n] = [m].
- Definition lt n m := [n] < [m].
- Definition le n m := [n] <= [m].
-
- Parameter compare : t -> t -> comparison.
- Parameter eqb : t -> t -> bool.
- Parameter ltb : t -> t -> bool.
- Parameter leb : t -> t -> bool.
- Parameter max : t -> t -> t.
- Parameter min : t -> t -> t.
- Parameter zero : t.
- Parameter one : t.
- Parameter two : t.
- Parameter succ : t -> t.
- Parameter pred : t -> t.
- Parameter add : t -> t -> t.
- Parameter sub : t -> t -> t.
- Parameter mul : t -> t -> t.
- Parameter square : t -> t.
- Parameter pow_pos : t -> positive -> t.
- Parameter pow_N : t -> N -> t.
- Parameter pow : t -> t -> t.
- Parameter sqrt : t -> t.
- Parameter log2 : t -> t.
- Parameter div_eucl : t -> t -> t * t.
- Parameter div : t -> t -> t.
- Parameter modulo : t -> t -> t.
- Parameter gcd : t -> t -> t.
- Parameter even : t -> bool.
- Parameter odd : t -> bool.
- Parameter testbit : t -> t -> bool.
- Parameter shiftr : t -> t -> t.
- Parameter shiftl : t -> t -> t.
- Parameter land : t -> t -> t.
- Parameter lor : t -> t -> t.
- Parameter ldiff : t -> t -> t.
- Parameter lxor : t -> t -> t.
- Parameter div2 : t -> t.
-
- Parameter spec_compare: forall x y, compare x y = ([x] ?= [y]).
- Parameter spec_eqb : forall x y, eqb x y = ([x] =? [y]).
- Parameter spec_ltb : forall x y, ltb x y = ([x] <? [y]).
- Parameter spec_leb : forall x y, leb x y = ([x] <=? [y]).
- Parameter spec_max : forall x y, [max x y] = Z.max [x] [y].
- Parameter spec_min : forall x y, [min x y] = Z.min [x] [y].
- Parameter spec_0: [zero] = 0.
- Parameter spec_1: [one] = 1.
- Parameter spec_2: [two] = 2.
- Parameter spec_succ: forall n, [succ n] = [n] + 1.
- Parameter spec_add: forall x y, [add x y] = [x] + [y].
- Parameter spec_pred: forall x, [pred x] = Z.max 0 ([x] - 1).
- Parameter spec_sub: forall x y, [sub x y] = Z.max 0 ([x] - [y]).
- Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
- Parameter spec_square: forall x, [square x] = [x] * [x].
- Parameter spec_pow_pos: forall x n, [pow_pos x n] = [x] ^ Zpos n.
- Parameter spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z.of_N n.
- Parameter spec_pow: forall x n, [pow x n] = [x] ^ [n].
- Parameter spec_sqrt: forall x, [sqrt x] = Z.sqrt [x].
- Parameter spec_log2: forall x, [log2 x] = Z.log2 [x].
- Parameter spec_div_eucl: forall x y,
- let (q,r) := div_eucl x y in ([q], [r]) = Z.div_eucl [x] [y].
- Parameter spec_div: forall x y, [div x y] = [x] / [y].
- Parameter spec_modulo: forall x y, [modulo x y] = [x] mod [y].
- Parameter spec_gcd: forall a b, [gcd a b] = Z.gcd [a] [b].
- Parameter spec_even: forall x, even x = Z.even [x].
- Parameter spec_odd: forall x, odd x = Z.odd [x].
- Parameter spec_testbit: forall x p, testbit x p = Z.testbit [x] [p].
- Parameter spec_shiftr: forall x p, [shiftr x p] = Z.shiftr [x] [p].
- Parameter spec_shiftl: forall x p, [shiftl x p] = Z.shiftl [x] [p].
- Parameter spec_land: forall x y, [land x y] = Z.land [x] [y].
- Parameter spec_lor: forall x y, [lor x y] = Z.lor [x] [y].
- Parameter spec_ldiff: forall x y, [ldiff x y] = Z.ldiff [x] [y].
- Parameter spec_lxor: forall x y, [lxor x y] = Z.lxor [x] [y].
- Parameter spec_div2: forall x, [div2 x] = Z.div2 [x].
-
-End NType.
-
-Module Type NType_Notation (Import N:NType).
- Notation "[ x ]" := (to_Z x).
- Infix "==" := eq (at level 70).
- Notation "0" := zero.
- Notation "1" := one.
- Notation "2" := two.
- Infix "+" := add.
- Infix "-" := sub.
- Infix "*" := mul.
- Infix "^" := pow.
- Infix "<=" := le.
- Infix "<" := lt.
-End NType_Notation.
-
-Module Type NType' := NType <+ NType_Notation.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
deleted file mode 100644
index 355da4cc6..000000000
--- a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
+++ /dev/null
@@ -1,487 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-Require Import ZArith OrdersFacts Nnat NAxioms NSig.
-
-(** * The interface [NSig.NType] implies the interface [NAxiomsSig] *)
-
-Module NTypeIsNAxioms (Import NN : NType').
-
-Hint Rewrite
- spec_0 spec_1 spec_2 spec_succ spec_add spec_mul spec_pred spec_sub
- spec_div spec_modulo spec_gcd spec_compare spec_eqb spec_ltb spec_leb
- spec_square spec_sqrt spec_log2 spec_max spec_min spec_pow_pos spec_pow_N
- spec_pow spec_even spec_odd spec_testbit spec_shiftl spec_shiftr
- spec_land spec_lor spec_ldiff spec_lxor spec_div2 spec_of_N
- : nsimpl.
-Ltac nsimpl := autorewrite with nsimpl.
-Ltac ncongruence := unfold eq, to_N; repeat red; intros; nsimpl; congruence.
-Ltac zify := unfold eq, lt, le, to_N in *; nsimpl.
-Ltac omega_pos n := generalize (spec_pos n); omega with *.
-
-Local Obligation Tactic := ncongruence.
-
-Instance eq_equiv : Equivalence eq.
-Proof. unfold eq. firstorder. Qed.
-
-Program Instance succ_wd : Proper (eq==>eq) succ.
-Program Instance pred_wd : Proper (eq==>eq) pred.
-Program Instance add_wd : Proper (eq==>eq==>eq) add.
-Program Instance sub_wd : Proper (eq==>eq==>eq) sub.
-Program Instance mul_wd : Proper (eq==>eq==>eq) mul.
-
-Theorem pred_succ : forall n, pred (succ n) == n.
-Proof.
-intros. zify. omega_pos n.
-Qed.
-
-Theorem one_succ : 1 == succ 0.
-Proof.
-now zify.
-Qed.
-
-Theorem two_succ : 2 == succ 1.
-Proof.
-now zify.
-Qed.
-
-Definition N_of_Z z := of_N (Z.to_N z).
-
-Lemma spec_N_of_Z z : (0<=z)%Z -> [N_of_Z z] = z.
-Proof.
- unfold N_of_Z. zify. apply Z2N.id.
-Qed.
-
-Section Induction.
-
-Variable A : NN.t -> Prop.
-Hypothesis A_wd : Proper (eq==>iff) A.
-Hypothesis A0 : A 0.
-Hypothesis AS : forall n, A n <-> A (succ n).
-
-Let B (z : Z) := A (N_of_Z z).
-
-Lemma B0 : B 0.
-Proof.
-unfold B, N_of_Z; simpl.
-rewrite <- (A_wd 0); auto.
-red; rewrite spec_0, spec_of_N; auto.
-Qed.
-
-Lemma BS : forall z : Z, (0 <= z)%Z -> B z -> B (z + 1).
-Proof.
-intros z H1 H2.
-unfold B in *. apply -> AS in H2.
-setoid_replace (N_of_Z (z + 1)) with (succ (N_of_Z z)); auto.
-unfold eq. rewrite spec_succ, 2 spec_N_of_Z; auto with zarith.
-Qed.
-
-Lemma B_holds : forall z : Z, (0 <= z)%Z -> B z.
-Proof.
-exact (natlike_ind B B0 BS).
-Qed.
-
-Theorem bi_induction : forall n, A n.
-Proof.
-intro n. setoid_replace n with (N_of_Z (to_Z n)).
-apply B_holds. apply spec_pos.
-red. now rewrite spec_N_of_Z by apply spec_pos.
-Qed.
-
-End Induction.
-
-Theorem add_0_l : forall n, 0 + n == n.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem add_succ_l : forall n m, (succ n) + m == succ (n + m).
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem sub_0_r : forall n, n - 0 == n.
-Proof.
-intros. zify. omega_pos n.
-Qed.
-
-Theorem sub_succ_r : forall n m, n - (succ m) == pred (n - m).
-Proof.
-intros. zify. omega with *.
-Qed.
-
-Theorem mul_0_l : forall n, 0 * n == 0.
-Proof.
-intros. zify. auto with zarith.
-Qed.
-
-Theorem mul_succ_l : forall n m, (succ n) * m == n * m + m.
-Proof.
-intros. zify. ring.
-Qed.
-
-(** Order *)
-
-Lemma eqb_eq x y : eqb x y = true <-> x == y.
-Proof.
- zify. apply Z.eqb_eq.
-Qed.
-
-Lemma leb_le x y : leb x y = true <-> x <= y.
-Proof.
- zify. apply Z.leb_le.
-Qed.
-
-Lemma ltb_lt x y : ltb x y = true <-> x < y.
-Proof.
- zify. apply Z.ltb_lt.
-Qed.
-
-Lemma compare_eq_iff n m : compare n m = Eq <-> n == m.
-Proof.
- intros. zify. apply Z.compare_eq_iff.
-Qed.
-
-Lemma compare_lt_iff n m : compare n m = Lt <-> n < m.
-Proof.
- intros. zify. reflexivity.
-Qed.
-
-Lemma compare_le_iff n m : compare n m <> Gt <-> n <= m.
-Proof.
- intros. zify. reflexivity.
-Qed.
-
-Lemma compare_antisym n m : compare m n = CompOpp (compare n m).
-Proof.
- intros. zify. apply Z.compare_antisym.
-Qed.
-
-Include BoolOrderFacts NN NN NN [no inline].
-
-Instance compare_wd : Proper (eq ==> eq ==> Logic.eq) compare.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance eqb_wd : Proper (eq ==> eq ==> Logic.eq) eqb.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance ltb_wd : Proper (eq ==> eq ==> Logic.eq) ltb.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance leb_wd : Proper (eq ==> eq ==> Logic.eq) leb.
-Proof.
-intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
-Qed.
-
-Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
-Proof.
-intros x x' Hx y y' Hy; unfold lt; rewrite Hx, Hy; intuition.
-Qed.
-
-Theorem lt_succ_r : forall n m, n < succ m <-> n <= m.
-Proof.
-intros. zify. omega.
-Qed.
-
-Theorem min_l : forall n m, n <= m -> min n m == n.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-Theorem min_r : forall n m, m <= n -> min n m == m.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-Theorem max_l : forall n m, m <= n -> max n m == n.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-Theorem max_r : forall n m, n <= m -> max n m == m.
-Proof.
-intros n m. zify. omega with *.
-Qed.
-
-(** Properties specific to natural numbers, not integers. *)
-
-Theorem pred_0 : pred 0 == 0.
-Proof.
-zify. auto.
-Qed.
-
-(** Power *)
-
-Program Instance pow_wd : Proper (eq==>eq==>eq) pow.
-
-Lemma pow_0_r : forall a, a^0 == 1.
-Proof.
- intros. now zify.
-Qed.
-
-Lemma pow_succ_r : forall a b, 0<=b -> a^(succ b) == a * a^b.
-Proof.
- intros a b. zify. intros. now Z.nzsimpl.
-Qed.
-
-Lemma pow_neg_r : forall a b, b<0 -> a^b == 0.
-Proof.
- intros a b. zify. intro Hb. exfalso. omega_pos b.
-Qed.
-
-Lemma pow_pow_N : forall a b, a^b == pow_N a (to_N b).
-Proof.
- intros. zify. f_equal.
- now rewrite Z2N.id by apply spec_pos.
-Qed.
-
-Lemma pow_N_pow : forall a b, pow_N a b == a^(of_N b).
-Proof.
- intros. now zify.
-Qed.
-
-Lemma pow_pos_N : forall a p, pow_pos a p == pow_N a (Npos p).
-Proof.
- intros. now zify.
-Qed.
-
-(** Square *)
-
-Lemma square_spec n : square n == n * n.
-Proof.
- now zify.
-Qed.
-
-(** Sqrt *)
-
-Lemma sqrt_spec : forall n, 0<=n ->
- (sqrt n)*(sqrt n) <= n /\ n < (succ (sqrt n))*(succ (sqrt n)).
-Proof.
- intros n. zify. apply Z.sqrt_spec.
-Qed.
-
-Lemma sqrt_neg : forall n, n<0 -> sqrt n == 0.
-Proof.
- intros n. zify. intro H. exfalso. omega_pos n.
-Qed.
-
-(** Log2 *)
-
-Lemma log2_spec : forall n, 0<n ->
- 2^(log2 n) <= n /\ n < 2^(succ (log2 n)).
-Proof.
- intros n. zify. change (Z.log2 [n]+1)%Z with (Z.succ (Z.log2 [n])).
- apply Z.log2_spec.
-Qed.
-
-Lemma log2_nonpos : forall n, n<=0 -> log2 n == 0.
-Proof.
- intros n. zify. apply Z.log2_nonpos.
-Qed.
-
-(** Even / Odd *)
-
-Definition Even n := exists m, n == 2*m.
-Definition Odd n := exists m, n == 2*m+1.
-
-Lemma even_spec n : even n = true <-> Even n.
-Proof.
- unfold Even. zify. rewrite Z.even_spec.
- split; intros (m,Hm).
- - exists (N_of_Z m). zify. rewrite spec_N_of_Z; trivial. omega_pos n.
- - exists [m]. revert Hm; now zify.
-Qed.
-
-Lemma odd_spec n : odd n = true <-> Odd n.
-Proof.
- unfold Odd. zify. rewrite Z.odd_spec.
- split; intros (m,Hm).
- - exists (N_of_Z m). zify. rewrite spec_N_of_Z; trivial. omega_pos n.
- - exists [m]. revert Hm; now zify.
-Qed.
-
-(** Div / Mod *)
-
-Program Instance div_wd : Proper (eq==>eq==>eq) div.
-Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
-
-Theorem div_mod : forall a b, ~b==0 -> a == b*(div a b) + (modulo a b).
-Proof.
-intros a b. zify. intros. apply Z.div_mod; auto.
-Qed.
-
-Theorem mod_bound_pos : forall a b, 0<=a -> 0<b ->
- 0 <= modulo a b /\ modulo a b < b.
-Proof.
-intros a b. zify. apply Z.mod_bound_pos.
-Qed.
-
-(** Gcd *)
-
-Definition divide n m := exists p, m == p*n.
-Local Notation "( x | y )" := (divide x y) (at level 0).
-
-Lemma spec_divide : forall n m, (n|m) <-> Z.divide [n] [m].
-Proof.
- intros n m. split.
- - intros (p,H). exists [p]. revert H; now zify.
- - intros (z,H). exists (of_N (Z.abs_N z)). zify.
- rewrite N2Z.inj_abs_N.
- rewrite <- (Z.abs_eq [m]), <- (Z.abs_eq [n]) by apply spec_pos.
- now rewrite H, Z.abs_mul.
-Qed.
-
-Lemma gcd_divide_l : forall n m, (gcd n m | n).
-Proof.
- intros n m. apply spec_divide. zify. apply Z.gcd_divide_l.
-Qed.
-
-Lemma gcd_divide_r : forall n m, (gcd n m | m).
-Proof.
- intros n m. apply spec_divide. zify. apply Z.gcd_divide_r.
-Qed.
-
-Lemma gcd_greatest : forall n m p, (p|n) -> (p|m) -> (p|gcd n m).
-Proof.
- intros n m p. rewrite !spec_divide. zify. apply Z.gcd_greatest.
-Qed.
-
-Lemma gcd_nonneg : forall n m, 0 <= gcd n m.
-Proof.
- intros. zify. apply Z.gcd_nonneg.
-Qed.
-
-(** Bitwise operations *)
-
-Program Instance testbit_wd : Proper (eq==>eq==>Logic.eq) testbit.
-
-Lemma testbit_odd_0 : forall a, testbit (2*a+1) 0 = true.
-Proof.
- intros. zify. apply Z.testbit_odd_0.
-Qed.
-
-Lemma testbit_even_0 : forall a, testbit (2*a) 0 = false.
-Proof.
- intros. zify. apply Z.testbit_even_0.
-Qed.
-
-Lemma testbit_odd_succ : forall a n, 0<=n ->
- testbit (2*a+1) (succ n) = testbit a n.
-Proof.
- intros a n. zify. apply Z.testbit_odd_succ.
-Qed.
-
-Lemma testbit_even_succ : forall a n, 0<=n ->
- testbit (2*a) (succ n) = testbit a n.
-Proof.
- intros a n. zify. apply Z.testbit_even_succ.
-Qed.
-
-Lemma testbit_neg_r : forall a n, n<0 -> testbit a n = false.
-Proof.
- intros a n. zify. apply Z.testbit_neg_r.
-Qed.
-
-Lemma shiftr_spec : forall a n m, 0<=m ->
- testbit (shiftr a n) m = testbit a (m+n).
-Proof.
- intros a n m. zify. apply Z.shiftr_spec.
-Qed.
-
-Lemma shiftl_spec_high : forall a n m, 0<=m -> n<=m ->
- testbit (shiftl a n) m = testbit a (m-n).
-Proof.
- intros a n m. zify. intros Hn H. rewrite Z.max_r by auto with zarith.
- now apply Z.shiftl_spec_high.
-Qed.
-
-Lemma shiftl_spec_low : forall a n m, m<n ->
- testbit (shiftl a n) m = false.
-Proof.
- intros a n m. zify. intros H. now apply Z.shiftl_spec_low.
-Qed.
-
-Lemma land_spec : forall a b n,
- testbit (land a b) n = testbit a n && testbit b n.
-Proof.
- intros a n m. zify. now apply Z.land_spec.
-Qed.
-
-Lemma lor_spec : forall a b n,
- testbit (lor a b) n = testbit a n || testbit b n.
-Proof.
- intros a n m. zify. now apply Z.lor_spec.
-Qed.
-
-Lemma ldiff_spec : forall a b n,
- testbit (ldiff a b) n = testbit a n && negb (testbit b n).
-Proof.
- intros a n m. zify. now apply Z.ldiff_spec.
-Qed.
-
-Lemma lxor_spec : forall a b n,
- testbit (lxor a b) n = xorb (testbit a n) (testbit b n).
-Proof.
- intros a n m. zify. now apply Z.lxor_spec.
-Qed.
-
-Lemma div2_spec : forall a, div2 a == shiftr a 1.
-Proof.
- intros a. zify. now apply Z.div2_spec.
-Qed.
-
-(** Recursion *)
-
-Definition recursion (A : Type) (a : A) (f : NN.t -> A -> A) (n : NN.t) :=
- N.peano_rect (fun _ => A) a (fun n a => f (NN.of_N n) a) (NN.to_N n).
-Arguments recursion [A] a f n.
-
-Instance recursion_wd (A : Type) (Aeq : relation A) :
- Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
-Proof.
-unfold eq.
-intros a a' Eaa' f f' Eff' x x' Exx'.
-unfold recursion.
-unfold NN.to_N.
-rewrite <- Exx'; clear x' Exx'.
-induction (Z.to_N [x]) using N.peano_ind.
-simpl; auto.
-rewrite 2 N.peano_rect_succ. now apply Eff'.
-Qed.
-
-Theorem recursion_0 :
- forall (A : Type) (a : A) (f : NN.t -> A -> A), recursion a f 0 = a.
-Proof.
-intros A a f; unfold recursion, NN.to_N; rewrite NN.spec_0; simpl; auto.
-Qed.
-
-Theorem recursion_succ :
- forall (A : Type) (Aeq : relation A) (a : A) (f : NN.t -> A -> A),
- Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
- forall n, Aeq (recursion a f (succ n)) (f n (recursion a f n)).
-Proof.
-unfold eq, recursion; intros A Aeq a f EAaa f_wd n.
-replace (to_N (succ n)) with (N.succ (to_N n)) by
- (zify; now rewrite <- Z2N.inj_succ by apply spec_pos).
-rewrite N.peano_rect_succ.
-apply f_wd; auto.
-zify. now rewrite Z2N.id by apply spec_pos.
-fold (recursion a f n). apply recursion_wd; auto. red; auto.
-Qed.
-
-End NTypeIsNAxioms.
-
-Module NType_NAxioms (NN : NType)
- <: NAxiomsSig <: OrderFunctions NN <: HasMinMax NN
- := NN <+ NTypeIsNAxioms.
diff --git a/theories/Numbers/Rational/BigQ/BigQ.v b/theories/Numbers/Rational/BigQ/BigQ.v
deleted file mode 100644
index 850afe534..000000000
--- a/theories/Numbers/Rational/BigQ/BigQ.v
+++ /dev/null
@@ -1,162 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** * BigQ: an efficient implementation of rational numbers *)
-
-(** Initial authors: Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-
-Require Export BigZ.
-Require Import Field Qfield QSig QMake Orders GenericMinMax.
-
-(** We choose for BigQ an implemention with
- multiple representation of 0: 0, 1/0, 2/0 etc.
- See [QMake.v] *)
-
-(** First, we provide translations functions between [BigN] and [BigZ] *)
-
-Module BigN_BigZ <: NType_ZType BigN.BigN BigZ.
- Definition Z_of_N := BigZ.Pos.
- Lemma spec_Z_of_N : forall n, BigZ.to_Z (Z_of_N n) = BigN.to_Z n.
- Proof.
- reflexivity.
- Qed.
- Definition Zabs_N := BigZ.to_N.
- Lemma spec_Zabs_N : forall z, BigN.to_Z (Zabs_N z) = Z.abs (BigZ.to_Z z).
- Proof.
- unfold Zabs_N; intros.
- rewrite BigZ.spec_to_Z, Z.mul_comm; apply Z.sgn_abs.
- Qed.
-End BigN_BigZ.
-
-(** This allows building [BigQ] out of [BigN] and [BigQ] via [QMake] *)
-
-Delimit Scope bigQ_scope with bigQ.
-
-Module BigQ <: QType <: OrderedTypeFull <: TotalOrder.
- Include QMake.Make BigN BigZ BigN_BigZ
- <+ !QProperties <+ HasEqBool2Dec
- <+ !MinMaxLogicalProperties <+ !MinMaxDecProperties.
- Ltac order := Private_Tac.order.
-End BigQ.
-
-(** Notations about [BigQ] *)
-
-Local Open Scope bigQ_scope.
-
-Notation bigQ := BigQ.t.
-Bind Scope bigQ_scope with bigQ BigQ.t BigQ.t_.
-(** As in QArith, we use [#] to denote fractions *)
-Notation "p # q" := (BigQ.Qq p q) (at level 55, no associativity) : bigQ_scope.
-Local Notation "0" := BigQ.zero : bigQ_scope.
-Local Notation "1" := BigQ.one : bigQ_scope.
-Infix "+" := BigQ.add : bigQ_scope.
-Infix "-" := BigQ.sub : bigQ_scope.
-Notation "- x" := (BigQ.opp x) : bigQ_scope.
-Infix "*" := BigQ.mul : bigQ_scope.
-Infix "/" := BigQ.div : bigQ_scope.
-Infix "^" := BigQ.power : bigQ_scope.
-Infix "?=" := BigQ.compare : bigQ_scope.
-Infix "==" := BigQ.eq : bigQ_scope.
-Notation "x != y" := (~x==y) (at level 70, no associativity) : bigQ_scope.
-Infix "<" := BigQ.lt : bigQ_scope.
-Infix "<=" := BigQ.le : bigQ_scope.
-Notation "x > y" := (BigQ.lt y x) (only parsing) : bigQ_scope.
-Notation "x >= y" := (BigQ.le y x) (only parsing) : bigQ_scope.
-Notation "x < y < z" := (x<y /\ y<z) : bigQ_scope.
-Notation "x < y <= z" := (x<y /\ y<=z) : bigQ_scope.
-Notation "x <= y < z" := (x<=y /\ y<z) : bigQ_scope.
-Notation "x <= y <= z" := (x<=y /\ y<=z) : bigQ_scope.
-Notation "[ q ]" := (BigQ.to_Q q) : bigQ_scope.
-
-(** [BigQ] is a field *)
-
-Lemma BigQfieldth :
- field_theory 0 1 BigQ.add BigQ.mul BigQ.sub BigQ.opp
- BigQ.div BigQ.inv BigQ.eq.
-Proof.
-constructor.
-constructor.
-exact BigQ.add_0_l. exact BigQ.add_comm. exact BigQ.add_assoc.
-exact BigQ.mul_1_l. exact BigQ.mul_comm. exact BigQ.mul_assoc.
-exact BigQ.mul_add_distr_r. exact BigQ.sub_add_opp.
-exact BigQ.add_opp_diag_r. exact BigQ.neq_1_0.
-exact BigQ.div_mul_inv. exact BigQ.mul_inv_diag_l.
-Qed.
-
-Declare Equivalent Keys pow_N pow_pos.
-
-Lemma BigQpowerth :
- power_theory 1 BigQ.mul BigQ.eq Z.of_N BigQ.power.
-Proof.
-constructor. intros. BigQ.qify.
-replace ([r] ^ Z.of_N n)%Q with (pow_N 1 Qmult [r] n)%Q by (now destruct n).
-destruct n. reflexivity.
-induction p; simpl; auto; rewrite ?BigQ.spec_mul, ?IHp; reflexivity.
-Qed.
-
-Ltac isBigQcst t :=
- match t with
- | BigQ.Qz ?t => isBigZcst t
- | BigQ.Qq ?n ?d => match isBigZcst n with
- | true => isBigNcst d
- | false => constr:(false)
- end
- | BigQ.zero => constr:(true)
- | BigQ.one => constr:(true)
- | BigQ.minus_one => constr:(true)
- | _ => constr:(false)
- end.
-
-Ltac BigQcst t :=
- match isBigQcst t with
- | true => constr:(t)
- | false => constr:(NotConstant)
- end.
-
-Add Field BigQfield : BigQfieldth
- (decidable BigQ.eqb_correct,
- completeness BigQ.eqb_complete,
- constants [BigQcst],
- power_tac BigQpowerth [Qpow_tac]).
-
-Section TestField.
-
-Let ex1 : forall x y z, (x+y)*z == (x*z)+(y*z).
- intros.
- ring.
-Qed.
-
-Let ex8 : forall x, x ^ 2 == x*x.
- intro.
- ring.
-Qed.
-
-Let ex10 : forall x y, y!=0 -> (x/y)*y == x.
-intros.
-field.
-auto.
-Qed.
-
-End TestField.
-
-(** [BigQ] can also benefit from an "order" tactic *)
-
-Ltac bigQ_order := BigQ.order.
-
-Section TestOrder.
-Let test : forall x y : bigQ, x<=y -> y<=x -> x==y.
-Proof. bigQ_order. Qed.
-End TestOrder.
-
-(** We can also reason by switching to QArith thanks to tactic
- BigQ.qify. *)
-
-Section TestQify.
-Let test : forall x : bigQ, 0+x == 1*x.
-Proof. intro x. BigQ.qify. ring. Qed.
-End TestQify.
diff --git a/theories/Numbers/Rational/BigQ/QMake.v b/theories/Numbers/Rational/BigQ/QMake.v
deleted file mode 100644
index b9fed9d56..000000000
--- a/theories/Numbers/Rational/BigQ/QMake.v
+++ /dev/null
@@ -1,1283 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** * QMake : a generic efficient implementation of rational numbers *)
-
-(** Initial authors : Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-
-Require Import BigNumPrelude ROmega.
-Require Import QArith Qcanon Qpower Qminmax.
-Require Import NSig ZSig QSig.
-
-(** We will build rationals out of an implementation of integers [ZType]
- for numerators and an implementation of natural numbers [NType] for
- denominators. But first we will need some glue between [NType] and
- [ZType]. *)
-
-Module Type NType_ZType (NN:NType)(ZZ:ZType).
- Parameter Z_of_N : NN.t -> ZZ.t.
- Parameter spec_Z_of_N : forall n, ZZ.to_Z (Z_of_N n) = NN.to_Z n.
- Parameter Zabs_N : ZZ.t -> NN.t.
- Parameter spec_Zabs_N : forall z, NN.to_Z (Zabs_N z) = Z.abs (ZZ.to_Z z).
-End NType_ZType.
-
-Module Make (NN:NType)(ZZ:ZType)(Import NZ:NType_ZType NN ZZ) <: QType.
-
- (** The notation of a rational number is either an integer x,
- interpreted as itself or a pair (x,y) of an integer x and a natural
- number y interpreted as x/y. The pairs (x,0) and (0,y) are all
- interpreted as 0. *)
-
- Inductive t_ :=
- | Qz : ZZ.t -> t_
- | Qq : ZZ.t -> NN.t -> t_.
-
- Definition t := t_.
-
- (** Specification with respect to [QArith] *)
-
- Local Open Scope Q_scope.
-
- Definition of_Z x: t := Qz (ZZ.of_Z x).
-
- Definition of_Q (q:Q) : t :=
- let (x,y) := q in
- match y with
- | 1%positive => Qz (ZZ.of_Z x)
- | _ => Qq (ZZ.of_Z x) (NN.of_N (Npos y))
- end.
-
- Definition to_Q (q: t) :=
- match q with
- | Qz x => ZZ.to_Z x # 1
- | Qq x y => if NN.eqb y NN.zero then 0
- else ZZ.to_Z x # Z.to_pos (NN.to_Z y)
- end.
-
- Notation "[ x ]" := (to_Q x).
-
- Lemma N_to_Z_pos :
- forall x, (NN.to_Z x <> NN.to_Z NN.zero)%Z -> (0 < NN.to_Z x)%Z.
- Proof.
- intros x; rewrite NN.spec_0; generalize (NN.spec_pos x). romega.
- Qed.
-
- Ltac destr_zcompare := case Z.compare_spec; intros ?H.
-
- Ltac destr_eqb :=
- match goal with
- | |- context [ZZ.eqb ?x ?y] =>
- rewrite (ZZ.spec_eqb x y);
- case (Z.eqb_spec (ZZ.to_Z x) (ZZ.to_Z y));
- destr_eqb
- | |- context [NN.eqb ?x ?y] =>
- rewrite (NN.spec_eqb x y);
- case (Z.eqb_spec (NN.to_Z x) (NN.to_Z y));
- [ | let H:=fresh "H" in
- try (intro H;generalize (N_to_Z_pos _ H); clear H)];
- destr_eqb
- | _ => idtac
- end.
-
- Hint Rewrite
- Z.add_0_r Z.add_0_l Z.mul_0_r Z.mul_0_l Z.mul_1_r Z.mul_1_l
- ZZ.spec_0 NN.spec_0 ZZ.spec_1 NN.spec_1 ZZ.spec_m1 ZZ.spec_opp
- ZZ.spec_compare NN.spec_compare
- ZZ.spec_add NN.spec_add ZZ.spec_mul NN.spec_mul ZZ.spec_div NN.spec_div
- ZZ.spec_gcd NN.spec_gcd Z.gcd_abs_l Z.gcd_1_r
- spec_Z_of_N spec_Zabs_N
- : nz.
-
- Ltac nzsimpl := autorewrite with nz in *.
-
- Ltac qsimpl := try red; unfold to_Q; simpl; intros;
- destr_eqb; simpl; nzsimpl; intros;
- rewrite ?Z2Pos.id by auto;
- auto.
-
- Theorem strong_spec_of_Q: forall q: Q, [of_Q q] = q.
- Proof.
- intros(x,y); destruct y; simpl; rewrite ?ZZ.spec_of_Z; auto;
- destr_eqb; now rewrite ?NN.spec_0, ?NN.spec_of_N.
- Qed.
-
- Theorem spec_of_Q: forall q: Q, [of_Q q] == q.
- Proof.
- intros; rewrite strong_spec_of_Q; red; auto.
- Qed.
-
- Definition eq x y := [x] == [y].
-
- Definition zero: t := Qz ZZ.zero.
- Definition one: t := Qz ZZ.one.
- Definition minus_one: t := Qz ZZ.minus_one.
-
- Lemma spec_0: [zero] == 0.
- Proof.
- simpl. nzsimpl. reflexivity.
- Qed.
-
- Lemma spec_1: [one] == 1.
- Proof.
- simpl. nzsimpl. reflexivity.
- Qed.
-
- Lemma spec_m1: [minus_one] == -(1).
- Proof.
- simpl. nzsimpl. reflexivity.
- Qed.
-
- Definition compare (x y: t) :=
- match x, y with
- | Qz zx, Qz zy => ZZ.compare zx zy
- | Qz zx, Qq ny dy =>
- if NN.eqb dy NN.zero then ZZ.compare zx ZZ.zero
- else ZZ.compare (ZZ.mul zx (Z_of_N dy)) ny
- | Qq nx dx, Qz zy =>
- if NN.eqb dx NN.zero then ZZ.compare ZZ.zero zy
- else ZZ.compare nx (ZZ.mul zy (Z_of_N dx))
- | Qq nx dx, Qq ny dy =>
- match NN.eqb dx NN.zero, NN.eqb dy NN.zero with
- | true, true => Eq
- | true, false => ZZ.compare ZZ.zero ny
- | false, true => ZZ.compare nx ZZ.zero
- | false, false => ZZ.compare (ZZ.mul nx (Z_of_N dy))
- (ZZ.mul ny (Z_of_N dx))
- end
- end.
-
- Theorem spec_compare: forall q1 q2, (compare q1 q2) = ([q1] ?= [q2]).
- Proof.
- intros [z1 | x1 y1] [z2 | x2 y2];
- unfold Qcompare, compare; qsimpl.
- Qed.
-
- Definition lt n m := [n] < [m].
- Definition le n m := [n] <= [m].
-
- Definition min n m := match compare n m with Gt => m | _ => n end.
- Definition max n m := match compare n m with Lt => m | _ => n end.
-
- Lemma spec_min : forall n m, [min n m] == Qmin [n] [m].
- Proof.
- unfold min, Qmin, GenericMinMax.gmin. intros.
- rewrite spec_compare; destruct Qcompare; auto with qarith.
- Qed.
-
- Lemma spec_max : forall n m, [max n m] == Qmax [n] [m].
- Proof.
- unfold max, Qmax, GenericMinMax.gmax. intros.
- rewrite spec_compare; destruct Qcompare; auto with qarith.
- Qed.
-
- Definition eq_bool n m :=
- match compare n m with Eq => true | _ => false end.
-
- Theorem spec_eq_bool: forall x y, eq_bool x y = Qeq_bool [x] [y].
- Proof.
- intros. unfold eq_bool. rewrite spec_compare. reflexivity.
- Qed.
-
- (** [check_int] : is a reduced fraction [n/d] in fact a integer ? *)
-
- Definition check_int n d :=
- match NN.compare NN.one d with
- | Lt => Qq n d
- | Eq => Qz n
- | Gt => zero (* n/0 encodes 0 *)
- end.
-
- Theorem strong_spec_check_int : forall n d, [check_int n d] = [Qq n d].
- Proof.
- intros; unfold check_int.
- nzsimpl.
- destr_zcompare.
- simpl. rewrite <- H; qsimpl. congruence.
- reflexivity.
- qsimpl. exfalso; romega.
- Qed.
-
- (** Normalisation function *)
-
- Definition norm n d : t :=
- let gcd := NN.gcd (Zabs_N n) d in
- match NN.compare NN.one gcd with
- | Lt => check_int (ZZ.div n (Z_of_N gcd)) (NN.div d gcd)
- | Eq => check_int n d
- | Gt => zero (* gcd = 0 => both numbers are 0 *)
- end.
-
- Theorem spec_norm: forall n q, [norm n q] == [Qq n q].
- Proof.
- intros p q; unfold norm.
- assert (Hp := NN.spec_pos (Zabs_N p)).
- assert (Hq := NN.spec_pos q).
- nzsimpl.
- destr_zcompare.
- (* Eq *)
- rewrite strong_spec_check_int; reflexivity.
- (* Lt *)
- rewrite strong_spec_check_int.
- qsimpl.
- generalize (Zgcd_div_pos (ZZ.to_Z p) (NN.to_Z q)). romega.
- replace (NN.to_Z q) with 0%Z in * by assumption.
- rewrite Zdiv_0_l in *; auto with zarith.
- apply Zgcd_div_swap0; romega.
- (* Gt *)
- qsimpl.
- assert (H' : Z.gcd (ZZ.to_Z p) (NN.to_Z q) = 0%Z).
- generalize (Z.gcd_nonneg (ZZ.to_Z p) (NN.to_Z q)); romega.
- symmetry; apply (Z.gcd_eq_0_l _ _ H'); auto.
- Qed.
-
- Theorem strong_spec_norm : forall p q, [norm p q] = Qred [Qq p q].
- Proof.
- intros.
- replace (Qred [Qq p q]) with (Qred [norm p q]) by
- (apply Qred_complete; apply spec_norm).
- symmetry; apply Qred_identity.
- unfold norm.
- assert (Hp := NN.spec_pos (Zabs_N p)).
- assert (Hq := NN.spec_pos q).
- nzsimpl.
- destr_zcompare; rewrite ?strong_spec_check_int.
- (* Eq *)
- qsimpl.
- (* Lt *)
- qsimpl.
- rewrite Zgcd_1_rel_prime.
- destruct (Z_lt_le_dec 0 (NN.to_Z q)).
- apply Zis_gcd_rel_prime; auto with zarith.
- apply Zgcd_is_gcd.
- replace (NN.to_Z q) with 0%Z in * by romega.
- rewrite Zdiv_0_l in *; romega.
- (* Gt *)
- simpl; auto with zarith.
- Qed.
-
- (** Reduction function : producing irreducible fractions *)
-
- Definition red (x : t) : t :=
- match x with
- | Qz z => x
- | Qq n d => norm n d
- end.
-
- Class Reduced x := is_reduced : [red x] = [x].
-
- Theorem spec_red : forall x, [red x] == [x].
- Proof.
- intros [ z | n d ].
- auto with qarith.
- unfold red.
- apply spec_norm.
- Qed.
-
- Theorem strong_spec_red : forall x, [red x] = Qred [x].
- Proof.
- intros [ z | n d ].
- unfold red.
- symmetry; apply Qred_identity; simpl; auto with zarith.
- unfold red; apply strong_spec_norm.
- Qed.
-
- Definition add (x y: t): t :=
- match x with
- | Qz zx =>
- match y with
- | Qz zy => Qz (ZZ.add zx zy)
- | Qq ny dy =>
- if NN.eqb dy NN.zero then x
- else Qq (ZZ.add (ZZ.mul zx (Z_of_N dy)) ny) dy
- end
- | Qq nx dx =>
- if NN.eqb dx NN.zero then y
- else match y with
- | Qz zy => Qq (ZZ.add nx (ZZ.mul zy (Z_of_N dx))) dx
- | Qq ny dy =>
- if NN.eqb dy NN.zero then x
- else
- let n := ZZ.add (ZZ.mul nx (Z_of_N dy)) (ZZ.mul ny (Z_of_N dx)) in
- let d := NN.mul dx dy in
- Qq n d
- end
- end.
-
- Theorem spec_add : forall x y, [add x y] == [x] + [y].
- Proof.
- intros [x | nx dx] [y | ny dy]; unfold Qplus; qsimpl;
- auto with zarith.
- rewrite Pos.mul_1_r, Z2Pos.id; auto.
- rewrite Pos.mul_1_r, Z2Pos.id; auto.
- rewrite Z.mul_eq_0 in *; intuition.
- rewrite Pos2Z.inj_mul, 2 Z2Pos.id; auto.
- Qed.
-
- Definition add_norm (x y: t): t :=
- match x with
- | Qz zx =>
- match y with
- | Qz zy => Qz (ZZ.add zx zy)
- | Qq ny dy =>
- if NN.eqb dy NN.zero then x
- else norm (ZZ.add (ZZ.mul zx (Z_of_N dy)) ny) dy
- end
- | Qq nx dx =>
- if NN.eqb dx NN.zero then y
- else match y with
- | Qz zy => norm (ZZ.add nx (ZZ.mul zy (Z_of_N dx))) dx
- | Qq ny dy =>
- if NN.eqb dy NN.zero then x
- else
- let n := ZZ.add (ZZ.mul nx (Z_of_N dy)) (ZZ.mul ny (Z_of_N dx)) in
- let d := NN.mul dx dy in
- norm n d
- end
- end.
-
- Theorem spec_add_norm : forall x y, [add_norm x y] == [x] + [y].
- Proof.
- intros x y; rewrite <- spec_add.
- destruct x; destruct y; unfold add_norm, add;
- destr_eqb; auto using Qeq_refl, spec_norm.
- Qed.
-
- Instance strong_spec_add_norm x y
- `(Reduced x, Reduced y) : Reduced (add_norm x y).
- Proof.
- unfold Reduced; intros.
- rewrite strong_spec_red.
- rewrite <- (Qred_complete [add x y]);
- [ | rewrite spec_add, spec_add_norm; apply Qeq_refl ].
- rewrite <- strong_spec_red.
- destruct x as [zx|nx dx]; destruct y as [zy|ny dy];
- simpl; destr_eqb; nzsimpl; simpl; auto.
- Qed.
-
- Definition opp (x: t): t :=
- match x with
- | Qz zx => Qz (ZZ.opp zx)
- | Qq nx dx => Qq (ZZ.opp nx) dx
- end.
-
- Theorem strong_spec_opp: forall q, [opp q] = -[q].
- Proof.
- intros [z | x y]; simpl.
- rewrite ZZ.spec_opp; auto.
- match goal with |- context[NN.eqb ?X ?Y] =>
- generalize (NN.spec_eqb X Y); case NN.eqb
- end; auto; rewrite NN.spec_0.
- rewrite ZZ.spec_opp; auto.
- Qed.
-
- Theorem spec_opp : forall q, [opp q] == -[q].
- Proof.
- intros; rewrite strong_spec_opp; red; auto.
- Qed.
-
- Instance strong_spec_opp_norm q `(Reduced q) : Reduced (opp q).
- Proof.
- unfold Reduced; intros.
- rewrite strong_spec_opp, <- H, !strong_spec_red, <- Qred_opp.
- apply Qred_complete; apply spec_opp.
- Qed.
-
- Definition sub x y := add x (opp y).
-
- Theorem spec_sub : forall x y, [sub x y] == [x] - [y].
- Proof.
- intros x y; unfold sub; rewrite spec_add; auto.
- rewrite spec_opp; ring.
- Qed.
-
- Definition sub_norm x y := add_norm x (opp y).
-
- Theorem spec_sub_norm : forall x y, [sub_norm x y] == [x] - [y].
- Proof.
- intros x y; unfold sub_norm; rewrite spec_add_norm; auto.
- rewrite spec_opp; ring.
- Qed.
-
- Instance strong_spec_sub_norm x y
- `(Reduced x, Reduced y) : Reduced (sub_norm x y).
- Proof.
- intros.
- unfold sub_norm.
- apply strong_spec_add_norm; auto.
- apply strong_spec_opp_norm; auto.
- Qed.
-
- Definition mul (x y: t): t :=
- match x, y with
- | Qz zx, Qz zy => Qz (ZZ.mul zx zy)
- | Qz zx, Qq ny dy => Qq (ZZ.mul zx ny) dy
- | Qq nx dx, Qz zy => Qq (ZZ.mul nx zy) dx
- | Qq nx dx, Qq ny dy => Qq (ZZ.mul nx ny) (NN.mul dx dy)
- end.
-
- Ltac nsubst :=
- match goal with E : NN.to_Z _ = _ |- _ => rewrite E in * end.
-
- Theorem spec_mul : forall x y, [mul x y] == [x] * [y].
- Proof.
- intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl; qsimpl.
- rewrite Pos.mul_1_r, Z2Pos.id; auto.
- rewrite Z.mul_eq_0 in *; intuition.
- nsubst; auto with zarith.
- nsubst; auto with zarith.
- nsubst; nzsimpl; auto with zarith.
- rewrite Pos2Z.inj_mul, 2 Z2Pos.id; auto.
- Qed.
-
- Definition norm_denum n d :=
- if NN.eqb d NN.one then Qz n else Qq n d.
-
- Lemma spec_norm_denum : forall n d,
- [norm_denum n d] == [Qq n d].
- Proof.
- unfold norm_denum; intros; simpl; qsimpl.
- congruence.
- nsubst; auto with zarith.
- Qed.
-
- Definition irred n d :=
- let gcd := NN.gcd (Zabs_N n) d in
- match NN.compare gcd NN.one with
- | Gt => (ZZ.div n (Z_of_N gcd), NN.div d gcd)
- | _ => (n, d)
- end.
-
- Lemma spec_irred : forall n d, exists g,
- let (n',d') := irred n d in
- (ZZ.to_Z n' * g = ZZ.to_Z n)%Z /\ (NN.to_Z d' * g = NN.to_Z d)%Z.
- Proof.
- intros.
- unfold irred; nzsimpl; simpl.
- destr_zcompare.
- exists 1%Z; nzsimpl; auto.
- exists 0%Z; nzsimpl.
- assert (Z.gcd (ZZ.to_Z n) (NN.to_Z d) = 0%Z).
- generalize (Z.gcd_nonneg (ZZ.to_Z n) (NN.to_Z d)); romega.
- clear H.
- split.
- symmetry; apply (Z.gcd_eq_0_l _ _ H0).
- symmetry; apply (Z.gcd_eq_0_r _ _ H0).
- exists (Z.gcd (ZZ.to_Z n) (NN.to_Z d)).
- simpl.
- split.
- nzsimpl.
- destruct (Zgcd_is_gcd (ZZ.to_Z n) (NN.to_Z d)).
- rewrite Z.mul_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
- nzsimpl.
- destruct (Zgcd_is_gcd (ZZ.to_Z n) (NN.to_Z d)).
- rewrite Z.mul_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
- Qed.
-
- Lemma spec_irred_zero : forall n d,
- (NN.to_Z d = 0)%Z <-> (NN.to_Z (snd (irred n d)) = 0)%Z.
- Proof.
- intros.
- unfold irred.
- split.
- nzsimpl; intros.
- destr_zcompare; auto.
- simpl.
- nzsimpl.
- rewrite H, Zdiv_0_l; auto.
- nzsimpl; destr_zcompare; simpl; auto.
- nzsimpl.
- intros.
- generalize (NN.spec_pos d); intros.
- destruct (NN.to_Z d); auto.
- assert (0 < 0)%Z.
- rewrite <- H0 at 2.
- apply Zgcd_div_pos; auto with zarith.
- compute; auto.
- discriminate.
- compute in H1; elim H1; auto.
- Qed.
-
- Lemma strong_spec_irred : forall n d,
- (NN.to_Z d <> 0%Z) ->
- let (n',d') := irred n d in Z.gcd (ZZ.to_Z n') (NN.to_Z d') = 1%Z.
- Proof.
- unfold irred; intros.
- nzsimpl.
- destr_zcompare; simpl; auto.
- elim H.
- apply (Z.gcd_eq_0_r (ZZ.to_Z n)).
- generalize (Z.gcd_nonneg (ZZ.to_Z n) (NN.to_Z d)); romega.
-
- nzsimpl.
- rewrite Zgcd_1_rel_prime.
- apply Zis_gcd_rel_prime.
- generalize (NN.spec_pos d); romega.
- generalize (Z.gcd_nonneg (ZZ.to_Z n) (NN.to_Z d)); romega.
- apply Zgcd_is_gcd; auto.
- Qed.
-
- Definition mul_norm_Qz_Qq z n d :=
- if ZZ.eqb z ZZ.zero then zero
- else
- let gcd := NN.gcd (Zabs_N z) d in
- match NN.compare gcd NN.one with
- | Gt =>
- let z := ZZ.div z (Z_of_N gcd) in
- let d := NN.div d gcd in
- norm_denum (ZZ.mul z n) d
- | _ => Qq (ZZ.mul z n) d
- end.
-
- Definition mul_norm (x y: t): t :=
- match x, y with
- | Qz zx, Qz zy => Qz (ZZ.mul zx zy)
- | Qz zx, Qq ny dy => mul_norm_Qz_Qq zx ny dy
- | Qq nx dx, Qz zy => mul_norm_Qz_Qq zy nx dx
- | Qq nx dx, Qq ny dy =>
- let (nx, dy) := irred nx dy in
- let (ny, dx) := irred ny dx in
- norm_denum (ZZ.mul ny nx) (NN.mul dx dy)
- end.
-
- Lemma spec_mul_norm_Qz_Qq : forall z n d,
- [mul_norm_Qz_Qq z n d] == [Qq (ZZ.mul z n) d].
- Proof.
- intros z n d; unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
- destr_eqb; nzsimpl; intros Hz.
- qsimpl; rewrite Hz; auto.
- destruct Z_le_gt_dec as [LE|GT].
- qsimpl.
- rewrite spec_norm_denum.
- qsimpl.
- rewrite Zdiv_gcd_zero in GT; auto with zarith.
- nsubst. rewrite Zdiv_0_l in *; discriminate.
- rewrite <- Z.mul_assoc, (Z.mul_comm (ZZ.to_Z n)), Z.mul_assoc.
- rewrite Zgcd_div_swap0; try romega.
- ring.
- Qed.
-
- Instance strong_spec_mul_norm_Qz_Qq z n d :
- forall `(Reduced (Qq n d)), Reduced (mul_norm_Qz_Qq z n d).
- Proof.
- unfold Reduced.
- rewrite 2 strong_spec_red, 2 Qred_iff.
- simpl; nzsimpl.
- destr_eqb; intros Hd H; simpl in *; nzsimpl.
-
- unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
- destr_eqb; intros Hz; simpl; nzsimpl; simpl; auto.
- destruct Z_le_gt_dec.
- simpl; nzsimpl.
- destr_eqb; simpl; nzsimpl; auto with zarith.
- unfold norm_denum. destr_eqb; simpl; nzsimpl.
- rewrite Hd, Zdiv_0_l; discriminate.
- intros _.
- destr_eqb; simpl; nzsimpl; auto.
- nzsimpl; rewrite Hd, Zdiv_0_l; auto with zarith.
-
- rewrite Z2Pos.id in H; auto.
- unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
- destr_eqb; intros Hz; simpl; nzsimpl; simpl; auto.
- destruct Z_le_gt_dec as [H'|H'].
- simpl; nzsimpl.
- destr_eqb; simpl; nzsimpl; auto.
- intros.
- rewrite Z2Pos.id; auto.
- apply Zgcd_mult_rel_prime; auto.
- generalize (Z.gcd_eq_0_l (ZZ.to_Z z) (NN.to_Z d))
- (Z.gcd_nonneg (ZZ.to_Z z) (NN.to_Z d)); romega.
- destr_eqb; simpl; nzsimpl; auto.
- unfold norm_denum.
- destr_eqb; nzsimpl; simpl; destr_eqb; simpl; auto.
- intros; nzsimpl.
- rewrite Z2Pos.id; auto.
- apply Zgcd_mult_rel_prime.
- rewrite Zgcd_1_rel_prime.
- apply Zis_gcd_rel_prime.
- generalize (NN.spec_pos d); romega.
- generalize (Z.gcd_nonneg (ZZ.to_Z z) (NN.to_Z d)); romega.
- apply Zgcd_is_gcd.
- destruct (Zgcd_is_gcd (ZZ.to_Z z) (NN.to_Z d)) as [ (z0,Hz0) (d0,Hd0) Hzd].
- replace (NN.to_Z d / Z.gcd (ZZ.to_Z z) (NN.to_Z d))%Z with d0.
- rewrite Zgcd_1_rel_prime in *.
- apply bezout_rel_prime.
- destruct (rel_prime_bezout _ _ H) as [u v Huv].
- apply Bezout_intro with u (v*(Z.gcd (ZZ.to_Z z) (NN.to_Z d)))%Z.
- rewrite <- Huv; rewrite Hd0 at 2; ring.
- rewrite Hd0 at 1.
- symmetry; apply Z_div_mult_full; auto with zarith.
- Qed.
-
- Theorem spec_mul_norm : forall x y, [mul_norm x y] == [x] * [y].
- Proof.
- intros x y; rewrite <- spec_mul; auto.
- unfold mul_norm, mul; destruct x; destruct y.
- apply Qeq_refl.
- apply spec_mul_norm_Qz_Qq.
- rewrite spec_mul_norm_Qz_Qq; qsimpl; ring.
-
- rename t0 into nx, t3 into dy, t2 into ny, t1 into dx.
- destruct (spec_irred nx dy) as (g & Hg).
- destruct (spec_irred ny dx) as (g' & Hg').
- assert (Hz := spec_irred_zero nx dy).
- assert (Hz':= spec_irred_zero ny dx).
- destruct irred as (n1,d1); destruct irred as (n2,d2).
- simpl @snd in *; destruct Hg as [Hg1 Hg2]; destruct Hg' as [Hg1' Hg2'].
- rewrite spec_norm_denum.
- qsimpl.
-
- match goal with E : (_ * _ = 0)%Z |- _ =>
- rewrite Z.mul_eq_0 in E; destruct E as [Eq|Eq] end.
- rewrite Eq in *; simpl in *.
- rewrite <- Hg2' in *; auto with zarith.
- rewrite Eq in *; simpl in *.
- rewrite <- Hg2 in *; auto with zarith.
-
- match goal with E : (_ * _ = 0)%Z |- _ =>
- rewrite Z.mul_eq_0 in E; destruct E as [Eq|Eq] end.
- rewrite Hz' in Eq; rewrite Eq in *; auto with zarith.
- rewrite Hz in Eq; rewrite Eq in *; auto with zarith.
-
- rewrite <- Hg1, <- Hg2, <- Hg1', <- Hg2'; ring.
- Qed.
-
- Instance strong_spec_mul_norm x y :
- forall `(Reduced x, Reduced y), Reduced (mul_norm x y).
- Proof.
- unfold Reduced; intros.
- rewrite strong_spec_red, Qred_iff.
- destruct x as [zx|nx dx]; destruct y as [zy|ny dy].
- simpl in *; auto with zarith.
- simpl.
- rewrite <- Qred_iff, <- strong_spec_red, strong_spec_mul_norm_Qz_Qq; auto.
- simpl.
- rewrite <- Qred_iff, <- strong_spec_red, strong_spec_mul_norm_Qz_Qq; auto.
- simpl.
- destruct (spec_irred nx dy) as [g Hg].
- destruct (spec_irred ny dx) as [g' Hg'].
- assert (Hz := spec_irred_zero nx dy).
- assert (Hz':= spec_irred_zero ny dx).
- assert (Hgc := strong_spec_irred nx dy).
- assert (Hgc' := strong_spec_irred ny dx).
- destruct irred as (n1,d1); destruct irred as (n2,d2).
- simpl @snd in *; destruct Hg as [Hg1 Hg2]; destruct Hg' as [Hg1' Hg2'].
-
- unfold norm_denum; qsimpl.
-
- assert (NEQ : NN.to_Z dy <> 0%Z) by
- (rewrite Hz; intros EQ; rewrite EQ in *; romega).
- specialize (Hgc NEQ).
-
- assert (NEQ' : NN.to_Z dx <> 0%Z) by
- (rewrite Hz'; intro EQ; rewrite EQ in *; romega).
- specialize (Hgc' NEQ').
-
- revert H H0.
- rewrite 2 strong_spec_red, 2 Qred_iff; simpl.
- destr_eqb; simpl; nzsimpl; try romega; intros.
- rewrite Z2Pos.id in *; auto.
-
- apply Zgcd_mult_rel_prime; rewrite Z.gcd_comm;
- apply Zgcd_mult_rel_prime; rewrite Z.gcd_comm; auto.
-
- rewrite Zgcd_1_rel_prime in *.
- apply bezout_rel_prime.
- destruct (rel_prime_bezout (ZZ.to_Z ny) (NN.to_Z dy)) as [u v Huv]; trivial.
- apply Bezout_intro with (u*g')%Z (v*g)%Z.
- rewrite <- Huv, <- Hg1', <- Hg2. ring.
-
- rewrite Zgcd_1_rel_prime in *.
- apply bezout_rel_prime.
- destruct (rel_prime_bezout (ZZ.to_Z nx) (NN.to_Z dx)) as [u v Huv]; trivial.
- apply Bezout_intro with (u*g)%Z (v*g')%Z.
- rewrite <- Huv, <- Hg2', <- Hg1. ring.
- Qed.
-
- Definition inv (x: t): t :=
- match x with
- | Qz z =>
- match ZZ.compare ZZ.zero z with
- | Eq => zero
- | Lt => Qq ZZ.one (Zabs_N z)
- | Gt => Qq ZZ.minus_one (Zabs_N z)
- end
- | Qq n d =>
- match ZZ.compare ZZ.zero n with
- | Eq => zero
- | Lt => Qq (Z_of_N d) (Zabs_N n)
- | Gt => Qq (ZZ.opp (Z_of_N d)) (Zabs_N n)
- end
- end.
-
- Theorem spec_inv : forall x, [inv x] == /[x].
- Proof.
- destruct x as [ z | n d ].
- (* Qz z *)
- simpl.
- rewrite ZZ.spec_compare; destr_zcompare.
- (* 0 = z *)
- rewrite <- H.
- simpl; nzsimpl; compute; auto.
- (* 0 < z *)
- simpl.
- destr_eqb; nzsimpl; [ intros; rewrite Z.abs_eq in *; romega | intros _ ].
- set (z':=ZZ.to_Z z) in *; clearbody z'.
- red; simpl.
- rewrite Z.abs_eq by romega.
- rewrite Z2Pos.id by auto.
- unfold Qinv; simpl; destruct z'; simpl; auto; discriminate.
- (* 0 > z *)
- simpl.
- destr_eqb; nzsimpl; [ intros; rewrite Z.abs_neq in *; romega | intros _ ].
- set (z':=ZZ.to_Z z) in *; clearbody z'.
- red; simpl.
- rewrite Z.abs_neq by romega.
- rewrite Z2Pos.id by romega.
- unfold Qinv; simpl; destruct z'; simpl; auto; discriminate.
- (* Qq n d *)
- simpl.
- rewrite ZZ.spec_compare; destr_zcompare.
- (* 0 = n *)
- rewrite <- H.
- simpl; nzsimpl.
- destr_eqb; intros; compute; auto.
- (* 0 < n *)
- simpl.
- destr_eqb; nzsimpl; intros.
- intros; rewrite Z.abs_eq in *; romega.
- intros; rewrite Z.abs_eq in *; romega.
- nsubst; compute; auto.
- set (n':=ZZ.to_Z n) in *; clearbody n'.
- rewrite Z.abs_eq by romega.
- red; simpl.
- rewrite Z2Pos.id by auto.
- unfold Qinv; simpl; destruct n'; simpl; auto; try discriminate.
- rewrite Pos2Z.inj_mul, Z2Pos.id; auto.
- (* 0 > n *)
- simpl.
- destr_eqb; nzsimpl; intros.
- intros; rewrite Z.abs_neq in *; romega.
- intros; rewrite Z.abs_neq in *; romega.
- nsubst; compute; auto.
- set (n':=ZZ.to_Z n) in *; clearbody n'.
- red; simpl; nzsimpl.
- rewrite Z.abs_neq by romega.
- rewrite Z2Pos.id by romega.
- unfold Qinv; simpl; destruct n'; simpl; auto; try discriminate.
- assert (T : forall x, Zneg x = Z.opp (Zpos x)) by auto.
- rewrite T, Pos2Z.inj_mul, Z2Pos.id; auto; ring.
- Qed.
-
- Definition inv_norm (x: t): t :=
- match x with
- | Qz z =>
- match ZZ.compare ZZ.zero z with
- | Eq => zero
- | Lt => Qq ZZ.one (Zabs_N z)
- | Gt => Qq ZZ.minus_one (Zabs_N z)
- end
- | Qq n d =>
- if NN.eqb d NN.zero then zero else
- match ZZ.compare ZZ.zero n with
- | Eq => zero
- | Lt =>
- match ZZ.compare n ZZ.one with
- | Gt => Qq (Z_of_N d) (Zabs_N n)
- | _ => Qz (Z_of_N d)
- end
- | Gt =>
- match ZZ.compare n ZZ.minus_one with
- | Lt => Qq (ZZ.opp (Z_of_N d)) (Zabs_N n)
- | _ => Qz (ZZ.opp (Z_of_N d))
- end
- end
- end.
-
- Theorem spec_inv_norm : forall x, [inv_norm x] == /[x].
- Proof.
- intros.
- rewrite <- spec_inv.
- destruct x as [ z | n d ].
- (* Qz z *)
- simpl.
- rewrite ZZ.spec_compare; destr_zcompare; auto with qarith.
- (* Qq n d *)
- simpl; nzsimpl; destr_eqb.
- destr_zcompare; simpl; auto with qarith.
- destr_eqb; nzsimpl; auto with qarith.
- intros _ Hd; rewrite Hd; auto with qarith.
- destr_eqb; nzsimpl; auto with qarith.
- intros _ Hd; rewrite Hd; auto with qarith.
- (* 0 < n *)
- destr_zcompare; auto with qarith.
- destr_zcompare; nzsimpl; simpl; auto with qarith; intros.
- destr_eqb; nzsimpl; [ intros; rewrite Z.abs_eq in *; romega | intros _ ].
- rewrite H0; auto with qarith.
- romega.
- (* 0 > n *)
- destr_zcompare; nzsimpl; simpl; auto with qarith.
- destr_eqb; nzsimpl; [ intros; rewrite Z.abs_neq in *; romega | intros _ ].
- rewrite H0; auto with qarith.
- romega.
- Qed.
-
- Instance strong_spec_inv_norm x : Reduced x -> Reduced (inv_norm x).
- Proof.
- unfold Reduced.
- intros.
- destruct x as [ z | n d ].
- (* Qz *)
- simpl; nzsimpl.
- rewrite strong_spec_red, Qred_iff.
- destr_zcompare; simpl; nzsimpl; auto.
- destr_eqb; nzsimpl; simpl; auto.
- destr_eqb; nzsimpl; simpl; auto.
- (* Qq n d *)
- rewrite strong_spec_red, Qred_iff in H; revert H.
- simpl; nzsimpl.
- destr_eqb; nzsimpl; auto with qarith.
- destr_zcompare; simpl; nzsimpl; auto; intros.
- (* 0 < n *)
- destr_zcompare; simpl; nzsimpl; auto.
- destr_eqb; nzsimpl; simpl; auto.
- rewrite Z.abs_eq; romega.
- intros _.
- rewrite strong_spec_norm; simpl; nzsimpl.
- destr_eqb; nzsimpl.
- rewrite Z.abs_eq; romega.
- intros _.
- rewrite Qred_iff.
- simpl.
- rewrite Z.abs_eq; auto with zarith.
- rewrite Z2Pos.id in *; auto.
- rewrite Z.gcd_comm; auto.
- (* 0 > n *)
- destr_eqb; nzsimpl; simpl; auto; intros.
- destr_zcompare; simpl; nzsimpl; auto.
- destr_eqb; nzsimpl.
- rewrite Z.abs_neq; romega.
- intros _.
- rewrite strong_spec_norm; simpl; nzsimpl.
- destr_eqb; nzsimpl.
- rewrite Z.abs_neq; romega.
- intros _.
- rewrite Qred_iff.
- simpl.
- rewrite Z2Pos.id in *; auto.
- intros.
- rewrite Z.gcd_comm, Z.gcd_abs_l, Z.gcd_comm.
- apply Zis_gcd_gcd; auto with zarith.
- apply Zis_gcd_minus.
- rewrite Z.opp_involutive, <- H1; apply Zgcd_is_gcd.
- rewrite Z.abs_neq; romega.
- Qed.
-
- Definition div x y := mul x (inv y).
-
- Theorem spec_div x y: [div x y] == [x] / [y].
- Proof.
- unfold div; rewrite spec_mul; auto.
- unfold Qdiv; apply Qmult_comp.
- apply Qeq_refl.
- apply spec_inv; auto.
- Qed.
-
- Definition div_norm x y := mul_norm x (inv_norm y).
-
- Theorem spec_div_norm x y: [div_norm x y] == [x] / [y].
- Proof.
- unfold div_norm; rewrite spec_mul_norm; auto.
- unfold Qdiv; apply Qmult_comp.
- apply Qeq_refl.
- apply spec_inv_norm; auto.
- Qed.
-
- Instance strong_spec_div_norm x y
- `(Reduced x, Reduced y) : Reduced (div_norm x y).
- Proof.
- intros; unfold div_norm.
- apply strong_spec_mul_norm; auto.
- apply strong_spec_inv_norm; auto.
- Qed.
-
- Definition square (x: t): t :=
- match x with
- | Qz zx => Qz (ZZ.square zx)
- | Qq nx dx => Qq (ZZ.square nx) (NN.square dx)
- end.
-
- Theorem spec_square : forall x, [square x] == [x] ^ 2.
- Proof.
- destruct x as [ z | n d ].
- simpl; rewrite ZZ.spec_square; red; auto.
- simpl.
- destr_eqb; nzsimpl; intros.
- apply Qeq_refl.
- rewrite NN.spec_square in *; nzsimpl.
- rewrite Z.mul_eq_0 in *; romega.
- rewrite NN.spec_square in *; nzsimpl; nsubst; romega.
- rewrite ZZ.spec_square, NN.spec_square.
- red; simpl.
- rewrite Pos2Z.inj_mul; rewrite !Z2Pos.id; auto.
- apply Z.mul_pos_pos; auto.
- Qed.
-
- Definition power_pos (x : t) p : t :=
- match x with
- | Qz zx => Qz (ZZ.pow_pos zx p)
- | Qq nx dx => Qq (ZZ.pow_pos nx p) (NN.pow_pos dx p)
- end.
-
- Theorem spec_power_pos : forall x p, [power_pos x p] == [x] ^ Zpos p.
- Proof.
- intros [ z | n d ] p; unfold power_pos.
- (* Qz *)
- simpl.
- rewrite ZZ.spec_pow_pos, Qpower_decomp.
- red; simpl; f_equal.
- now rewrite Pos2Z.inj_pow, Z.pow_1_l.
- (* Qq *)
- simpl.
- rewrite ZZ.spec_pow_pos.
- destr_eqb; nzsimpl; intros.
- - apply Qeq_sym; apply Qpower_positive_0.
- - rewrite NN.spec_pow_pos in *.
- assert (0 < NN.to_Z d ^ ' p)%Z by
- (apply Z.pow_pos_nonneg; auto with zarith).
- romega.
- - exfalso.
- rewrite NN.spec_pow_pos in *. nsubst.
- rewrite Z.pow_0_l' in *; [romega|discriminate].
- - rewrite Qpower_decomp.
- red; simpl; do 3 f_equal.
- apply Pos2Z.inj. rewrite Pos2Z.inj_pow.
- rewrite 2 Z2Pos.id by (generalize (NN.spec_pos d); romega).
- now rewrite NN.spec_pow_pos.
- Qed.
-
- Instance strong_spec_power_pos x p `(Reduced x) : Reduced (power_pos x p).
- Proof.
- destruct x as [z | n d]; simpl; intros.
- red; simpl; auto.
- red; simpl; intros.
- rewrite strong_spec_norm; simpl.
- destr_eqb; nzsimpl; intros.
- simpl; auto.
- rewrite Qred_iff.
- revert H.
- unfold Reduced; rewrite strong_spec_red, Qred_iff; simpl.
- destr_eqb; nzsimpl; simpl; intros.
- exfalso.
- rewrite NN.spec_pow_pos in *. nsubst.
- rewrite Z.pow_0_l' in *; [romega|discriminate].
- rewrite Z2Pos.id in *; auto.
- rewrite NN.spec_pow_pos, ZZ.spec_pow_pos; auto.
- rewrite Zgcd_1_rel_prime in *.
- apply rel_prime_Zpower; auto with zarith.
- Qed.
-
- Definition power (x : t) (z : Z) : t :=
- match z with
- | Z0 => one
- | Zpos p => power_pos x p
- | Zneg p => inv (power_pos x p)
- end.
-
- Theorem spec_power : forall x z, [power x z] == [x]^z.
- Proof.
- destruct z.
- simpl; nzsimpl; red; auto.
- apply spec_power_pos.
- simpl.
- rewrite spec_inv, spec_power_pos; apply Qeq_refl.
- Qed.
-
- Definition power_norm (x : t) (z : Z) : t :=
- match z with
- | Z0 => one
- | Zpos p => power_pos x p
- | Zneg p => inv_norm (power_pos x p)
- end.
-
- Theorem spec_power_norm : forall x z, [power_norm x z] == [x]^z.
- Proof.
- destruct z.
- simpl; nzsimpl; red; auto.
- apply spec_power_pos.
- simpl.
- rewrite spec_inv_norm, spec_power_pos; apply Qeq_refl.
- Qed.
-
- Instance strong_spec_power_norm x z :
- Reduced x -> Reduced (power_norm x z).
- Proof.
- destruct z; simpl.
- intros _; unfold Reduced; rewrite strong_spec_red.
- unfold one.
- simpl to_Q; nzsimpl; auto.
- intros; apply strong_spec_power_pos; auto.
- intros; apply strong_spec_inv_norm; apply strong_spec_power_pos; auto.
- Qed.
-
-
- (** Interaction with [Qcanon.Qc] *)
-
- Open Scope Qc_scope.
-
- Definition of_Qc q := of_Q (this q).
-
- Definition to_Qc q := Q2Qc [q].
-
- Notation "[[ x ]]" := (to_Qc x).
-
- Theorem strong_spec_of_Qc : forall q, [of_Qc q] = q.
- Proof.
- intros (q,Hq); intros.
- unfold of_Qc; rewrite strong_spec_of_Q; auto.
- Qed.
-
- Instance strong_spec_of_Qc_bis q : Reduced (of_Qc q).
- Proof.
- intros; red; rewrite strong_spec_red, strong_spec_of_Qc.
- destruct q; simpl; auto.
- Qed.
-
- Theorem spec_of_Qc: forall q, [[of_Qc q]] = q.
- Proof.
- intros; apply Qc_decomp; simpl; intros.
- rewrite strong_spec_of_Qc. apply canon.
- Qed.
-
- Theorem spec_oppc: forall q, [[opp q]] = -[[q]].
- Proof.
- intros q; unfold Qcopp, to_Qc, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- rewrite spec_opp, <- Qred_opp, Qred_correct.
- apply Qeq_refl.
- Qed.
-
- Theorem spec_oppc_bis : forall q : Qc, [opp (of_Qc q)] = - q.
- Proof.
- intros.
- rewrite <- strong_spec_opp_norm by apply strong_spec_of_Qc_bis.
- rewrite strong_spec_red.
- symmetry; apply (Qred_complete (-q)%Q).
- rewrite spec_opp, strong_spec_of_Qc; auto with qarith.
- Qed.
-
- Theorem spec_comparec: forall q1 q2,
- compare q1 q2 = ([[q1]] ?= [[q2]]).
- Proof.
- unfold Qccompare, to_Qc.
- intros q1 q2; rewrite spec_compare; simpl; auto.
- apply Qcompare_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_addc x y:
- [[add x y]] = [[x]] + [[y]].
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc ([x] + [y])).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_add; auto.
- unfold Qcplus, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qplus_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_add_normc x y:
- [[add_norm x y]] = [[x]] + [[y]].
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc ([x] + [y])).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_add_norm; auto.
- unfold Qcplus, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qplus_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_add_normc_bis : forall x y : Qc,
- [add_norm (of_Qc x) (of_Qc y)] = x+y.
- Proof.
- intros.
- rewrite <- strong_spec_add_norm by apply strong_spec_of_Qc_bis.
- rewrite strong_spec_red.
- symmetry; apply (Qred_complete (x+y)%Q).
- rewrite spec_add_norm, ! strong_spec_of_Qc; auto with qarith.
- Qed.
-
- Theorem spec_subc x y: [[sub x y]] = [[x]] - [[y]].
- Proof.
- unfold sub; rewrite spec_addc; auto.
- rewrite spec_oppc; ring.
- Qed.
-
- Theorem spec_sub_normc x y:
- [[sub_norm x y]] = [[x]] - [[y]].
- Proof.
- unfold sub_norm; rewrite spec_add_normc; auto.
- rewrite spec_oppc; ring.
- Qed.
-
- Theorem spec_sub_normc_bis : forall x y : Qc,
- [sub_norm (of_Qc x) (of_Qc y)] = x-y.
- Proof.
- intros.
- rewrite <- strong_spec_sub_norm by apply strong_spec_of_Qc_bis.
- rewrite strong_spec_red.
- symmetry; apply (Qred_complete (x+(-y)%Qc)%Q).
- rewrite spec_sub_norm, ! strong_spec_of_Qc.
- unfold Qcopp, Q2Qc, this. rewrite Qred_correct ; auto with qarith.
- Qed.
-
- Theorem spec_mulc x y:
- [[mul x y]] = [[x]] * [[y]].
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc ([x] * [y])).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_mul; auto.
- unfold Qcmult, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qmult_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_mul_normc x y:
- [[mul_norm x y]] = [[x]] * [[y]].
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc ([x] * [y])).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_mul_norm; auto.
- unfold Qcmult, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qmult_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_mul_normc_bis : forall x y : Qc,
- [mul_norm (of_Qc x) (of_Qc y)] = x*y.
- Proof.
- intros.
- rewrite <- strong_spec_mul_norm by apply strong_spec_of_Qc_bis.
- rewrite strong_spec_red.
- symmetry; apply (Qred_complete (x*y)%Q).
- rewrite spec_mul_norm, ! strong_spec_of_Qc; auto with qarith.
- Qed.
-
- Theorem spec_invc x:
- [[inv x]] = /[[x]].
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc (/[x])).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_inv; auto.
- unfold Qcinv, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qinv_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_inv_normc x:
- [[inv_norm x]] = /[[x]].
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc (/[x])).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_inv_norm; auto.
- unfold Qcinv, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qinv_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_inv_normc_bis : forall x : Qc,
- [inv_norm (of_Qc x)] = /x.
- Proof.
- intros.
- rewrite <- strong_spec_inv_norm by apply strong_spec_of_Qc_bis.
- rewrite strong_spec_red.
- symmetry; apply (Qred_complete (/x)%Q).
- rewrite spec_inv_norm, ! strong_spec_of_Qc; auto with qarith.
- Qed.
-
- Theorem spec_divc x y: [[div x y]] = [[x]] / [[y]].
- Proof.
- unfold div; rewrite spec_mulc; auto.
- unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto.
- apply spec_invc; auto.
- Qed.
-
- Theorem spec_div_normc x y: [[div_norm x y]] = [[x]] / [[y]].
- Proof.
- unfold div_norm; rewrite spec_mul_normc; auto.
- unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto.
- apply spec_inv_normc; auto.
- Qed.
-
- Theorem spec_div_normc_bis : forall x y : Qc,
- [div_norm (of_Qc x) (of_Qc y)] = x/y.
- Proof.
- intros.
- rewrite <- strong_spec_div_norm by apply strong_spec_of_Qc_bis.
- rewrite strong_spec_red.
- symmetry; apply (Qred_complete (x*(/y)%Qc)%Q).
- rewrite spec_div_norm, ! strong_spec_of_Qc.
- unfold Qcinv, Q2Qc, this; rewrite Qred_correct; auto with qarith.
- Qed.
-
- Theorem spec_squarec x: [[square x]] = [[x]]^2.
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc ([x]^2)).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_square; auto.
- simpl Qcpower.
- replace (Q2Qc [x] * 1) with (Q2Qc [x]); try ring.
- simpl.
- unfold Qcmult, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- apply Qmult_comp; apply Qeq_sym; apply Qred_correct.
- Qed.
-
- Theorem spec_power_posc x p:
- [[power_pos x p]] = [[x]] ^ Pos.to_nat p.
- Proof.
- unfold to_Qc.
- transitivity (Q2Qc ([x]^Zpos p)).
- unfold Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete; apply spec_power_pos; auto.
- induction p using Pos.peano_ind.
- simpl; ring.
- rewrite Pos2Nat.inj_succ; simpl Qcpower.
- rewrite <- IHp; clear IHp.
- unfold Qcmult, Q2Qc.
- apply Qc_decomp; unfold this.
- apply Qred_complete.
- setoid_replace ([x] ^ ' Pos.succ p)%Q with ([x] * [x] ^ ' p)%Q.
- apply Qmult_comp; apply Qeq_sym; apply Qred_correct.
- simpl.
- rewrite <- Pos.add_1_l.
- rewrite Qpower_plus_positive; simpl; apply Qeq_refl.
- Qed.
-
-End Make.
diff --git a/theories/Numbers/Rational/SpecViaQ/QSig.v b/theories/Numbers/Rational/SpecViaQ/QSig.v
deleted file mode 100644
index 8e20fd060..000000000
--- a/theories/Numbers/Rational/SpecViaQ/QSig.v
+++ /dev/null
@@ -1,229 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-Require Import QArith Qpower Qminmax Orders RelationPairs GenericMinMax.
-
-Open Scope Q_scope.
-
-(** * QSig *)
-
-(** Interface of a rich structure about rational numbers.
- Specifications are written via translation to Q.
-*)
-
-Module Type QType.
-
- Parameter t : Type.
-
- Parameter to_Q : t -> Q.
- Local Notation "[ x ]" := (to_Q x).
-
- Definition eq x y := [x] == [y].
- Definition lt x y := [x] < [y].
- Definition le x y := [x] <= [y].
-
- Parameter of_Q : Q -> t.
- Parameter spec_of_Q: forall x, to_Q (of_Q x) == x.
-
- Parameter red : t -> t.
- Parameter compare : t -> t -> comparison.
- Parameter eq_bool : t -> t -> bool.
- Parameter max : t -> t -> t.
- Parameter min : t -> t -> t.
- Parameter zero : t.
- Parameter one : t.
- Parameter minus_one : t.
- Parameter add : t -> t -> t.
- Parameter sub : t -> t -> t.
- Parameter opp : t -> t.
- Parameter mul : t -> t -> t.
- Parameter square : t -> t.
- Parameter inv : t -> t.
- Parameter div : t -> t -> t.
- Parameter power : t -> Z -> t.
-
- Parameter spec_red : forall x, [red x] == [x].
- Parameter strong_spec_red : forall x, [red x] = Qred [x].
- Parameter spec_compare : forall x y, compare x y = ([x] ?= [y]).
- Parameter spec_eq_bool : forall x y, eq_bool x y = Qeq_bool [x] [y].
- Parameter spec_max : forall x y, [max x y] == Qmax [x] [y].
- Parameter spec_min : forall x y, [min x y] == Qmin [x] [y].
- Parameter spec_0: [zero] == 0.
- Parameter spec_1: [one] == 1.
- Parameter spec_m1: [minus_one] == -(1).
- Parameter spec_add: forall x y, [add x y] == [x] + [y].
- Parameter spec_sub: forall x y, [sub x y] == [x] - [y].
- Parameter spec_opp: forall x, [opp x] == - [x].
- Parameter spec_mul: forall x y, [mul x y] == [x] * [y].
- Parameter spec_square: forall x, [square x] == [x] ^ 2.
- Parameter spec_inv : forall x, [inv x] == / [x].
- Parameter spec_div: forall x y, [div x y] == [x] / [y].
- Parameter spec_power: forall x z, [power x z] == [x] ^ z.
-
-End QType.
-
-(** NB: several of the above functions come with [..._norm] variants
- that expect reduced arguments and return reduced results. *)
-
-(** TODO : also speak of specifications via Qcanon ... *)
-
-Module Type QType_Notation (Import Q : QType).
- Notation "[ x ]" := (to_Q x).
- Infix "==" := eq (at level 70).
- Notation "x != y" := (~x==y) (at level 70).
- Infix "<=" := le.
- Infix "<" := lt.
- Notation "0" := zero.
- Notation "1" := one.
- Infix "+" := add.
- Infix "-" := sub.
- Infix "*" := mul.
- Notation "- x" := (opp x).
- Infix "/" := div.
- Notation "/ x" := (inv x).
- Infix "^" := power.
-End QType_Notation.
-
-Module Type QType' := QType <+ QType_Notation.
-
-
-Module QProperties (Import Q : QType').
-
-(** Conversion to Q *)
-
-Hint Rewrite
- spec_red spec_compare spec_eq_bool spec_min spec_max
- spec_add spec_sub spec_opp spec_mul spec_square spec_inv spec_div
- spec_power : qsimpl.
-Ltac qify := unfold eq, lt, le in *; autorewrite with qsimpl;
- try rewrite spec_0 in *; try rewrite spec_1 in *; try rewrite spec_m1 in *.
-
-(** NB: do not add [spec_0] in the autorewrite database. Otherwise,
- after instantiation in BigQ, this lemma become convertible to 0=0,
- and autorewrite loops. Idem for [spec_1] and [spec_m1] *)
-
-(** Morphisms *)
-
-Ltac solve_wd1 := intros x x' Hx; qify; now rewrite Hx.
-Ltac solve_wd2 := intros x x' Hx y y' Hy; qify; now rewrite Hx, Hy.
-
-Local Obligation Tactic := solve_wd2 || solve_wd1.
-
-Instance : Measure to_Q.
-Instance eq_equiv : Equivalence eq.
-Proof.
- change eq with (RelCompFun Qeq to_Q); apply _.
-Defined.
-
-Program Instance lt_wd : Proper (eq==>eq==>iff) lt.
-Program Instance le_wd : Proper (eq==>eq==>iff) le.
-Program Instance red_wd : Proper (eq==>eq) red.
-Program Instance compare_wd : Proper (eq==>eq==>Logic.eq) compare.
-Program Instance eq_bool_wd : Proper (eq==>eq==>Logic.eq) eq_bool.
-Program Instance min_wd : Proper (eq==>eq==>eq) min.
-Program Instance max_wd : Proper (eq==>eq==>eq) max.
-Program Instance add_wd : Proper (eq==>eq==>eq) add.
-Program Instance sub_wd : Proper (eq==>eq==>eq) sub.
-Program Instance opp_wd : Proper (eq==>eq) opp.
-Program Instance mul_wd : Proper (eq==>eq==>eq) mul.
-Program Instance square_wd : Proper (eq==>eq) square.
-Program Instance inv_wd : Proper (eq==>eq) inv.
-Program Instance div_wd : Proper (eq==>eq==>eq) div.
-Program Instance power_wd : Proper (eq==>Logic.eq==>eq) power.
-
-(** Let's implement [HasCompare] *)
-
-Lemma compare_spec : forall x y, CompareSpec (x==y) (x<y) (y<x) (compare x y).
-Proof. intros. qify. destruct (Qcompare_spec [x] [y]); auto. Qed.
-
-(** Let's implement [TotalOrder] *)
-
-Definition lt_compat := lt_wd.
-Instance lt_strorder : StrictOrder lt.
-Proof.
- change lt with (RelCompFun Qlt to_Q); apply _.
-Qed.
-
-Lemma le_lteq : forall x y, x<=y <-> x<y \/ x==y.
-Proof. intros. qify. apply Qle_lteq. Qed.
-
-Lemma lt_total : forall x y, x<y \/ x==y \/ y<x.
-Proof. intros. destruct (compare_spec x y); auto. Qed.
-
-(** Let's implement [HasEqBool] *)
-
-Definition eqb := eq_bool.
-
-Lemma eqb_eq : forall x y, eq_bool x y = true <-> x == y.
-Proof. intros. qify. apply Qeq_bool_iff. Qed.
-
-Lemma eqb_correct : forall x y, eq_bool x y = true -> x == y.
-Proof. now apply eqb_eq. Qed.
-
-Lemma eqb_complete : forall x y, x == y -> eq_bool x y = true.
-Proof. now apply eqb_eq. Qed.
-
-(** Let's implement [HasMinMax] *)
-
-Lemma max_l : forall x y, y<=x -> max x y == x.
-Proof. intros x y. qify. apply Qminmax.Q.max_l. Qed.
-
-Lemma max_r : forall x y, x<=y -> max x y == y.
-Proof. intros x y. qify. apply Qminmax.Q.max_r. Qed.
-
-Lemma min_l : forall x y, x<=y -> min x y == x.
-Proof. intros x y. qify. apply Qminmax.Q.min_l. Qed.
-
-Lemma min_r : forall x y, y<=x -> min x y == y.
-Proof. intros x y. qify. apply Qminmax.Q.min_r. Qed.
-
-(** Q is a ring *)
-
-Lemma add_0_l : forall x, 0+x == x.
-Proof. intros. qify. apply Qplus_0_l. Qed.
-
-Lemma add_comm : forall x y, x+y == y+x.
-Proof. intros. qify. apply Qplus_comm. Qed.
-
-Lemma add_assoc : forall x y z, x+(y+z) == x+y+z.
-Proof. intros. qify. apply Qplus_assoc. Qed.
-
-Lemma mul_1_l : forall x, 1*x == x.
-Proof. intros. qify. apply Qmult_1_l. Qed.
-
-Lemma mul_comm : forall x y, x*y == y*x.
-Proof. intros. qify. apply Qmult_comm. Qed.
-
-Lemma mul_assoc : forall x y z, x*(y*z) == x*y*z.
-Proof. intros. qify. apply Qmult_assoc. Qed.
-
-Lemma mul_add_distr_r : forall x y z, (x+y)*z == x*z + y*z.
-Proof. intros. qify. apply Qmult_plus_distr_l. Qed.
-
-Lemma sub_add_opp : forall x y, x-y == x+(-y).
-Proof. intros. qify. now unfold Qminus. Qed.
-
-Lemma add_opp_diag_r : forall x, x+(-x) == 0.
-Proof. intros. qify. apply Qplus_opp_r. Qed.
-
-(** Q is a field *)
-
-Lemma neq_1_0 : 1!=0.
-Proof. intros. qify. apply Q_apart_0_1. Qed.
-
-Lemma div_mul_inv : forall x y, x/y == x*(/y).
-Proof. intros. qify. now unfold Qdiv. Qed.
-
-Lemma mul_inv_diag_l : forall x, x!=0 -> /x * x == 1.
-Proof. intros x. qify. rewrite Qmult_comm. apply Qmult_inv_r. Qed.
-
-End QProperties.
-
-Module QTypeExt (Q : QType)
- <: QType <: TotalOrder <: HasCompare Q <: HasMinMax Q <: HasEqBool Q
- := Q <+ QProperties.