diff options
467 files changed, 11585 insertions, 17049 deletions
diff --git a/.gitignore b/.gitignore index 84b9844a5..58e1d346c 100644 --- a/.gitignore +++ b/.gitignore @@ -52,6 +52,7 @@ config/Info-*.plist dev/ocamldebug-coq dev/camlp4.dbg plugins/micromega/csdpcert +plugins/micromega/.micromega.ml.generated kernel/byterun/dllcoqrun.so coqdoc.sty .csdp.cache @@ -72,6 +73,8 @@ test-suite/coq-makefile/*/mlihtml test-suite/coq-makefile/*/subdir/done test-suite/coq-makefile/latex1/all.pdf test-suite/coq-makefile/merlin1/.merlin +test-suite/coq-makefile/plugin-reach-outside-API-and-fail/_CoqProject +test-suite/coq-makefile/plugin-reach-outside-API-and-succeed-by-bypassing-the-API/_CoqProject # documentation @@ -140,13 +143,14 @@ plugins/ltac/extraargs.ml plugins/ltac/profile_ltac_tactics.ml ide/coqide_main.ml plugins/ssrmatching/ssrmatching.ml +plugins/ssr/ssrparser.ml +plugins/ssr/ssrvernac.ml # other auto-generated files 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 @@ -173,9 +177,9 @@ dev/myinclude user-contrib .*.sw* +.#* test-suite/.lia.cache test-suite/.nra.cache -# these files are generated from plugins/micromega/MExtraction.v -plugins/micromega/micromega.ml -plugins/micromega/micromega.mli +plugins/ssr/ssrparser.ml +plugins/ssr/ssrvernac.ml diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index a6a27194a..1de9e7f7c 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -241,10 +241,7 @@ validate:32bit: COMPILER: "$COMPILER_32BIT" EXTRA_PACKAGES: "gcc-multilib" -ci-bedrock-src: - <<: *ci-template - -ci-bedrock-facade: +ci-bignums: <<: *ci-template ci-color: @@ -36,6 +36,8 @@ S vernac B vernac S plugins/ltac B plugins/ltac +S API +B API S tools B tools diff --git a/.travis.yml b/.travis.yml index e79498124..e7082a9ee 100644 --- a/.travis.yml +++ b/.travis.yml @@ -37,10 +37,10 @@ env: - TEST_TARGET="test-suite" COMPILER="4.02.3+32bit" - TEST_TARGET="validate" TW="travis_wait" - TEST_TARGET="validate" COMPILER="4.02.3+32bit" TW="travis_wait" - - TEST_TARGET="ci-bedrock-src" - - TEST_TARGET="ci-bedrock-facade" + - TEST_TARGET="ci-bignums" - TEST_TARGET="ci-color" - TEST_TARGET="ci-compcert" + - TEST_TARGET="ci-coq-dpdgraph" EXTRA_OPAM="ocamlgraph" - TEST_TARGET="ci-coquelicot" - TEST_TARGET="ci-geocoq" - TEST_TARGET="ci-fiat-crypto" @@ -62,7 +62,9 @@ env: matrix: allow_failures: + - env: TEST_TARGET="ci-coq-dpdgraph" EXTRA_OPAM="ocamlgraph" - env: TEST_TARGET="ci-geocoq" + - env: TEST_TARGET="ci-fiat-parsers" include: # Full Coq test-suite with two compilers @@ -155,7 +157,7 @@ script: - set -e - echo 'Configuring Coq...' && echo -en 'travis_fold:start:coq.config\\r' -- ./configure -local -usecamlp5 -native-compiler ${NATIVE_COMP} ${EXTRA_CONF} +- ./configure -local -native-compiler ${NATIVE_COMP} ${EXTRA_CONF} - echo -en 'travis_fold:end:coq.config\\r' - echo 'Building Coq...' && echo -en 'travis_fold:start:coq.build\\r' diff --git a/API/API.ml b/API/API.ml new file mode 100644 index 000000000..2b7bbd561 --- /dev/null +++ b/API/API.ml @@ -0,0 +1,215 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +module Ppvernac = Ppvernac +module Command = Command +module States = States +module Kindops = Kindops +module Coq_config = Coq_config +module Esubst = Esubst +module Evar = Evar +module Constrexpr = Constrexpr +module Libobject = Libobject +module Evd = Evd +module Libnames = Libnames +module Nameops = Nameops +module Topfmt = Topfmt +module Locus = Locus +module Locusops = Locusops +module Lemmas = Lemmas +module Clenv = Clenv +module Elimschemes = Elimschemes +module Classes = Classes +module Class_tactics = Class_tactics +module Eauto = Eauto +module Keys = Keys +module Vernac_classifier = Vernac_classifier +module Autorewrite = Autorewrite +module Redops = Redops +module Elim = Elim +module Geninterp = Geninterp +module Obligations = Obligations +module Retroknowledge = Retroknowledge +module Evar_refiner = Evar_refiner +module Hipattern = Hipattern +module Auto = Auto +module Hints = Hints +module Contradiction = Contradiction +module Tacticals = Tacticals +module Tactics = Tactics +module Inv = Inv +module Leminv = Leminv +module Equality = Equality +module Redexpr = Redexpr +module Pfedit = Pfedit +module Stm = Stm +module Stateid = Stateid +module Declaremods = Declaremods +module Miscops = Miscops +module Miscprint = Miscprint +module Genprint = Genprint +module Ppconstr = Ppconstr +module Pputils = Pputils +module Extend = Extend +module Logic = Logic +module Himsg = Himsg +module Tacred = Tacred +module Names = Names +module Indrec = Indrec +module Glob_ops = Glob_ops +module Constrexpr_ops = Constrexpr_ops +module Eqdecide = Eqdecide +module Genredexpr = Genredexpr +module Detyping = Detyping +module Tactypes = Tactypes +module ExplainErr = ExplainErr +module Printer = Printer +module Constrextern = Constrextern +module Locality = Locality +module Impargs = Impargs +module Termops = Termops +module Refiner = Refiner +module Ppextend = Ppextend +module Nametab = Nametab +module Vernacentries = Vernacentries +module Mltop = Mltop +module Goal = Goal +module Proof_global = Proof_global +module Proof = Proof +module Smartlocate = Smartlocate +module Dumpglob = Dumpglob +module Constrintern = Constrintern +module Topconstr = Topconstr +module Notation_ops = Notation_ops +module Patternops = Patternops +module Mod_typing = Mod_typing +module Modops = Modops +module Opaqueproof = Opaqueproof +module Ind_tables = Ind_tables +module Typeops = Typeops +module Inductive = Inductive +module Vars = Vars +module Reduction = Reduction +module Mod_subst = Mod_subst +module Sorts = Sorts +module Univ = Univ +module Constr = Constr +module CClosure = CClosure +module Type_errors = Type_errors +module Safe_typing = Safe_typing +module UGraph = UGraph +module Namegen = Namegen +module Ftactic = Ftactic +module UState = UState +module Proofview_monad = Proofview_monad +module Classops = Classops +module Global = Global +module Goptions = Goptions +module Lib = Lib +module Library = Library +module Summary = Summary +module Universes = Universes +module Declare = Declare +module Refine = Refine +module Find_subterm = Find_subterm +module Evar_kinds = Evar_kinds +module Decl_kinds = Decl_kinds +module Misctypes = Misctypes +module Pattern = Pattern +module Vernacexpr = Vernacexpr +module Search = Search +module Notation_term = Notation_term +module Reductionops = Reductionops +module Inductiveops = Inductiveops +module Recordops = Recordops +module Retyping = Retyping +module Typing = Typing +module Evarsolve = Evarsolve +module Constr_matching = Constr_matching +module Pretyping = Pretyping +module Evarconv = Evarconv +module Unification = Unification +module Typeclasses = Typeclasses +module Pretype_errors = Pretype_errors +module Notation = Notation +module Declarations = Declarations +module Declareops = Declareops +module Globnames = Globnames +module Environ = Environ +module Term = Term +module Coqlib = Coqlib +module Glob_term = Glob_term +module Context = Context +module Stdarg = Stdarg +module Tacmach = Tacmach +module Proofview = Proofview +module Evarutil = Evarutil +module EConstr = EConstr + +module Prelude = + struct + type global_reference = Globnames.global_reference + type metavariable = int + type meta_value_map = (metavariable * Constr.constr) list + type named_context_val = Environ.named_context_val + type conv_pb = Reduction.conv_pb = + | CONV + | CUMUL + type constr = Constr.constr + type types = Constr.types + type evar = Constr.existential_key + type 'constr pexistential = 'constr Constr.pexistential + type env = Environ.env + type evar_map = Evd.evar_map + type rigid = Evd.rigid = + | UnivRigid + | UnivFlexible of bool + type reference = Libnames.reference = + | Qualid of Libnames.qualid Loc.located + | Ident of Names.Id.t Loc.located + end + +(* NOTE: It does not make sense to replace the following "module expression" + simply with "module Proof_type = Proof_type" because + there is only "kernel/entries.mli"; + there is no "kernel/entries.ml" file *) +module Entries = + struct + type mutual_inductive_entry = Entries.mutual_inductive_entry + type inline = int option + type 'a proof_output = Constr.constr Univ.in_universe_context_set * 'a + type 'a const_entry_body = 'a proof_output Future.computation + type 'a definition_entry = 'a Entries.definition_entry = + { const_entry_body : 'a const_entry_body; + const_entry_secctx : Context.Named.t option; + const_entry_feedback : Stateid.t option; + const_entry_type : Term.types option; + const_entry_polymorphic : bool; + const_entry_universes : Univ.universe_context; + const_entry_opaque : bool; + const_entry_inline_code : bool } + type parameter_entry = Entries.parameter_entry + type projection_entry = Entries.projection_entry + type 'a constant_entry = 'a Entries.constant_entry = + | DefinitionEntry of 'a definition_entry + | ParameterEntry of parameter_entry + | ProjectionEntry of projection_entry + end + +(* NOTE: It does not make sense to replace the following "module expression" + simply with "module Proof_type = Proof_type" because + there is only "proofs/proof_type.mli"; + there is no "proofs/proof_type.ml" file *) +module Proof_type = + struct + type goal = Goal.goal + type tactic = goal Evd.sigma -> goal list Evd.sigma + type rule = Proof_type.prim_rule = + | Cut of bool * bool * Names.Id.t * Term.types + | Refine of Term.constr + end diff --git a/API/API.mli b/API/API.mli new file mode 100644 index 000000000..69278e7c9 --- /dev/null +++ b/API/API.mli @@ -0,0 +1,4784 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +module Prelude : +sig + (* None of the items in this modules are meant to be used by plugin-writers. + This module is here only for "technical reasons" + (it will disappear when we take advantage of mutually-recursive modules) *) + + (* API.Term.constr *) + type constr = Constr.t + + (* API.Term.types *) + type types = Constr.t + + (* API.Evar.t *) + type evar = Evar.t + + (* 'constr API.Term.pexistential *) + type 'constr pexistential = evar * 'constr array + + (* API.Environ.env *) + type env = Environ.env + + (* API.Evar.Map.t *) + type evar_map = Evd.evar_map + + (* API.Globnames.global_reference *) + type global_reference = Globnames.global_reference + + type rigid = Evd.rigid = + | UnivRigid + | UnivFlexible of bool + + type conv_pb = Reduction.conv_pb = + | CONV + | CUMUL + + type named_context_val = Environ.named_context_val + + type metavariable = int + + (* Termops.meta_value_map *) + type meta_value_map = (metavariable * constr) list + + (* API.Libnames.reference *) + type reference = Libnames.reference = + | Qualid of Libnames.qualid Loc.located + | Ident of Names.Id.t Loc.located +end + +module Univ : +sig + module Level : + sig + type t = Univ.Level.t + val set : t + val pr : t -> Pp.std_ppcmds + end + + module Instance : + sig + type t = Univ.Instance.t + val empty : t + val of_array : Level.t array -> t + val to_array : t -> Level.t array + val pr : (Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds + end + type 'a puniverses = 'a * Instance.t + + module Constraint : module type of struct include Univ.Constraint end + + type 'a constrained = 'a * Constraint.t + + module UContext : + sig + type t = Univ.UContext.t + val empty : t + end + + type universe_context = UContext.t + [@@ocaml.deprecated "alias of API.Names.UContext.t"] + + module LSet : module type of struct include Univ.LSet end + module ContextSet : + sig + type t = Univ.ContextSet.t + val empty : t + val of_context : UContext.t -> t + val to_context : t -> UContext.t + end + + type 'a in_universe_context_set = 'a * ContextSet.t + type 'a in_universe_context = 'a * UContext.t + type constraint_type = Univ.constraint_type + + module Universe : + sig + type t = Univ.Universe.t + val pr : t -> Pp.std_ppcmds + end + + type universe_context_set = ContextSet.t + [@@ocaml.deprecated "alias of API.Names.ContextSet.t"] + + type universe_set = LSet.t + [@@ocaml.deprecated "alias of API.Names.LSet.t"] + + type 'a constraint_function = 'a -> 'a -> Constraint.t -> Constraint.t + type universe_subst = Univ.universe_subst + type universe_level_subst = Univ.universe_level_subst + + val enforce_leq : Universe.t constraint_function + val pr_uni : Universe.t -> Pp.std_ppcmds + val pr_universe_context : (Level.t -> Pp.std_ppcmds) -> UContext.t -> Pp.std_ppcmds + val pr_universe_context_set : (Level.t -> Pp.std_ppcmds) -> ContextSet.t -> Pp.std_ppcmds + val pr_universe_subst : universe_subst -> Pp.std_ppcmds + val pr_universe_level_subst : universe_level_subst -> Pp.std_ppcmds + val pr_constraints : (Level.t -> Pp.std_ppcmds) -> Constraint.t -> Pp.std_ppcmds +end + +module UState : +sig + type t = UState.t + val context : t -> Univ.UContext.t + val context_set : t -> Univ.ContextSet.t + val of_context_set : Univ.ContextSet.t -> t +end + +module Sorts : +sig + type contents = Sorts.contents = Pos | Null + type t = Sorts.t = + | Prop of contents + | Type of Univ.Universe.t + val is_prop : t -> bool + val hash : t -> int + + type family = Sorts.family = InProp | InSet | InType + val family : t -> family +end + +module Names : +sig + module Id : module type of struct include Names.Id end + + module MBId : sig + type t = Names.MBId.t + val equal : t -> t -> bool + val to_id : t -> Names.Id.t + val repr : t -> int * Names.Id.t * Names.DirPath.t + val debug_to_string : t -> string + end + + type evaluable_global_reference = Names.evaluable_global_reference = + | EvalVarRef of Id.t + | EvalConstRef of Names.Constant.t + + module Name : module type of struct include Names.Name end + + type name = Name.t = + | Anonymous + | Name of Id.t + [@@ocaml.deprecated "alias of API.Name.t"] + + module DirPath : + sig + type t = Names.DirPath.t + val empty : t + val make : Id.t list -> t + val repr : t -> Id.t list + val equal : t -> t -> bool + val to_string : t -> string + end + + module Label : + sig + type t = Names.Label.t + val make : string -> t + val equal : t -> t -> bool + val compare : t -> t -> int + val of_id : Names.Id.t -> t + val to_id : t -> Names.Id.t + val to_string : t -> string + end + + module ModPath : + sig + type t = Names.ModPath.t = + | MPfile of Names.DirPath.t + | MPbound of MBId.t + | MPdot of t * Label.t + val compare : t -> t -> int + val equal : t -> t -> bool + val hash : t -> int + val initial : t + val to_string : t -> string + val debug_to_string : t -> string + end + + module KerName : + sig + type t = Names.KerName.t + val make : ModPath.t -> DirPath.t -> Label.t -> t + val make2 : ModPath.t -> Label.t -> t + val modpath : t -> ModPath.t + val equal : t -> t -> bool + val compare : t -> t -> int + val label : t -> Label.t + val repr : t -> ModPath.t * DirPath.t * Label.t + val print : t -> Pp.std_ppcmds + val to_string : t -> string + end + + type kernel_name = KerName.t + [@@ocaml.deprecated "alias of API.Names.KerName.t"] + + module Constant : + sig + type t = Names.Constant.t + val equal : t -> t -> bool + val make1 : Names.KerName.t -> t + val make2 : Names.ModPath.t -> Label.t -> t + val make3 : Names.ModPath.t -> Names.DirPath.t -> Label.t -> t + val repr3 : t -> Names.ModPath.t * Names.DirPath.t * Label.t + val canonical : t -> Names.KerName.t + val user : t -> Names.KerName.t + val label : t -> Label.t + end + + module MutInd : + sig + type t = Names.MutInd.t + val make1 : Names.KerName.t -> t + val make2 : Names.ModPath.t -> Label.t -> t + val equal : t -> t -> bool + val repr3 : t -> Names.ModPath.t * Names.DirPath.t * Label.t + val canonical : t -> Names.KerName.t + val modpath : t -> Names.ModPath.t + val label : t -> Label.t + val user : t -> Names.KerName.t + val print : t -> Pp.std_ppcmds + end + + module Projection : + sig + type t = Names.Projection.t + val make : Constant.t -> bool -> t + val map : (Constant.t -> Constant.t) -> t -> t + val constant : t -> Constant.t + val equal : t -> t -> bool + end + + type inductive = MutInd.t * int + val eq_ind : inductive -> inductive -> bool + + type constructor = inductive * int + val eq_constructor : constructor -> constructor -> bool + val constructor_hash : constructor -> int + + module MPset : module type of struct include Names.MPset end + module MPmap : module type of struct include Names.MPmap end + module KNset : module type of struct include Names.KNset end + module KNmap : module type of struct include Names.KNmap end + module Cset : module type of struct include Names.Cset end + module Cset_env : module type of struct include Names.Cset_env end + module Cmap : module type of struct include Names.Cmap end + module Cmap_env : module type of struct include Names.Cmap_env end + module Cpred : module type of struct include Names.Cpred end + module Mindset : module type of struct include Names.Mindset end + module Mindmap : module type of struct include Names.Mindmap end + module Mindmap_env : module type of struct include Names.Mindmap_env end + module Indmap : module type of struct include Names.Indmap end + with type key = inductive + module Indmap_env : module type of struct include Names.Indmap_env end + module Constrmap : module type of struct include Names.Constrmap end + module Constrmap_env : module type of struct include Names.Constrmap_env end + + type transparent_state = Id.Pred.t * Cpred.t + val empty_transparent_state : transparent_state + val full_transparent_state : transparent_state + val var_full_transparent_state : transparent_state + val cst_full_transparent_state : transparent_state + + val pr_kn : KerName.t -> Pp.std_ppcmds + [@@ocaml.deprecated "alias of API.Names.KerName.print"] + + val eq_constant : Constant.t -> Constant.t -> bool + [@@ocaml.deprecated "alias of API.Names.Constant.equal"] + + type module_path = ModPath.t = + | MPfile of DirPath.t + | MPbound of MBId.t + | MPdot of ModPath.t * Label.t + [@@ocaml.deprecated "alias of API.Names.ModPath.t"] + + type variable = Id.t + [@@ocaml.deprecated "alias of API.Names.Id.t"] + + type 'a tableKey = 'a Names.tableKey = + | ConstKey of 'a + | VarKey of Id.t + | RelKey of Int.t + + val id_of_string : string -> Id.t + [@@ocaml.deprecated "alias of API.Names.Id.of_string"] + + val string_of_id : Id.t -> string + [@@ocaml.deprecated "alias of API.Names.Id.to_string"] + + type mutual_inductive = MutInd.t + [@@ocaml.deprecated "alias of API.Names.MutInd.t"] + + val eq_mind : MutInd.t -> MutInd.t -> bool + [@@ocaml.deprecated "alias of API.Names.MutInd.equal"] + + val repr_con : Constant.t -> ModPath.t * DirPath.t * Label.t + [@@ocaml.deprecated "alias of API.Names.Constant.repr3"] + + val repr_mind : MutInd.t -> ModPath.t * DirPath.t * Label.t + [@@ocaml.deprecated "alias of API.Names.MutInd.repr3"] + + val initial_path : ModPath.t + [@@ocaml.deprecated "alias of API.Names.ModPath.initial"] + + val con_label : Constant.t -> Label.t + [@@ocaml.deprecated "alias of API.Names.Constant.label"] + + val mind_label : MutInd.t -> Label.t + [@@ocaml.deprecated "alias of API.Names.MutInd.label"] + + val string_of_mp : ModPath.t -> string + [@@ocaml.deprecated "alias of API.Names.ModPath.to_string"] + + val mind_of_kn : KerName.t -> MutInd.t + [@@ocaml.deprecated "alias of API.Names.MutInd.make1"] + + type constant = Constant.t + [@@ocaml.deprecated "alias of API.Names.Constant.t"] + + val mind_modpath : MutInd.t -> ModPath.t + [@@ocaml.deprecated "alias of API.Names.MutInd.modpath"] + + val canonical_mind : MutInd.t -> KerName.t + [@@ocaml.deprecated "alias of API.Names.MutInd.canonical"] + + val user_mind : MutInd.t -> KerName.t + [@@ocaml.deprecated "alias of API.Names.MutInd.user"] + + val repr_kn : KerName.t -> ModPath.t * DirPath.t * Label.t + [@@ocaml.deprecated "alias of API.Names.KerName.repr"] + + val constant_of_kn : KerName.t -> Constant.t + [@@ocaml.deprecated "alias of API.Names.Constant.make1"] + + val user_con : Constant.t -> KerName.t + [@@ocaml.deprecated "alias of API.Names.Constant.user"] + + val modpath : KerName.t -> ModPath.t + [@@ocaml.deprecated "alias of API.Names.KerName.modpath"] + + val canonical_con : Constant.t -> KerName.t + [@@ocaml.deprecated "alias of API.Names.Constant.canonical"] + + val make_kn : ModPath.t -> DirPath.t -> Label.t -> KerName.t + [@@ocaml.deprecated "alias of API.Names.KerName.make"] + + val make_con : ModPath.t -> DirPath.t -> Label.t -> Constant.t + [@@ocaml.deprecated "alias of API.Names.Constant.make3"] + + val debug_pr_con : Constant.t -> Pp.std_ppcmds + + val debug_pr_mind : MutInd.t -> Pp.std_ppcmds + + val pr_con : Constant.t -> Pp.std_ppcmds + + val string_of_con : Constant.t -> string + + val string_of_mind : MutInd.t -> string + + val debug_string_of_mind : MutInd.t -> string + + val debug_string_of_con : Constant.t -> string + + module Idset : module type of struct include Id.Set end +end + +module Context : +sig + + module Rel : + sig + module Declaration : + sig + (* local declaration *) + (* local declaration *) + type ('constr, 'types) pt = ('constr, 'types) Context.Rel.Declaration.pt = + | LocalAssum of Names.Name.t * 'types (** name, type *) + | LocalDef of Names.Name.t * 'constr * 'types (** name, value, type *) + + type t = (Prelude.constr, Prelude.types) pt + + (** Return the name bound by a given declaration. *) + val get_name : ('c, 't) pt -> Names.Name.t + + (** Return the type of the name bound by a given declaration. *) + val get_type : ('c, 't) pt -> 't + + (** Set the name that is bound by a given declaration. *) + val set_name : Names.Name.t -> ('c, 't) pt -> ('c, 't) pt + + (** Set the type of the bound variable in a given declaration. *) + val set_type : 't -> ('c, 't) pt -> ('c, 't) pt + + (** Return [true] iff a given declaration is a local assumption. *) + val is_local_assum : ('c, 't) pt -> bool + + (** Return [true] iff a given declaration is a local definition. *) + val is_local_def : ('c, 't) pt -> bool + + (** Check whether the two given declarations are equal. *) + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool + + (** Map the name bound by a given declaration. *) + val map_name : (Names.Name.t -> Names.Name.t) -> ('c, 't) pt -> ('c, 't) pt + + (** For local assumptions, this function returns the original local assumptions. + For local definitions, this function maps the value in the local definition. *) + val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt + + (** Map the type of the name bound by a given declaration. *) + val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt + + (** Map all terms in a given declaration. *) + val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt + + (** Perform a given action on all terms in a given declaration. *) + val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit + + (** Reduce all terms in a given declaration to a single value. *) + val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a + end + + (** Rel-context is represented as a list of declarations. + Inner-most declarations are at the beginning of the list. + Outer-most declarations are at the end of the list. *) + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list + type t = Declaration.t list + + (** empty rel-context *) + val empty : ('c, 't) pt + + (** Return a new rel-context enriched by with a given inner-most declaration. *) + val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt + + (** Return the number of {e local declarations} in a given context. *) + val length : ('c, 't) pt -> int + + (** Check whether given two rel-contexts are equal. *) + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool + + (** Return the number of {e local assumptions} in a given rel-context. *) + val nhyps : ('c, 't) pt -> int + + (** Return a declaration designated by a given de Bruijn index. + @raise Not_found if the designated de Bruijn index outside the range. *) + val lookup : int -> ('c, 't) pt -> ('c, 't) Declaration.pt + + (** Map all terms in a given rel-context. *) + val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt + + (** Perform a given action on every declaration in a given rel-context. *) + val iter : ('c -> unit) -> ('c, 'c) pt -> unit + + (** Reduce all terms in a given rel-context to a single value. + Innermost declarations are processed first. *) + val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a + + (** Reduce all terms in a given rel-context to a single value. + Outermost declarations are processed first. *) + val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a + + (** [extended_vect n Γ] does the same, returning instead an array. *) + val to_extended_vect : (int -> 'r) -> int -> ('c, 't) pt -> 'r array + end + module Named : + sig + module Declaration : + sig + (** local declaration *) + type ('constr, 'types) pt = ('constr, 'types) Context.Named.Declaration.pt = + | LocalAssum of Names.Id.t * 'types (** identifier, type *) + | LocalDef of Names.Id.t * 'constr * 'types (** identifier, value, type *) + + type t = (Prelude.constr, Prelude.types) pt + + (** Return the identifier bound by a given declaration. *) + val get_id : ('c, 't) pt -> Names.Id.t + + (** Return the type of the name bound by a given declaration. *) + val get_type : ('c, 't) pt -> 't + + (** Set the identifier that is bound by a given declaration. *) + val set_id : Names.Id.t -> ('c, 't) pt -> ('c, 't) pt + + (** Set the type of the bound variable in a given declaration. *) + val set_type : 't -> ('c, 't) pt -> ('c, 't) pt + + (** Return [true] iff a given declaration is a local assumption. *) + val is_local_assum : ('c, 't) pt -> bool + + (** Return [true] iff a given declaration is a local definition. *) + val is_local_def : ('c, 't) pt -> bool + + (** Check whether any term in a given declaration satisfies a given predicate. *) + val exists : ('c -> bool) -> ('c, 'c) pt -> bool + + (** Check whether all terms in a given declaration satisfy a given predicate. *) + val for_all : ('c -> bool) -> ('c, 'c) pt -> bool + + (** Check whether the two given declarations are equal. *) + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool + + (** Map the identifier bound by a given declaration. *) + val map_id : (Names.Id.t -> Names.Id.t) -> ('c, 't) pt -> ('c, 't) pt + + (** For local assumptions, this function returns the original local assumptions. + For local definitions, this function maps the value in the local definition. *) + val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt + + (** Map the type of the name bound by a given declaration. *) + val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt + + (** Map all terms in a given declaration. *) + val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt + + (** Perform a given action on all terms in a given declaration. *) + val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit + + (** Reduce all terms in a given declaration to a single value. *) + val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a + + val to_rel_decl : ('c, 't) pt -> ('c, 't) Rel.Declaration.pt + end + (** Named-context is represented as a list of declarations. + Inner-most declarations are at the beginning of the list. + Outer-most declarations are at the end of the list. *) + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list + type t = Declaration.t list + + (** empty named-context *) + val empty : ('c, 't) pt + + (** Return a new named-context enriched by with a given inner-most declaration. *) + val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt + + (** Return the number of {e local declarations} in a given named-context. *) + val length : ('c, 't) pt -> int + + (** Return a declaration designated by an identifier of the variable bound in that declaration. + @raise Not_found if the designated identifier is not bound in a given named-context. *) + val lookup : Names.Id.t -> ('c, 't) pt -> ('c, 't) Declaration.pt + + (** Check whether given two named-contexts are equal. *) + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool + + (** Map all terms in a given named-context. *) + val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt + + (** Perform a given action on every declaration in a given named-context. *) + val iter : ('c -> unit) -> ('c, 'c) pt -> unit + + (** Reduce all terms in a given named-context to a single value. + Innermost declarations are processed first. *) + val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a + + (** Reduce all terms in a given named-context to a single value. + Outermost declarations are processed first. *) + val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a + + (** Return the set of all identifiers bound in a given named-context. *) + val to_vars : ('c, 't) pt -> Names.Id.Set.t + + (** [to_instance Ω] builds an instance [args] such + that [Ω ⊢ args:Ω] where [Ω] is a named-context and with the local + definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it + gives [Var id1, Var id3]. All [idj] are supposed distinct. *) + val to_instance : (Names.Id.t -> 'r) -> ('c, 't) pt -> 'r list + end +end + +module Term : +sig + type sorts_family = Sorts.family = InProp | InSet | InType + [@@deprecated "alias of API.Sorts.family"] + + type metavariable = Prelude.metavariable + + type contents = Sorts.contents = Pos | Null + + type sorts = Sorts.t = + | Prop of contents + | Type of Univ.Universe.t + [@@ocaml.deprecated "alias of API.Sorts.t"] + + type constr = Prelude.constr + type types = Prelude.types + type ('constr, 'types) prec_declaration = Names.Name.t array * 'types array * 'constr array + type 'constr pexistential = 'constr Prelude.pexistential + type cast_kind = Term.cast_kind = + | VMcast + | NATIVEcast + | DEFAULTcast + | REVERTcast + type 'a puniverses = 'a Univ.puniverses + type pconstant = Names.Constant.t puniverses + type pinductive = Names.inductive puniverses + type pconstructor = Names.constructor puniverses + type case_style = Term.case_style = + | LetStyle + | IfStyle + | LetPatternStyle + | MatchStyle + | RegularStyle + type case_printing = Term.case_printing = + { ind_tags : bool list; + cstr_tags : bool list array; + style : case_style + } + type case_info = Term.case_info = + { ci_ind : Names.inductive; + ci_npar : int; + ci_cstr_ndecls : int array; + ci_cstr_nargs : int array; + ci_pp_info : case_printing + } + type ('constr, 'types) pfixpoint = + (int array * int) * ('constr, 'types) prec_declaration + type ('constr, 'types) pcofixpoint = + int * ('constr, 'types) prec_declaration + type ('constr, 'types, 'sort, 'univs) kind_of_term = ('constr, 'types, 'sort, 'univs) Term.kind_of_term = + | Rel of int + | Var of Names.Id.t + | Meta of metavariable + | Evar of 'constr pexistential + | Sort of 'sort + | Cast of 'constr * cast_kind * 'types + | Prod of Names.Name.t * 'types * 'types + | Lambda of Names.Name.t * 'types * 'constr + | LetIn of Names.Name.t * 'constr * 'types * 'constr + | App of 'constr * 'constr array + | Const of (Names.Constant.t * 'univs) + | Ind of (Names.inductive * 'univs) + | Construct of (Names.constructor * 'univs) + | Case of case_info * 'constr * 'constr * 'constr array + | Fix of ('constr, 'types) pfixpoint + | CoFix of ('constr, 'types) pcofixpoint + | Proj of Names.Projection.t * 'constr + type existential = Prelude.evar * constr array + type rec_declaration = Names.Name.t array * constr array * constr array + type fixpoint = (int array * int) * rec_declaration + type cofixpoint = int * rec_declaration + val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term + val applistc : constr -> constr list -> constr + + val applist : constr * constr list -> constr + [@@ocaml.deprecated "(sort of an) alias of API.Term.applistc"] + + val mkArrow : types -> types -> constr + val mkRel : int -> constr + val mkVar : Names.Id.t -> constr + + val mkMeta : Prelude.metavariable -> constr + + val mkEvar : existential -> constr + val mkSort : Sorts.t -> types + val mkProp : types + val mkSet : types + val mkType : Univ.Universe.t -> types + val mkCast : constr * cast_kind * constr -> constr + val mkProd : Names.Name.t * types * types -> types + val mkLambda : Names.Name.t * types * constr -> constr + val mkLetIn : Names.Name.t * constr * types * constr -> constr + val mkApp : constr * constr array -> constr + val mkConst : Names.Constant.t -> constr + val mkProj : Names.Projection.t * constr -> constr + val mkInd : Names.inductive -> constr + val mkConstruct : Names.constructor -> constr + val mkConstructU : Names.constructor puniverses -> constr + val mkConstructUi : (pinductive * int) -> constr + val mkCase : case_info * constr * constr * constr array -> constr + val mkFix : fixpoint -> constr + val mkCoFix : cofixpoint -> constr + val mkNamedLambda : Names.Id.t -> types -> constr -> constr + val mkNamedLetIn : Names.Id.t -> constr -> types -> constr -> constr + val mkNamedProd : Names.Id.t -> types -> types -> types + + val decompose_app : constr -> constr * constr list + val decompose_prod : constr -> (Names.Name.t*constr) list * constr + val decompose_prod_n : int -> constr -> (Names.Name.t * constr) list * constr + val decompose_prod_assum : types -> Context.Rel.t * types + val decompose_lam : constr -> (Names.Name.t * constr) list * constr + val decompose_lam_n : int -> constr -> (Names.Name.t * constr) list * constr + val decompose_prod_n_assum : int -> types -> Context.Rel.t * types + + val compose_prod : (Names.Name.t * constr) list -> constr -> constr + val compose_lam : (Names.Name.t * constr) list -> constr -> constr + + val destSort : constr -> Sorts.t + val destVar : constr -> Names.Id.t + val destApp : constr -> constr * constr array + val destProd : types -> Names.Name.t * types * types + val destLetIn : constr -> Names.Name.t * constr * types * constr + val destEvar : constr -> existential + val destRel : constr -> int + val destConst : constr -> Names.Constant.t puniverses + val destCast : constr -> constr * cast_kind * constr + val destLambda : constr -> Names.Name.t * types * constr + + val isRel : constr -> bool + val isVar : constr -> bool + val isEvar : constr -> bool + val isLetIn : constr -> bool + val isLambda : constr -> bool + val isConst : constr -> bool + val isEvar_or_Meta : constr -> bool + val isCast : constr -> bool + val isMeta : constr -> bool + val isApp : constr -> bool + + val fold_constr : ('a -> constr -> 'a) -> 'a -> constr -> 'a + + val eq_constr : constr -> constr -> bool + + val hash_constr : constr -> int + val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr + val it_mkProd_or_LetIn : types -> Context.Rel.t -> types + val prod_applist : constr -> constr list -> constr + exception DestKO + val map_constr : (constr -> constr) -> constr -> constr + + val mkIndU : pinductive -> constr + val mkConstU : pconstant -> constr + val map_constr_with_binders : + ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr + val iter_constr : (constr -> unit) -> constr -> unit + + (* Quotients away universes: really needed? + * Can't we just call eq_c_univs_infer and discard the inferred csts? + *) + val eq_constr_nounivs : constr -> constr -> bool + + type ('constr, 'types) kind_of_type = ('constr, 'types) Term.kind_of_type = + | SortType of Sorts.t + | CastType of 'types * 'types + | ProdType of Names.Name.t * 'types * 'types + | LetInType of Names.Name.t * 'constr * 'types * 'types + | AtomicType of 'constr * 'constr array + val kind_of_type : types -> (constr, types) kind_of_type + + val is_prop_sort : Sorts.t -> bool + [@@ocaml.deprecated "alias of API.Sorts.is_prop"] + + type existential_key = Prelude.evar + + val family_of_sort : Sorts.t -> Sorts.family + + val compare : constr -> constr -> int + + val constr_ord : constr -> constr -> int + [@@ocaml.deprecated "alias of API.Term.compare"] + + val destInd : constr -> Names.inductive puniverses + val univ_of_sort : Sorts.t -> Univ.Universe.t + + val strip_lam : constr -> constr + val strip_prod_assum : types -> types + + val decompose_lam_assum : constr -> Context.Rel.t * constr + val destFix : constr -> fixpoint + + val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool +end + +module EConstr : +sig + type t = EConstr.t + type constr = t + type types = t + type unsafe_judgment = EConstr.unsafe_judgment + type named_declaration = (constr, types) Context.Named.Declaration.pt + type named_context = (constr, types) Context.Named.pt + type rel_context = (constr, types) Context.Rel.pt + type rel_declaration = (constr, types) Context.Rel.Declaration.pt + type existential = constr Term.pexistential + module ESorts : + sig + type t = EConstr.ESorts.t + (** Type of sorts up-to universe unification. Essentially a wrapper around + Sorts.t so that normalization is ensured statically. *) + + val make : Sorts.t -> t + (** Turn a sort into an up-to sort. *) + + val kind : Prelude.evar_map -> t -> Sorts.t + (** Returns the view into the current sort. Note that the kind of a variable + may change if the unification state of the evar map changes. *) + + end + + module EInstance : + sig + type t = EConstr.EInstance.t + (** Type of universe instances up-to universe unification. Similar to + {ESorts.t} for {Univ.Instance.t}. *) + + val make : Univ.Instance.t -> t + val kind : Prelude.evar_map -> t -> Univ.Instance.t + val empty : t + val is_empty : t -> bool + end + + val of_constr : Term.constr -> constr + + val kind : Prelude.evar_map -> constr -> (constr, constr, ESorts.t, EInstance.t) Term.kind_of_term + + val mkArrow : constr -> constr -> constr + val mkInd : Names.inductive -> t + val mkProp : constr + val mkProd : Names.Name.t * constr * constr -> constr + val mkRel : int -> constr + val mkSort : Sorts.t -> constr + val mkVar : Names.Id.t -> constr + val mkLambda : Names.Name.t * constr * constr -> constr + val mkLambda_or_LetIn : rel_declaration -> constr -> constr + val mkApp : constr * constr array -> constr + val mkEvar : constr Term.pexistential -> constr + + val mkMeta : Prelude.metavariable -> constr + + val mkConstructU : Names.constructor * EInstance.t -> constr + val mkLetIn : Names.Name.t * constr * constr * constr -> constr + val mkProd_or_LetIn : rel_declaration -> constr -> constr + val mkCast : constr * Term.cast_kind * constr -> constr + val mkNamedLambda : Names.Id.t -> types -> constr -> constr + val mkNamedProd : Names.Id.t -> types -> types -> types + + val isCast : Evd.evar_map -> t -> bool + val isEvar : Prelude.evar_map -> constr -> bool + val isInd : Prelude.evar_map -> constr -> bool + val isRel : Prelude.evar_map -> constr -> bool + val isSort : Prelude.evar_map -> constr -> bool + val isVar : Prelude.evar_map -> constr -> bool + val isConst : Prelude.evar_map -> constr -> bool + val isConstruct : Prelude.evar_map -> constr -> bool + + val destInd : Prelude.evar_map -> constr -> Names.inductive * EInstance.t + val destVar : Prelude.evar_map -> constr -> Names.Id.t + val destEvar : Prelude.evar_map -> constr -> constr Term.pexistential + val destRel : Prelude.evar_map -> constr -> int + val destProd : Prelude.evar_map -> constr -> Names.Name.t * types * types + val destLambda : Prelude.evar_map -> constr -> Names.Name.t * types * constr + val destApp : Prelude.evar_map -> constr -> constr * constr array + val destConst : Prelude.evar_map -> constr -> Names.Constant.t * EInstance.t + val destConstruct : Prelude.evar_map -> constr -> Names.constructor * EInstance.t + val destFix : Evd.evar_map -> t -> (t, t) Term.pfixpoint + val destCast : Evd.evar_map -> t -> t * Term.cast_kind * t + + val mkConstruct : Names.constructor -> constr + + val compose_lam : (Names.Name.t * constr) list -> constr -> constr + + val decompose_lam : Prelude.evar_map -> constr -> (Names.Name.t * constr) list * constr + val decompose_lam_n_assum : Prelude.evar_map -> int -> constr -> rel_context * constr + val decompose_app : Prelude.evar_map -> constr -> constr * constr list + val decompose_prod : Prelude.evar_map -> constr -> (Names.Name.t * constr) list * constr + val decompose_prod_assum : Prelude.evar_map -> constr -> rel_context * constr + + val applist : constr * constr list -> constr + + val to_constr : Prelude.evar_map -> constr -> Constr.t + + val push_rel : rel_declaration -> Prelude.env -> Prelude.env + + module Unsafe : + sig + val to_constr : constr -> Term.constr + + val to_rel_decl : (constr, types) Context.Rel.Declaration.pt -> (Prelude.constr, Prelude.types) Context.Rel.Declaration.pt + + (** Physical identity. Does not care for defined evars. *) + + val to_named_decl : (constr, types) Context.Named.Declaration.pt -> (Prelude.constr, Prelude.types) Context.Named.Declaration.pt + + val to_instance : EInstance.t -> Univ.Instance.t + end + + module Vars : + sig + val substnl : t list -> int -> t -> t + val noccurn : Prelude.evar_map -> int -> constr -> bool + val closed0 : Prelude.evar_map -> constr -> bool + val subst1 : constr -> constr -> constr + val substl : constr list -> constr -> constr + val lift : int -> constr -> constr + val liftn : int -> int -> t -> t + val subst_var : Names.Id.t -> t -> t + val subst_vars : Names.Id.t list -> t -> t + end + + val fresh_global : + ?loc:Loc.t -> ?rigid:Prelude.rigid -> ?names:Univ.Instance.t -> Environ.env -> + Evd.evar_map -> Prelude.global_reference -> Evd.evar_map * t + +val of_named_decl : (Term.constr, Term.types) Context.Named.Declaration.pt -> (constr, types) Context.Named.Declaration.pt + val of_rel_decl : (Term.constr, Term.types) Context.Rel.Declaration.pt -> (constr, types) Context.Rel.Declaration.pt + val kind_of_type : Prelude.evar_map -> constr -> (constr, constr) Term.kind_of_type + val to_lambda : Prelude.evar_map -> int -> constr -> constr + val it_mkLambda_or_LetIn : constr -> rel_context -> constr + val push_rel_context : rel_context -> Prelude.env -> Prelude.env + val eq_constr : Prelude.evar_map -> constr -> constr -> bool + val iter_with_binders : Prelude.evar_map -> ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit + val fold : Prelude.evar_map -> ('a -> constr -> 'a) -> 'a -> constr -> 'a + val existential_type : Prelude.evar_map -> existential -> types + val iter : Prelude.evar_map -> (constr -> unit) -> constr -> unit + val eq_constr_universes : Prelude.evar_map -> constr -> constr -> Universes.universe_constraints option + val eq_constr_nounivs : Prelude.evar_map -> constr -> constr -> bool + val compare_constr : Evd.evar_map -> (constr -> constr -> bool) -> constr -> constr -> bool + val isApp : Prelude.evar_map -> constr -> bool + val it_mkProd_or_LetIn : constr -> rel_context -> constr + val push_named : named_declaration -> Prelude.env -> Prelude.env + val destCase : Prelude.evar_map -> constr -> Term.case_info * constr * constr * constr array + val decompose_lam_assum : Prelude.evar_map -> constr -> rel_context * constr + val mkConst : Names.Constant.t -> constr + val mkCase : Term.case_info * constr * constr * constr array -> constr + val named_context : Prelude.env -> named_context + val val_of_named_context : named_context -> Prelude.named_context_val + val mkFix : (t, t) Term.pfixpoint -> t + val decompose_prod_n_assum : Evd.evar_map -> int -> t -> rel_context * t + val isMeta : Evd.evar_map -> t -> bool + + val destMeta : Evd.evar_map -> t -> Term.metavariable + + val map_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> t) -> 'a -> t -> t + val mkNamedLetIn : Names.Id.t -> constr -> types -> constr -> constr + val map : Evd.evar_map -> (t -> t) -> t -> t + val mkConstU : Names.Constant.t * EInstance.t -> t + val isProd : Evd.evar_map -> t -> bool + val mkConstructUi : (Names.inductive * EInstance.t) * int -> t + val isLambda : Evd.evar_map -> t -> bool +end + +module Mod_subst : +sig + type substitution = Mod_subst.substitution + type 'a substituted = 'a Mod_subst.substituted + type delta_resolver = Mod_subst.delta_resolver + + val force_constr : Term.constr substituted -> Term.constr + + val empty_delta_resolver : delta_resolver + val constant_of_delta_kn : delta_resolver -> Names.KerName.t -> Names.Constant.t + val mind_of_delta_kn : delta_resolver -> Names.KerName.t -> Names.MutInd.t + val subst_kn : substitution -> Names.KerName.t -> Names.KerName.t + val subst_evaluable_reference : + substitution -> Names.evaluable_global_reference -> Names.evaluable_global_reference + val subst_mps : substitution -> Term.constr -> Term.constr + val subst_constant : substitution -> Names.Constant.t -> Names.Constant.t + val subst_ind : substitution -> Names.inductive -> Names.inductive + val debug_pr_subst : substitution -> Pp.std_ppcmds + val debug_pr_delta : delta_resolver -> Pp.std_ppcmds +end + +module Retroknowledge : +sig + type action = Retroknowledge.action + type nat_field = Retroknowledge.nat_field = + | NatType + | NatPlus + | NatTimes + type n_field = Retroknowledge.n_field = + | NPositive + | NType + | NTwice + | NTwicePlusOne + | NPhi + | NPhiInv + | NPlus + | NTimes + type int31_field = Retroknowledge.int31_field = + | Int31Bits + | Int31Type + | Int31Constructor + | Int31Twice + | Int31TwicePlusOne + | Int31Phi + | Int31PhiInv + | Int31Plus + | Int31PlusC + | Int31PlusCarryC + | Int31Minus + | Int31MinusC + | Int31MinusCarryC + | Int31Times + | Int31TimesC + | Int31Div21 + | Int31Div + | Int31Diveucl + | Int31AddMulDiv + | Int31Compare + | Int31Head0 + | Int31Tail0 + | Int31Lor + | Int31Land + | Int31Lxor + type field = Retroknowledge.field = + | KInt31 of string * int31_field +end + +module Declarations : +sig + type recarg = Declarations.recarg = + | Norec + | Mrec of Names.inductive + | Imbr of Names.inductive + type wf_paths = recarg Rtree.t + type inline = Declarations.inline + type constant_def = Declarations.constant_def = + | Undef of inline + | Def of Term.constr Mod_subst.substituted + | OpaqueDef of Opaqueproof.opaque + type constant_type = Declarations.constant_type + type constant_universes = Declarations.constant_universes + type projection_body = Declarations.projection_body = { + proj_ind : Names.MutInd.t; + proj_npars : int; + proj_arg : int; + proj_type : Term.types; + proj_eta : Term.constr * Term.types; + proj_body : Term.constr; + } + type typing_flags = Declarations.typing_flags + type constant_body = Declarations.constant_body = { + const_hyps : Context.Named.t; + const_body : constant_def; + const_type : constant_type; + const_body_code : Cemitcodes.to_patch_substituted option; + const_polymorphic : bool; + const_universes : constant_universes; + const_proj : projection_body option; + const_inline_code : bool; + const_typing_flags : typing_flags; + } + type one_inductive_body = Declarations.one_inductive_body = { + mind_typename : Names.Id.t; + mind_arity_ctxt : Context.Rel.t; + mind_arity : Declarations.inductive_arity; + mind_consnames : Names.Id.t array; + mind_user_lc : Term.types array; + mind_nrealargs : int; + mind_nrealdecls : int; + mind_kelim : Sorts.family list; + mind_nf_lc : Term.types array; + mind_consnrealargs : int array; + mind_consnrealdecls : int array; + mind_recargs : wf_paths; + mind_nb_constant : int; + mind_nb_args : int; + mind_reloc_tbl : Cbytecodes.reloc_table; + } + type ('ty,'a) functorize = ('ty,'a) Declarations.functorize = + | NoFunctor of 'a + | MoreFunctor of Names.MBId.t * 'ty * ('ty,'a) functorize + type with_declaration = Declarations.with_declaration = + | WithMod of Names.Id.t list * Names.ModPath.t + | WithDef of Names.Id.t list * Term.constr Univ.in_universe_context + type module_alg_expr = Declarations.module_alg_expr = + | MEident of Names.ModPath.t + | MEapply of module_alg_expr * Names.ModPath.t + | MEwith of module_alg_expr * with_declaration + type mutual_inductive_body = Declarations.mutual_inductive_body = { + mind_packets : one_inductive_body array; + mind_record : Declarations.record_body option; + mind_finite : Decl_kinds.recursivity_kind; + mind_ntypes : int; + mind_hyps : Context.Named.t; + mind_nparams : int; + mind_nparams_rec : int; + mind_params_ctxt : Context.Rel.t; + mind_polymorphic : bool; + mind_universes : Univ.UContext.t; + mind_private : bool option; + mind_typing_flags : Declarations.typing_flags; + } + and module_expression = (module_type_body,module_alg_expr) functorize + and module_implementation = Declarations.module_implementation = + | Abstract + | Algebraic of module_expression + | Struct of module_signature + | FullStruct + and module_body = Declarations.module_body = + { mod_mp : Names.ModPath.t; + mod_expr : module_implementation; + mod_type : module_signature; + mod_type_alg : module_expression option; + mod_constraints : Univ.ContextSet.t; + mod_delta : Mod_subst.delta_resolver; + mod_retroknowledge : Retroknowledge.action list + } + and module_signature = (module_type_body,structure_body) functorize + and module_type_body = module_body + and structure_body = (Names.Label.t * structure_field_body) list + and structure_field_body = Declarations.structure_field_body = + | SFBconst of constant_body + | SFBmind of mutual_inductive_body + | SFBmodule of module_body + | SFBmodtype of module_type_body +end + +module Environ : +sig + type env = Prelude.env + type named_context_val = Prelude.named_context_val + type ('constr, 'types) punsafe_judgment = ('constr, 'types) Environ.punsafe_judgment = + { + uj_val : 'constr; + uj_type : 'types + } + val empty_env : env + val lookup_mind : Names.MutInd.t -> env -> Declarations.mutual_inductive_body + val push_rel : Context.Rel.Declaration.t -> env -> env + val push_rel_context : Context.Rel.t -> env -> env + val push_rec_types : Term.rec_declaration -> env -> env + val lookup_rel : int -> env -> Context.Rel.Declaration.t + val lookup_named : Names.Id.t -> env -> Context.Named.Declaration.t + val lookup_named_val : Names.Id.t -> Environ.named_context_val -> Context.Named.Declaration.t + val lookup_constant : Names.Constant.t -> env -> Declarations.constant_body + val opaque_tables : env -> Opaqueproof.opaquetab + val is_projection : Names.Constant.t -> env -> bool + val lookup_projection : Names.Projection.t -> env -> Declarations.projection_body + val named_context_of_val : named_context_val -> Context.Named.t + val push_named : Context.Named.Declaration.t -> env -> env + val named_context : env -> Context.Named.t + val named_context_val : env -> named_context_val + val push_named_context_val : Context.Named.Declaration.t -> named_context_val -> named_context_val + val reset_with_named_context : named_context_val -> env -> env + val rel_context : env -> Context.Rel.t + val constant_value_in : env -> Names.Constant.t Univ.puniverses -> Term.constr + val named_type : Names.Id.t -> env -> Term.types + val constant_opt_value_in : env -> Names.Constant.t Univ.puniverses -> Term.constr option + val fold_named_context_reverse : + ('a -> Context.Named.Declaration.t -> 'a) -> init:'a -> env -> 'a + val evaluable_named : Names.Id.t -> Environ.env -> bool +end + +module UGraph : +sig + type t = UGraph.t + val pr_universes : (Univ.Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds +end + +module Reduction : +sig + exception NotConvertible + type conv_pb = Prelude.conv_pb = + | CONV + | CUMUL + + val whd_all : Environ.env -> Term.constr -> Term.constr + + val whd_betaiotazeta : Environ.env -> Term.constr -> Term.constr + + val is_arity : Environ.env -> Term.types -> bool + + val dest_prod : Environ.env -> Term.types -> Context.Rel.t * Term.types + + type 'a extended_conversion_function = + ?l2r:bool -> ?reds:Names.transparent_state -> Environ.env -> + ?evars:((Term.existential->Term.constr option) * UGraph.t) -> + 'a -> 'a -> unit + val conv : Term.constr extended_conversion_function +end + +module Vars : +sig + type substl = Term.constr list + + val substl : substl -> Term.constr -> Term.constr + + val subst1 : Term.constr -> Term.constr -> Term.constr + + val lift : int -> Term.constr -> Term.constr + + val closed0 : Term.constr -> bool + + val closedn : int -> Term.constr -> bool + + val replace_vars : (Names.Id.t * Term.constr) list -> Term.constr -> Term.constr + + val noccurn : int -> Term.constr -> bool + val subst_var : Names.Id.t -> Term.constr -> Term.constr + val subst_vars : Names.Id.t list -> Term.constr -> Term.constr + val substnl : substl -> int -> Term.constr -> Term.constr +end + +module Inductive : +sig + type mind_specif = Declarations.mutual_inductive_body * Declarations.one_inductive_body + val type_of_inductive : Environ.env -> mind_specif Univ.puniverses -> Term.types + exception SingletonInductiveBecomesProp of Names.Id.t + val lookup_mind_specif : Environ.env -> Names.inductive -> mind_specif + val find_inductive : Environ.env -> Term.types -> Term.pinductive * Term.constr list +end + +module Typeops : +sig + val type_of_constant_type : Environ.env -> Declarations.constant_type -> Term.types + val type_of_constant_in : Environ.env -> Term.pconstant -> Term.types +end + +module Opaqueproof : +sig + type opaquetab = Opaqueproof.opaquetab + type opaque = Opaqueproof.opaque + val empty_opaquetab : opaquetab + val force_proof : opaquetab -> opaque -> Term.constr +end + +module Modops : +sig + val destr_nofunctor : ('ty,'a) Declarations.functorize -> 'a + val add_structure : + Names.ModPath.t -> Declarations.structure_body -> Mod_subst.delta_resolver -> + Environ.env -> Environ.env + val add_module_type : Names.ModPath.t -> Declarations.module_type_body -> Environ.env -> Environ.env +end + +module Entries : +sig + type mutual_inductive_entry = Entries.mutual_inductive_entry + type inline = int option + type 'a proof_output = Term.constr Univ.in_universe_context_set * 'a + type 'a const_entry_body = 'a proof_output Future.computation + type 'a definition_entry = 'a Entries.definition_entry = + { const_entry_body : 'a const_entry_body; + (* List of section variables *) + const_entry_secctx : Context.Named.t option; + (* State id on which the completion of type checking is reported *) + const_entry_feedback : Stateid.t option; + const_entry_type : Term.types option; + const_entry_polymorphic : bool; + const_entry_universes : Univ.UContext.t; + const_entry_opaque : bool; + const_entry_inline_code : bool } + type parameter_entry = Context.Named.t option * bool * Term.types Univ.in_universe_context * inline + type projection_entry = Entries.projection_entry + type 'a constant_entry = 'a Entries.constant_entry = + | DefinitionEntry of 'a definition_entry + | ParameterEntry of parameter_entry + | ProjectionEntry of projection_entry +end + +module Mod_typing : +sig + type 'alg translation = + Declarations.module_signature * 'alg * Mod_subst.delta_resolver * Univ.ContextSet.t + val translate_mse : + Environ.env -> Names.ModPath.t option -> Entries.inline -> Declarations.module_alg_expr -> + Declarations.module_alg_expr translation +end + +module Esubst : +sig + type 'a subs = 'a Esubst.subs + val subs_id : int -> 'a subs +end + +module CClosure : +sig + type fconstr = CClosure.fconstr + type clos_infos = CClosure.clos_infos + type table_key = Names.Constant.t Univ.puniverses Names.tableKey + type fterm = CClosure.fterm = + | FRel of int + | FAtom of Term.constr (** Metas and Sorts *) + | FCast of fconstr * Term.cast_kind * fconstr + | FFlex of table_key + | FInd of Names.inductive Univ.puniverses + | FConstruct of Names.constructor Univ.puniverses + | FApp of fconstr * fconstr array + | FProj of Names.Projection.t * fconstr + | FFix of Term.fixpoint * fconstr Esubst.subs + | FCoFix of Term.cofixpoint * fconstr Esubst.subs + | FCaseT of Term.case_info * Term.constr * fconstr * Term.constr array * fconstr Esubst.subs (* predicate and branches are closures *) + | FLambda of int * (Names.Name.t * Term.constr) list * Term.constr * fconstr Esubst.subs + | FProd of Names.Name.t * fconstr * fconstr + | FLetIn of Names.Name.t * fconstr * fconstr * Term.constr * fconstr Esubst.subs + | FEvar of Term.existential * fconstr Esubst.subs + | FLIFT of int * fconstr + | FCLOS of Term.constr * fconstr Esubst.subs + | FLOCKED + module RedFlags : sig + type reds = CClosure.RedFlags.reds + type red_kind = CClosure.RedFlags.red_kind + val mkflags : red_kind list -> reds + val fBETA : red_kind + val fCOFIX : red_kind + val fCONST : Names.Constant.t -> CClosure.RedFlags.red_kind + val fFIX : red_kind + val fMATCH : red_kind + val fZETA : red_kind + val red_add_transparent : reds -> Names.transparent_state -> reds + end + val mk_clos : fconstr Esubst.subs -> Term.constr -> fconstr + val mk_atom : Term.constr -> fconstr + val mk_clos_deep : + (fconstr Esubst.subs -> Term.constr -> fconstr) -> + fconstr Esubst.subs -> Term.constr -> fconstr + val mk_red : fterm -> fconstr + val all : RedFlags.reds + val beta : RedFlags.reds + val betaiota : RedFlags.reds + val betaiotazeta : RedFlags.reds + + val create_clos_infos : ?evars:(Term.existential -> Term.constr option) -> RedFlags.reds -> Environ.env -> clos_infos + + val whd_val : clos_infos -> fconstr -> Term.constr + + val inject : Term.constr -> fconstr + + val kl : clos_infos -> fconstr -> Term.constr + val term_of_fconstr : fconstr -> Term.constr +end + +module Type_errors : +sig + type type_error = Type_errors.type_error + exception TypeError of Environ.env * type_error +end + +module Evar : +sig + (** Unique identifier of some {i evar} *) + type t = Prelude.evar + + (** Recover the underlying integer. *) + val repr : t -> int + + val equal : t -> t -> bool + + (** a set of unique identifiers of some {i evars} *) + module Set : module type of struct include Evar.Set end +end + +module Evd : +sig + val string_of_existential : Evar.t -> string + type evar_constraint = Prelude.conv_pb * Environ.env * Term.constr * Term.constr + + (* --------------------------------- *) + + (* evar info *) + + module Store : + sig + type t = Evd.Store.t + val empty : t + end + + module Filter : + sig + type t = Evd.Filter.t + val repr : t -> bool list option + end + + (** This value defines the refinement of a given {i evar} *) + type evar_body = Evd.evar_body = + | Evar_empty (** given {i evar} was not yet refined *) + | Evar_defined of Term.constr (** given {i var} was refined to the indicated term *) + + (** all the information we have concerning some {i evar} *) + type evar_info = Evd.evar_info = + { + evar_concl : Term.constr; + evar_hyps : Environ.named_context_val; + evar_body : evar_body; + evar_filter : Filter.t; + evar_source : Evar_kinds.t Loc.located; + evar_candidates : Term.constr list option; (* if not None, list of allowed instances *) + evar_extra : Store.t + } + + val evar_concl : evar_info -> Term.constr + val evar_body : evar_info -> evar_body + val evar_context : evar_info -> Context.Named.t + val instantiate_evar_array : evar_info -> Term.constr -> Term.constr array -> Term.constr + val evar_filtered_env : evar_info -> Environ.env + val evar_hyps : evar_info -> Environ.named_context_val + + (* ------------------------------------ *) + + (* evar map *) + + type evar_map = Prelude.evar_map + type open_constr = evar_map * Term.constr + + type rigid = Prelude.rigid = + | UnivRigid + | UnivFlexible of bool + + + type 'a freelisted = 'a Evd.freelisted = { + rebus : 'a; + freemetas : Evd.Metaset.t + } + type instance_status = Evd.instance_status + type clbinding = Evd.clbinding = + | Cltyp of Names.Name.t * Term.constr freelisted + | Clval of Names.Name.t * (Term.constr freelisted * instance_status) * Term.constr freelisted + val empty : evar_map + val from_env : Environ.env -> evar_map + val find : evar_map -> Evar.t -> evar_info + val find_undefined : evar_map -> Prelude.evar -> evar_info + val is_defined : evar_map -> Evar.t -> bool + val mem : evar_map -> Evar.t -> bool + val add : evar_map -> Evar.t -> evar_info -> evar_map + val evar_universe_context : evar_map -> UState.t + val set_universe_context : evar_map -> UState.t -> evar_map + val universes : evar_map -> UGraph.t + val define : Evar.t -> Term.constr -> evar_map -> evar_map + val fold : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a + val evar_key : Names.Id.t -> evar_map -> Evar.t + + val create_evar_defs : evar_map -> evar_map + + val meta_declare : Prelude.metavariable -> Term.types -> ?name:Names.Name.t -> evar_map -> evar_map + + val clear_metas : evar_map -> evar_map + + (** Allocates a new evar that represents a {i sort}. *) + val new_sort_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * Sorts.t + + val remove : evar_map -> Evar.t -> evar_map + val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> Environ.env -> + evar_map -> Prelude.global_reference -> evar_map * Term.constr + val evar_filtered_context : evar_info -> Context.Named.t + val fresh_inductive_instance : ?loc:Loc.t -> Environ.env -> evar_map -> Names.inductive -> evar_map * Term.pinductive + val fold_undefined : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a + + val universe_context_set : evar_map -> Univ.ContextSet.t + val evar_ident : Prelude.evar -> evar_map -> Names.Id.t option + val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list + val universe_context : ?names:(Names.Id.t Loc.located) list -> evar_map -> + (Names.Id.t * Univ.Level.t) list * Univ.UContext.t + val nf_constraints : evar_map -> evar_map + val from_ctx : UState.t -> evar_map + + val meta_list : evar_map -> (Prelude.metavariable * clbinding) list + + val meta_defined : evar_map -> Prelude.metavariable -> bool + + val meta_name : evar_map -> Prelude.metavariable -> Names.Name.t + + module MonadR : + sig + module List : + sig + val map_right : ('a -> evar_map -> evar_map * 'b) -> 'a list -> evar_map -> evar_map * 'b list + end + end + + type 'a sigma = 'a Evd.sigma = { + it : 'a ; + sigma : evar_map + } + + val sig_sig : 'a sigma -> evar_map + + val sig_it : 'a sigma -> 'a + + type 'a in_evar_universe_context = 'a * UState.t + + val univ_flexible : rigid + val univ_flexible_alg : rigid + val empty_evar_universe_context : UState.t + val union_evar_universe_context : UState.t -> UState.t -> UState.t + val merge_universe_context : evar_map -> UState.t -> evar_map + + type unsolvability_explanation = Evd.unsolvability_explanation = + | SeveralInstancesFound of int + + module Metaset : module type of struct include Evd.Metaset end + with type elt = Prelude.metavariable + + (** Return {i ids} of all {i evars} that occur in a given term. *) + val evars_of_term : Term.constr -> Evar.Set.t + + val evar_universe_context_of : Univ.ContextSet.t -> UState.t + [@@ocaml.deprecated "alias of API.UState.of_context_set"] + + val evar_context_universe_context : UState.t -> Univ.UContext.t + [@@ocaml.deprecated "alias of API.UState.context"] + + type evar_universe_context = UState.t + [@@ocaml.deprecated "alias of API.UState.t"] + + val existential_opt_value : evar_map -> Term.existential -> Term.constr option + val existential_value : evar_map -> Term.existential -> Term.constr + + exception NotInstantiatedEvar + + val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> Environ.env -> evar_map -> Sorts.family -> evar_map * Sorts.t +end + +module Namegen : +sig + (** *) + + (** [next_ident_away original_id unwanted_ids] returns a new identifier as close as possible + to the [original_id] while avoiding all [unwanted_ids]. + + In particular: + {ul {- if [original_id] does not appear in the list of [unwanted_ids], then [original_id] is returned.} + {- if [original_id] appears in the list of [unwanted_ids], + then this function returns a new id that: + {ul {- has the same {i root} as the [original_id],} + {- does not occur in the list of [unwanted_ids],} + {- has the smallest possible {i subscript}.}}}} + + where by {i subscript} of some identifier we mean last part of it that is composed + only from (decimal) digits and by {i root} of some identifier we mean + the whole identifier except for the {i subscript}. + + E.g. if we take [foo42], then [42] is the {i subscript}, and [foo] is the root. *) + val next_ident_away : Names.Id.t -> Names.Id.t list -> Names.Id.t + + val hdchar : Environ.env -> Evd.evar_map -> EConstr.types -> string + val id_of_name_using_hdchar : Environ.env -> Evd.evar_map -> EConstr.types -> Names.Name.t -> Names.Id.t + val next_ident_away_in_goal : Names.Id.t -> Names.Id.t list -> Names.Id.t + val default_dependent_ident : Names.Id.t + val next_global_ident_away : Names.Id.t -> Names.Id.t list -> Names.Id.t + val rename_bound_vars_as_displayed : + Evd.evar_map -> Names.Id.t list -> Names.Name.t list -> EConstr.types -> EConstr.types +end + +module Safe_typing : +sig + type private_constants = Safe_typing.private_constants + val mk_pure_proof : Term.constr -> Safe_typing.private_constants Entries.proof_output +end + +module Proofview_monad : +sig + type lazy_msg = unit -> Pp.std_ppcmds + module Info : + sig + type tree = Proofview_monad.Info.tree + end +end + +(* All items in the Goal modules are deprecated. *) +module Goal : +sig + type goal = Evar.t + + val pr_goal : goal -> Pp.std_ppcmds + + module V82 : + sig + val new_goal_with : Evd.evar_map -> goal -> Context.Named.t -> goal Evd.sigma + + val nf_hyps : Evd.evar_map -> goal -> Environ.named_context_val + + val env : Evd.evar_map -> goal -> Environ.env + + val concl : Evd.evar_map -> goal -> EConstr.constr + + val mk_goal : Evd.evar_map -> + Environ.named_context_val -> + EConstr.constr -> + Evd.Store.t -> + goal * EConstr.constr * Evd.evar_map + + val extra : Evd.evar_map -> goal -> Evd.Store.t + + val partial_solution_to : Evd.evar_map -> goal -> goal -> EConstr.constr -> Evd.evar_map + + val partial_solution : Evd.evar_map -> goal -> EConstr.constr -> Evd.evar_map + + val hyps : Evd.evar_map -> goal -> Environ.named_context_val + + val abstract_type : Evd.evar_map -> goal -> EConstr.types + end +end + +module Proofview : +sig + type proofview = Proofview.proofview + type entry = Proofview.entry + type +'a tactic = 'a Proofview.tactic + type telescope = Proofview.telescope = + | TNil of Evd.evar_map + | TCons of Environ.env * Evd.evar_map * EConstr.types * (Evd.evar_map -> EConstr.constr -> telescope) + module NonLogical : + sig + type +'a t = 'a Proofview.NonLogical.t + val make : (unit -> 'a) -> 'a t + val return : 'a -> 'a t + val ( >> ) : unit t -> 'a t -> 'a t + val ( >>= ) : 'a t -> ('a -> 'b t) -> 'b t + val print_char : char -> unit t + val print_debug : Pp.std_ppcmds -> unit t + val print_warning : Pp.std_ppcmds -> unit t + val print_notice : Pp.std_ppcmds -> unit t + val print_info : Pp.std_ppcmds -> unit t + val run : 'a t -> 'a + type 'a ref = 'a Proofview.NonLogical.ref + val ref : 'a -> 'a ref t + val ( := ) : 'a ref -> 'a -> unit t + val ( ! ) : 'a ref -> 'a t + val raise : ?info:Exninfo.info -> exn -> 'a t + val catch : 'a t -> (Exninfo.iexn -> 'a t) -> 'a t + val read_line : string t + end + val proofview : proofview -> Goal.goal list * Evd.evar_map + val cycle : int -> unit tactic + val swap : int -> int -> unit tactic + val revgoals : unit tactic + val give_up : unit tactic + val init : Evd.evar_map -> (Environ.env * EConstr.types) list -> entry * proofview + val shelve : unit tactic + val tclZERO : ?info:Exninfo.info -> exn -> 'a tactic + val tclUNIT : 'a -> 'a tactic + val tclBIND : 'a tactic -> ('a -> 'b tactic) -> 'b tactic + val tclORELSE : 'a tactic -> (Util.iexn -> 'a tactic) -> 'a tactic + val tclFOCUS : int -> int -> 'a tactic -> 'a tactic + val tclEVARMAP : Evd.evar_map tactic + val tclTHEN : unit tactic -> 'a tactic -> 'a tactic + val tclLIFT : 'a NonLogical.t -> 'a tactic + val tclOR : 'a tactic -> (Exninfo.iexn -> 'a tactic) -> 'a tactic + val tclIFCATCH : 'a tactic -> ('a -> 'b tactic) -> (Exninfo.iexn -> 'b tactic) -> 'b tactic + val tclINDEPENDENT : unit tactic -> unit tactic + val tclDISPATCH : unit tactic list -> unit tactic + val tclEXTEND : unit tactic list -> unit tactic -> unit tactic list -> unit tactic + val tclBREAK : (Exninfo.iexn -> Exninfo.iexn option) -> 'a tactic -> 'a tactic + val tclENV : Environ.env tactic + val tclONCE : 'a tactic -> 'a tactic + val tclPROGRESS : 'a tactic -> 'a tactic + val shelve_unifiable : unit tactic + val apply : Environ.env -> 'a tactic -> proofview -> 'a + * proofview + * (bool*Goal.goal list*Goal.goal list) + * Proofview_monad.Info.tree + val numgoals : int tactic + val with_shelf : 'a tactic -> (Goal.goal list * 'a) tactic + + module Unsafe : + sig + val tclEVARS : Evd.evar_map -> unit tactic + + val tclGETGOALS : Goal.goal list tactic + + val tclSETGOALS : Goal.goal list -> unit tactic + + val tclNEWGOALS : Goal.goal list -> unit tactic + end + + module Goal : + sig + type 'a t = 'a Proofview.Goal.t + val enter : ([ `LZ ] t -> unit tactic) -> unit tactic + val hyps : 'a t -> EConstr.named_context + val nf_enter : ([ `NF ] t -> unit tactic) -> unit tactic + val enter_one : ([ `LZ ] t -> 'a tactic) -> 'a tactic + val concl : 'a t -> EConstr.constr + val sigma : 'a t -> Evd.evar_map + val goal : [ `NF ] t -> Evar.t + val env : 'a t -> Environ.env + val assume : 'a t -> [ `NF ] t + end + + module Notations : + sig + val (>>=) : 'a tactic -> ('a -> 'b tactic) -> 'b tactic + val (<*>) : unit tactic -> 'a tactic -> 'a tactic + val (<+>) : 'a tactic -> 'a tactic -> 'a tactic + end + module V82 : + sig + type tac = Evar.t Evd.sigma -> Evar.t list Evd.sigma + + val tactic : tac -> unit tactic + + val of_tactic : 'a tactic -> tac + + val nf_evar_goals : unit tactic + + val wrap_exceptions : (unit -> 'a tactic) -> 'a tactic + + val catchable_exception : exn -> bool + end + module Trace : + sig + val name_tactic : Proofview_monad.lazy_msg -> 'a tactic -> 'a tactic + val log : Proofview_monad.lazy_msg -> unit tactic + end +end + +module Ftactic : +sig + type +'a focus = 'a Ftactic.focus + type +'a t = 'a focus Proofview.tactic + val return : 'a -> 'a t + val run : 'a t -> ('a -> unit Proofview.tactic) -> unit Proofview.tactic + val enter : ([ `LZ ] Proofview.Goal.t -> 'a t) -> 'a t + val nf_enter : ([ `NF ] Proofview.Goal.t -> 'a t) -> 'a t + val bind : 'a t -> ('a -> 'b t) -> 'b t + val (>>=) : 'a t -> ('a -> 'b t) -> 'b t + val lift : 'a Proofview.tactic -> 'a t + val with_env : 'a t -> (Environ.env * 'a) t + module List : + sig + val map : ('a -> 'b t) -> 'a list -> 'b list t + val map_right : ('a -> 'b t) -> 'a list -> 'b list t + end + module Notations : + sig + val (>>=) : 'a t -> ('a -> 'b t) -> 'b t + val (<*>) : unit t -> 'a t -> 'a t + end +end + +module Evarutil : +sig + val e_new_global : Evd.evar_map ref -> Globnames.global_reference -> EConstr.constr + + val nf_evars_and_universes : Evd.evar_map -> Evd.evar_map * (Term.constr -> Term.constr) + val nf_evar : Evd.evar_map -> EConstr.constr -> EConstr.constr + val nf_evar_info : Evd.evar_map -> Evd.evar_info -> Evd.evar_info + + val mk_new_meta : unit -> EConstr.constr + + (** [new_meta] is a generator of unique meta variables *) + val new_meta : unit -> Prelude.metavariable + + val new_Type : ?rigid:Evd.rigid -> Environ.env -> Evd.evar_map -> Evd.evar_map * EConstr.constr + val new_global : Evd.evar_map -> Prelude.global_reference -> Evd.evar_map * EConstr.constr + + val new_evar : + Environ.env -> Evd.evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t -> + ?candidates:EConstr.constr list -> ?store:Evd.Store.t -> + ?naming:Misctypes.intro_pattern_naming_expr -> + ?principal:bool -> EConstr.types -> Evd.evar_map * EConstr.constr + + val new_evar_instance : + Environ.named_context_val -> Evd.evar_map -> EConstr.types -> + ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t -> ?candidates:EConstr.constr list -> + ?store:Evd.Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> + ?principal:bool -> + EConstr.constr list -> Evd.evar_map * EConstr.constr + + val clear_hyps_in_evi : Environ.env -> Evd.evar_map ref -> Environ.named_context_val -> + EConstr.types -> Names.Id.Set.t -> Environ.named_context_val * EConstr.types + + exception ClearDependencyError of Names.Id.t * Evarutil.clear_dependency_error + val undefined_evars_of_term : Evd.evar_map -> EConstr.constr -> Evar.Set.t + val e_new_evar : + Environ.env -> Evd.evar_map ref -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t -> + ?candidates:EConstr.constr list -> ?store:Evd.Store.t -> + ?naming:Misctypes.intro_pattern_naming_expr -> + ?principal:bool -> EConstr.types -> EConstr.constr + val new_type_evar : + Environ.env -> Evd.evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Evd.Filter.t -> + ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> Evd.rigid -> + Evd.evar_map * (EConstr.constr * Sorts.t) + val nf_evars_universes : Evd.evar_map -> Term.constr -> Term.constr + val safe_evar_value : Evd.evar_map -> Term.existential -> Term.constr option + val evd_comb1 : (Evd.evar_map -> 'b -> Evd.evar_map * 'a) -> Evd.evar_map ref -> 'b -> 'a +end + +module Geninterp : +sig + module Val : + sig + type 'a typ = 'a Geninterp.Val.typ + type t = Geninterp.Val.t = Dyn : 'a typ * 'a -> t + type 'a tag = 'a Geninterp.Val.tag = + | Base : 'a typ -> 'a tag + | List : 'a tag -> 'a list tag + | Opt : 'a tag -> 'a option tag + | Pair : 'a tag * 'b tag -> ('a * 'b) tag + val create : string -> 'a typ + val pr : 'a typ -> Pp.std_ppcmds + val eq : 'a typ -> 'b typ -> ('a, 'b) CSig.eq option + val typ_list : t list typ + val typ_opt : t option typ + val typ_pair : (t * t) typ + val repr : 'a typ -> string + val inject : 'a tag -> 'a -> t + end + module TacStore : + sig + type t = Geninterp.TacStore.t + type 'a field = 'a Geninterp.TacStore.field + val empty : t + val field : unit -> 'a field + val get : t -> 'a field -> 'a option + val set : t -> 'a field -> 'a -> t + val remove : t -> 'a field -> t + val merge : t -> t -> t + end + type interp_sign = Geninterp.interp_sign = + {lfun : Val.t Names.Id.Map.t; + extra : TacStore.t } + type ('glb, 'top) interp_fun = interp_sign -> 'glb -> 'top Ftactic.t + val register_interp0 : + ('raw, 'glb, 'top) Genarg.genarg_type -> ('glb, Val.t) interp_fun -> unit + val register_val0 : ('raw, 'glb, 'top) Genarg.genarg_type -> 'top Val.tag option -> unit + val val_tag : 'a Genarg.typed_abstract_argument_type -> 'a Val.tag + val interp : ('raw, 'glb, 'top) Genarg.genarg_type -> ('glb, Val.t) interp_fun +end + +module Globnames : +sig + type global_reference = Globnames.global_reference = + | VarRef of Names.Id.t + | ConstRef of Names.Constant.t + | IndRef of Names.inductive + | ConstructRef of Names.constructor + + type extended_global_reference = Globnames.extended_global_reference = + | TrueGlobal of global_reference + | SynDef of Names.KerName.t + + (* Long term: change implementation so that only 1 kind of order is needed. + * Today: _env ones are fine grained, which one to pick depends. Eg. + * - conversion rule are implemented by the non_env ones + * - pretty printing (of user provided names/aliases) are implemented by + * the _env ones + *) + module Refset : module type of struct include Globnames.Refset end + module Refmap : module type of struct include Globnames.Refmap end + module Refset_env : module type of struct include Globnames.Refset_env end + module Refmap_env : module type of struct include Globnames.Refmap_env end + module RefOrdered : + sig + type t = global_reference + val compare : t -> t -> int + end + + val pop_global_reference : global_reference -> global_reference + val eq_gr : global_reference -> global_reference -> bool + val destIndRef : global_reference -> Names.inductive + + val encode_mind : Names.DirPath.t -> Names.Id.t -> Names.MutInd.t + val encode_con : Names.DirPath.t -> Names.Id.t -> Names.Constant.t + + val global_of_constr : Term.constr -> global_reference + + val subst_global : Mod_subst.substitution -> global_reference -> global_reference * Term.constr + val destConstructRef : Globnames.global_reference -> Names.constructor + + val reference_of_constr : Term.constr -> global_reference + [@@ocaml.deprecated "alias of API.Globnames.global_of_constr"] + + val is_global : global_reference -> Term.constr -> bool +end + +module Evar_kinds : +sig + type obligation_definition_status = Evar_kinds.obligation_definition_status = + | Define of bool + | Expand + + type matching_var_kind = Evar_kinds.matching_var_kind = + | FirstOrderPatVar of Names.Id.t + | SecondOrderPatVar of Names.Id.t + + type t = Evar_kinds.t = + | ImplicitArg of Globnames.global_reference * (int * Names.Id.t option) + * bool (** Force inference *) + | BinderType of Names.Name.t + | NamedHole of Names.Id.t (* coming from some ?[id] syntax *) + | QuestionMark of obligation_definition_status * Names.Name.t + | CasesType of bool (* true = a subterm of the type *) + | InternalHole + | TomatchTypeParameter of Names.inductive * int + | GoalEvar + | ImpossibleCase + | MatchingVar of matching_var_kind + | VarInstance of Names.Id.t + | SubEvar of Prelude.evar +end + +module Decl_kinds : +sig + type polymorphic = bool + type recursivity_kind = Decl_kinds.recursivity_kind = + | Finite + | CoFinite + | BiFinite + type locality = Decl_kinds.locality = + | Discharge + | Local + | Global + type definition_object_kind = Decl_kinds.definition_object_kind = + | Definition + | Coercion + | SubClass + | CanonicalStructure + | Example + | Fixpoint + | CoFixpoint + | Scheme + | StructureComponent + | IdentityCoercion + | Instance + | Method + type theorem_kind = Decl_kinds.theorem_kind = + | Theorem + | Lemma + | Fact + | Remark + | Property + | Proposition + | Corollary + type goal_object_kind = Decl_kinds.goal_object_kind = + | DefinitionBody of definition_object_kind + | Proof of theorem_kind + type goal_kind = locality * polymorphic * goal_object_kind + type assumption_object_kind = Decl_kinds.assumption_object_kind = + | Definitional + | Logical + | Conjectural + type logical_kind = Decl_kinds.logical_kind = + | IsAssumption of assumption_object_kind + | IsDefinition of definition_object_kind + | IsProof of theorem_kind + type binding_kind = Decl_kinds.binding_kind = + | Explicit + | Implicit + type private_flag = bool + type definition_kind = locality * polymorphic * definition_object_kind +end + +module Misctypes : +sig + type evars_flag = bool + type clear_flag = bool option + type advanced_flag = bool + type rec_flag = bool + + type 'a or_by_notation = 'a Misctypes.or_by_notation = + | AN of 'a + | ByNotation of (string * string option) Loc.located + type 'a or_var = 'a Misctypes.or_var = + | ArgArg of 'a + | ArgVar of Names.Id.t Loc.located + type 'a and_short_name = 'a * Names.Id.t Loc.located option + type glob_level = Misctypes.glob_level + type 'a glob_sort_gen = 'a Misctypes.glob_sort_gen = + | GProp + | GSet + | GType of 'a + type sort_info = Names.Name.t Loc.located list + type glob_sort = sort_info glob_sort_gen + type 'a cast_type = 'a Misctypes.cast_type = + | CastConv of 'a + | CastVM of 'a + | CastCoerce + | CastNative of 'a + type 'constr intro_pattern_expr = 'constr Misctypes.intro_pattern_expr = + | IntroForthcoming of bool + | IntroNaming of intro_pattern_naming_expr + | IntroAction of 'constr intro_pattern_action_expr + and intro_pattern_naming_expr = Misctypes.intro_pattern_naming_expr = + | IntroIdentifier of Names.Id.t + | IntroFresh of Names.Id.t + | IntroAnonymous + and 'constr intro_pattern_action_expr = 'constr Misctypes.intro_pattern_action_expr = + | IntroWildcard + | IntroOrAndPattern of 'constr or_and_intro_pattern_expr + | IntroInjection of ('constr intro_pattern_expr) Loc.located list + | IntroApplyOn of 'constr Loc.located * 'constr intro_pattern_expr Loc.located + | IntroRewrite of bool + and 'constr or_and_intro_pattern_expr = 'constr Misctypes.or_and_intro_pattern_expr = + | IntroOrPattern of ('constr intro_pattern_expr) Loc.located list list + | IntroAndPattern of ('constr intro_pattern_expr) Loc.located list + type quantified_hypothesis = Misctypes.quantified_hypothesis = + | AnonHyp of int + | NamedHyp of Names.Id.t + type 'a explicit_bindings = (quantified_hypothesis * 'a) Loc.located list + type 'a bindings = 'a Misctypes.bindings = + | ImplicitBindings of 'a list + | ExplicitBindings of 'a explicit_bindings + | NoBindings + type 'a with_bindings = 'a * 'a bindings + type 'a core_destruction_arg = 'a Misctypes.core_destruction_arg = + | ElimOnConstr of 'a + | ElimOnIdent of Names.Id.t Loc.located + | ElimOnAnonHyp of int + type inversion_kind = Misctypes.inversion_kind = + | SimpleInversion + | FullInversion + | FullInversionClear + type multi = Misctypes.multi = + | Precisely of int + | UpTo of int + | RepeatStar + | RepeatPlus + type 'id move_location = 'id Misctypes.move_location = + | MoveAfter of 'id + | MoveBefore of 'id + | MoveFirst + | MoveLast + type 'a destruction_arg = clear_flag * 'a core_destruction_arg +end + +module Pattern : +sig + type case_info_pattern = Pattern.case_info_pattern + type constr_pattern = Pattern.constr_pattern = + | PRef of Globnames.global_reference + | PVar of Names.Id.t + | PEvar of Evar.t * constr_pattern array + | PRel of int + | PApp of constr_pattern * constr_pattern array + | PSoApp of Names.Id.t * constr_pattern list + | PProj of Names.Projection.t * constr_pattern + | PLambda of Names.Name.t * constr_pattern * constr_pattern + | PProd of Names.Name.t * constr_pattern * constr_pattern + | PLetIn of Names.Name.t * constr_pattern * constr_pattern option * constr_pattern + | PSort of Misctypes.glob_sort + | PMeta of Names.Id.t option + | PIf of constr_pattern * constr_pattern * constr_pattern + | PCase of case_info_pattern * constr_pattern * constr_pattern * + (int * bool list * constr_pattern) list (** index of constructor, nb of args *) + | PFix of Term.fixpoint + | PCoFix of Term.cofixpoint + type constr_under_binders = Names.Id.t list * EConstr.constr + type extended_patvar_map = constr_under_binders Names.Id.Map.t + type patvar_map = EConstr.constr Names.Id.Map.t +end + +module Constrexpr : +sig + type binder_kind = Constrexpr.binder_kind = + | Default of Decl_kinds.binding_kind + | Generalized of Decl_kinds.binding_kind * Decl_kinds.binding_kind * bool + type explicitation = Constrexpr.explicitation = + | ExplByPos of int * Names.Id.t option + | ExplByName of Names.Id.t + type sign = bool + type raw_natural_number = string + type prim_token = Constrexpr.prim_token = + | Numeral of raw_natural_number * sign + | String of string + type notation = string + type instance_expr = Misctypes.glob_level list + type proj_flag = int option + type abstraction_kind = Constrexpr.abstraction_kind = + | AbsLambda + | AbsPi + type cases_pattern_expr_r = Constrexpr.cases_pattern_expr_r = + | CPatAlias of cases_pattern_expr * Names.Id.t + | CPatCstr of Prelude.reference + * cases_pattern_expr list option * cases_pattern_expr list + (** [CPatCstr (_, c, Some l1, l2)] represents (@c l1) l2 *) + | CPatAtom of Prelude.reference option + | CPatOr of cases_pattern_expr list + | CPatNotation of notation * cases_pattern_notation_substitution + * cases_pattern_expr list + | CPatPrim of prim_token + | CPatRecord of (Prelude.reference * cases_pattern_expr) list + | CPatDelimiters of string * cases_pattern_expr + | CPatCast of cases_pattern_expr * constr_expr + and cases_pattern_expr = cases_pattern_expr_r CAst.t + + and cases_pattern_notation_substitution = + cases_pattern_expr list * cases_pattern_expr list list + + and constr_expr_r = Constrexpr.constr_expr_r = + | CRef of Prelude.reference * instance_expr option + | CFix of Names.Id.t Loc.located * fix_expr list + | CCoFix of Names.Id.t Loc.located * cofix_expr list + | CProdN of binder_expr list * constr_expr + | CLambdaN of binder_expr list * constr_expr + | CLetIn of Names.Name.t Loc.located * constr_expr * constr_expr option * constr_expr + | CAppExpl of (proj_flag * Prelude.reference * instance_expr option) * constr_expr list + | CApp of (proj_flag * constr_expr) * + (constr_expr * explicitation Loc.located option) list + | CRecord of (Prelude.reference * constr_expr) list + | CCases of Term.case_style + * constr_expr option + * case_expr list + * branch_expr list + | CLetTuple of Names.Name.t Loc.located list * (Names.Name.t Loc.located option * constr_expr option) * + constr_expr * constr_expr + | CIf of constr_expr * (Names.Name.t Loc.located option * constr_expr option) + * constr_expr * constr_expr + | CHole of Evar_kinds.t option * Misctypes.intro_pattern_naming_expr * Genarg.raw_generic_argument option + | CPatVar of Names.Id.t + | CEvar of Glob_term.existential_name * (Names.Id.t * constr_expr) list + | CSort of Misctypes.glob_sort + | CCast of constr_expr * constr_expr Misctypes.cast_type + | CNotation of notation * constr_notation_substitution + | CGeneralization of Decl_kinds.binding_kind * abstraction_kind option * constr_expr + | CPrim of prim_token + | CDelimiters of string * constr_expr + and constr_expr = constr_expr_r CAst.t + + and case_expr = constr_expr * Names.Name.t Loc.located option * cases_pattern_expr option + + and branch_expr = + (cases_pattern_expr list Loc.located list * constr_expr) Loc.located + + and binder_expr = + Names.Name.t Loc.located list * binder_kind * constr_expr + + and fix_expr = + Names.Id.t Loc.located * (Names.Id.t Loc.located option * recursion_order_expr) * + local_binder_expr list * constr_expr * constr_expr + + and cofix_expr = + Names.Id.t Loc.located * local_binder_expr list * constr_expr * constr_expr + + and recursion_order_expr = Constrexpr.recursion_order_expr = + | CStructRec + | CWfRec of constr_expr + | CMeasureRec of constr_expr * constr_expr option + + and local_binder_expr = Constrexpr.local_binder_expr = + | CLocalAssum of Names.Name.t Loc.located list * binder_kind * constr_expr + | CLocalDef of Names.Name.t Loc.located * constr_expr * constr_expr option + | CLocalPattern of (cases_pattern_expr * constr_expr option) Loc.located + + and constr_notation_substitution = + constr_expr list * + constr_expr list list * + local_binder_expr list list + + type typeclass_constraint = (Names.Name.t Loc.located * Names.Id.t Loc.located list option) * Decl_kinds.binding_kind * constr_expr + type constr_pattern_expr = constr_expr +end + +module Goptions : +sig + type option_name = string list + type 'a option_sig = 'a Goptions.option_sig = + { + optdepr : bool; + optname : string; + optkey : option_name; + optread : unit -> 'a; + optwrite : 'a -> unit + } + type 'a write_function = 'a Goptions.write_function + val declare_bool_option : ?preprocess:(bool -> bool) -> + bool option_sig -> bool write_function + val declare_int_option : ?preprocess:(int option -> int option) -> + int option option_sig -> int option write_function + val declare_string_option: ?preprocess:(string -> string) -> + string option_sig -> string write_function + val set_bool_option_value : option_name -> bool -> unit +end + +module Locus : +sig + type 'a occurrences_gen = 'a Locus.occurrences_gen = + | AllOccurrences + | AllOccurrencesBut of 'a list (** non-empty *) + | NoOccurrences + | OnlyOccurrences of 'a list (** non-empty *) + type occurrences = int occurrences_gen + type occurrences_expr = (int Misctypes.or_var) occurrences_gen + type 'a with_occurrences = occurrences_expr * 'a + type hyp_location_flag = Locus.hyp_location_flag = + InHyp | InHypTypeOnly | InHypValueOnly + type 'a hyp_location_expr = 'a with_occurrences * hyp_location_flag + type 'id clause_expr = 'id Locus.clause_expr = + { onhyps : 'id hyp_location_expr list option; + concl_occs : occurrences_expr } + type clause = Names.Id.t clause_expr + type hyp_location = Names.Id.t * hyp_location_flag + type goal_location = hyp_location option +end + +module Genredexpr : +sig + + (** The parsing produces initially a list of [red_atom] *) + + type 'a red_atom = 'a Genredexpr.red_atom = + | FBeta + | FMatch + | FFix + | FCofix + | FZeta + | FConst of 'a list + | FDeltaBut of 'a list + + (** This list of atoms is immediately converted to a [glob_red_flag] *) + + type 'a glob_red_flag = 'a Genredexpr.glob_red_flag = { + rBeta : bool; + rMatch : bool; + rFix : bool; + rCofix : bool; + rZeta : bool; + rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*) + rConst : 'a list + } + + (** Generic kinds of reductions *) + + type ('a,'b,'c) red_expr_gen = ('a,'b,'c) Genredexpr.red_expr_gen = + | Red of bool + | Hnf + | Simpl of 'b glob_red_flag*('b,'c) Util.union Locus.with_occurrences option + | Cbv of 'b glob_red_flag + | Cbn of 'b glob_red_flag + | Lazy of 'b glob_red_flag + | Unfold of 'b Locus.with_occurrences list + | Fold of 'a list + | Pattern of 'a Locus.with_occurrences list + | ExtraRedExpr of string + | CbvVm of ('b,'c) Util.union Locus.with_occurrences option + | CbvNative of ('b,'c) Util.union Locus.with_occurrences option + + type ('a,'b,'c) may_eval = ('a,'b,'c) Genredexpr.may_eval = + | ConstrTerm of 'a + | ConstrEval of ('a,'b,'c) red_expr_gen * 'a + | ConstrContext of Names.Id.t Loc.located * 'a + | ConstrTypeOf of 'a + + type r_trm = Constrexpr.constr_expr + type r_pat = Constrexpr.constr_pattern_expr + type r_cst = Prelude.reference Misctypes.or_by_notation + type raw_red_expr = (r_trm, r_cst, r_pat) red_expr_gen +end + +module Vernacexpr : +sig + type instance_flag = bool option + type coercion_flag = bool + type inductive_flag = Decl_kinds.recursivity_kind + type lname = Names.Name.t Loc.located + type lident = Names.Id.t Loc.located + type opacity_flag = Vernacexpr.opacity_flag = + | Opaque of lident list option + | Transparent + type locality_flag = bool + type inductive_kind = Vernacexpr.inductive_kind = + | Inductive_kw | CoInductive | Variant | Record | Structure | Class of bool + type 'a hint_info_gen = 'a Vernacexpr.hint_info_gen = + { hint_priority : int option; + hint_pattern : 'a option } + type vernac_type = Vernacexpr.vernac_type = + | VtStartProof of vernac_start + | VtSideff of vernac_sideff_type + | VtQed of vernac_qed_type + | VtProofStep of proof_step + | VtProofMode of string + | VtQuery of vernac_part_of_script * report_with + | VtStm of vernac_control * vernac_part_of_script + | VtUnknown + and report_with = Stateid.t * Feedback.route_id + and vernac_qed_type = Vernacexpr.vernac_qed_type = + | VtKeep + | VtKeepAsAxiom + | VtDrop + and vernac_start = string * opacity_guarantee * Names.Id.t list + and vernac_sideff_type = Names.Id.t list + and vernac_part_of_script = bool + and vernac_control = Vernacexpr.vernac_control = + | VtWait + | VtJoinDocument + | VtBack of Stateid.t + and opacity_guarantee = Vernacexpr.opacity_guarantee = + | GuaranteesOpacity + | Doesn'tGuaranteeOpacity + and proof_step = Vernacexpr.proof_step = { + parallel : [ `Yes of solving_tac * anon_abstracting_tac | `No ]; + proof_block_detection : proof_block_name option + } + and solving_tac = bool + and anon_abstracting_tac = bool + and proof_block_name = string + type vernac_when = Vernacexpr.vernac_when = + | VtNow + | VtLater + type verbose_flag = bool + + type obsolete_locality = bool + + type lstring = Vernacexpr.lstring + type 'a with_coercion = coercion_flag * 'a + type scope_name = string + type decl_notation = lstring * Constrexpr.constr_expr * scope_name option + type constructor_expr = (lident * Constrexpr.constr_expr) with_coercion + type 'a with_notation = 'a * decl_notation list + type local_decl_expr = Vernacexpr.local_decl_expr = + | AssumExpr of lname * Constrexpr.constr_expr + | DefExpr of lname * Constrexpr.constr_expr * Constrexpr.constr_expr option + type 'a with_priority = 'a * int option + type 'a with_instance = instance_flag * 'a + type constructor_list_or_record_decl_expr = Vernacexpr.constructor_list_or_record_decl_expr = + | Constructors of constructor_expr list + | RecordDecl of lident option * local_decl_expr with_instance with_priority with_notation list + type plident = lident * lident list option + type inductive_expr = plident with_coercion * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * inductive_kind * constructor_list_or_record_decl_expr + + type syntax_modifier = Vernacexpr.syntax_modifier + type class_rawexpr = Vernacexpr.class_rawexpr + type definition_expr = Vernacexpr.definition_expr + type hint_info_expr = Constrexpr.constr_pattern_expr hint_info_gen + type proof_expr = Vernacexpr.proof_expr + type proof_end = Vernacexpr.proof_end = + | Admitted + | Proved of opacity_flag * lident option + type inline = Vernacexpr.inline + type fixpoint_expr = plident * (Names.Id.t Loc.located option * Constrexpr.recursion_order_expr) * Constrexpr.local_binder_expr list * Constrexpr.constr_expr * Constrexpr.constr_expr option + type cofixpoint_expr = Vernacexpr.cofixpoint_expr + type scheme = Vernacexpr.scheme + type section_subset_expr = Vernacexpr.section_subset_expr + type module_binder = Vernacexpr.module_binder + type vernac_argument_status = Vernacexpr.vernac_argument_status + type vernac_implicit_status = Vernacexpr.vernac_implicit_status + type module_ast_inl = Vernacexpr.module_ast_inl + type 'a module_signature = 'a Vernacexpr.module_signature + type extend_name = string * int + type simple_binder = Vernacexpr.simple_binder + type option_value = Vernacexpr.option_value + type showable = Vernacexpr.showable + type bullet = Vernacexpr.bullet + type stm_vernac = Vernacexpr.stm_vernac + type comment = Vernacexpr.comment + type register_kind = Vernacexpr.register_kind + type locatable = Vernacexpr.locatable + type search_restriction = Vernacexpr.search_restriction + type searchable = Vernacexpr.searchable + type printable = Vernacexpr.printable + type option_ref_value = Vernacexpr.option_ref_value + type onlyparsing_flag = Vernacexpr.onlyparsing_flag + type reference_or_constr = Vernacexpr.reference_or_constr + type hint_mode = Vernacexpr.hint_mode + type hints_expr = Vernacexpr.hints_expr = + | HintsResolve of (hint_info_expr * bool * reference_or_constr) list + | HintsImmediate of reference_or_constr list + | HintsUnfold of Prelude.reference list + | HintsTransparency of Prelude.reference list * bool + | HintsMode of Prelude.reference * hint_mode list + | HintsConstructors of Prelude.reference list + | HintsExtern of int * Constrexpr.constr_expr option * Genarg.raw_generic_argument + type vernac_expr = Vernacexpr.vernac_expr = + | VernacLoad of verbose_flag * string + | VernacTime of vernac_expr Loc.located + | VernacRedirect of string * vernac_expr Loc.located + | VernacTimeout of int * vernac_expr + | VernacFail of vernac_expr + | VernacSyntaxExtension of + obsolete_locality * (lstring * syntax_modifier list) + | VernacOpenCloseScope of obsolete_locality * (bool * scope_name) + | VernacDelimiters of scope_name * string option + | VernacBindScope of scope_name * class_rawexpr list + | VernacInfix of obsolete_locality * (lstring * syntax_modifier list) * + Constrexpr.constr_expr * scope_name option + | VernacNotation of + obsolete_locality * Constrexpr.constr_expr * (lstring * syntax_modifier list) * + scope_name option + | VernacNotationAddFormat of string * string * string + | VernacDefinition of + (Decl_kinds.locality option * Decl_kinds.definition_object_kind) * plident * definition_expr + | VernacStartTheoremProof of Decl_kinds.theorem_kind * proof_expr list * bool + | VernacEndProof of proof_end + | VernacExactProof of Constrexpr.constr_expr + | VernacAssumption of (Decl_kinds.locality option * Decl_kinds.assumption_object_kind) * + inline * (plident list * Constrexpr.constr_expr) with_coercion list + | VernacInductive of Decl_kinds.private_flag * inductive_flag * (inductive_expr * decl_notation list) list + | VernacFixpoint of + Decl_kinds.locality option * (fixpoint_expr * decl_notation list) list + | VernacCoFixpoint of + Decl_kinds.locality option * (cofixpoint_expr * decl_notation list) list + | VernacScheme of (lident option * scheme) list + | VernacCombinedScheme of lident * lident list + | VernacUniverse of lident list + | VernacConstraint of (Misctypes.glob_level * Univ.constraint_type * Misctypes.glob_level) list + | VernacBeginSection of lident + | VernacEndSegment of lident + | VernacRequire of + Prelude.reference option * bool option * Prelude.reference list + | VernacImport of bool * Prelude.reference list + | VernacCanonical of Prelude.reference Misctypes.or_by_notation + | VernacCoercion of obsolete_locality * Prelude.reference Misctypes.or_by_notation * + class_rawexpr * class_rawexpr + | VernacIdentityCoercion of obsolete_locality * lident * + class_rawexpr * class_rawexpr + | VernacNameSectionHypSet of lident * section_subset_expr + | VernacInstance of + bool * + Constrexpr.local_binder_expr list * + Constrexpr.typeclass_constraint * + (bool * Constrexpr.constr_expr) option * + hint_info_expr + | VernacContext of Constrexpr.local_binder_expr list + | VernacDeclareInstances of + (Prelude.reference * hint_info_expr) list + | VernacDeclareClass of Prelude.reference + | VernacDeclareModule of bool option * lident * + module_binder list * module_ast_inl + | VernacDefineModule of bool option * lident * module_binder list * + module_ast_inl module_signature * module_ast_inl list + | VernacDeclareModuleType of lident * + module_binder list * module_ast_inl list * module_ast_inl list + | VernacInclude of module_ast_inl list + | VernacSolveExistential of int * Constrexpr.constr_expr + | VernacAddLoadPath of bool * string * Names.DirPath.t option + | VernacRemoveLoadPath of string + | VernacAddMLPath of bool * string + | VernacDeclareMLModule of string list + | VernacChdir of string option + | VernacWriteState of string + | VernacRestoreState of string + | VernacResetName of lident + | VernacResetInitial + | VernacBack of int + | VernacBackTo of int + | VernacCreateHintDb of string * bool + | VernacRemoveHints of string list * Prelude.reference list + | VernacHints of obsolete_locality * string list * hints_expr + | VernacSyntacticDefinition of Names.Id.t Loc.located * (Names.Id.t list * Constrexpr.constr_expr) * + obsolete_locality * onlyparsing_flag + | VernacDeclareImplicits of Prelude.reference Misctypes.or_by_notation * + (Constrexpr.explicitation * bool * bool) list list + | VernacArguments of Prelude.reference Misctypes.or_by_notation * + vernac_argument_status list * + (Names.Name.t * vernac_implicit_status) list list * + int option * + [ `ReductionDontExposeCase | `ReductionNeverUnfold | `Rename | + `ExtraScopes | `Assert | `ClearImplicits | `ClearScopes | + `DefaultImplicits ] list + | VernacArgumentsScope of Prelude.reference Misctypes.or_by_notation * + scope_name option list + | VernacReserve of simple_binder list + | VernacGeneralizable of (lident list) option + | VernacSetOpacity of (Conv_oracle.level * Prelude.reference Misctypes.or_by_notation list) + | VernacSetStrategy of + (Conv_oracle.level * Prelude.reference Misctypes.or_by_notation list) list + | VernacUnsetOption of Goptions.option_name + | VernacSetOption of Goptions.option_name * option_value + | VernacSetAppendOption of Goptions.option_name * string + | VernacAddOption of Goptions.option_name * option_ref_value list + | VernacRemoveOption of Goptions.option_name * option_ref_value list + | VernacMemOption of Goptions.option_name * option_ref_value list + | VernacPrintOption of Goptions.option_name + | VernacCheckMayEval of Genredexpr.raw_red_expr option * goal_selector option * Constrexpr.constr_expr + | VernacGlobalCheck of Constrexpr.constr_expr + | VernacDeclareReduction of string * Genredexpr.raw_red_expr + | VernacPrint of printable + | VernacSearch of searchable * goal_selector option * search_restriction + | VernacLocate of locatable + | VernacRegister of lident * register_kind + | VernacComments of comment list + | VernacStm of stm_vernac + | VernacGoal of Constrexpr.constr_expr + | VernacAbort of lident option + | VernacAbortAll + | VernacRestart + | VernacUndo of int + | VernacUndoTo of int + | VernacBacktrack of int*int*int + | VernacFocus of int option + | VernacUnfocus + | VernacUnfocused + | VernacBullet of bullet + | VernacSubproof of int option + | VernacEndSubproof + | VernacShow of showable + | VernacCheckGuard + | VernacProof of Genarg.raw_generic_argument option * section_subset_expr option + | VernacProofMode of string + | VernacToplevelControl of exn + | VernacExtend of extend_name * Genarg.raw_generic_argument list + | VernacProgram of vernac_expr + | VernacPolymorphic of bool * vernac_expr + | VernacLocal of bool * vernac_expr + and goal_selector = Vernacexpr.goal_selector = + | SelectNth of int + | SelectList of (int * int) list + | SelectId of Names.Id.t + | SelectAll + and vernac_classification = vernac_type * vernac_when + and one_inductive_expr = + plident * Constrexpr.local_binder_expr list * Constrexpr.constr_expr option * constructor_expr list +end + +module Glob_term : +sig + type cases_pattern_r = Glob_term.cases_pattern_r = + | PatVar of Names.Name.t + | PatCstr of Names.constructor * cases_pattern list * Names.Name.t + and cases_pattern = cases_pattern_r CAst.t + type existential_name = Names.Id.t + type glob_constr_r = Glob_term.glob_constr_r = + | GRef of Globnames.global_reference * Misctypes.glob_level list option + (** An identifier that represents a reference to an object defined + either in the (global) environment or in the (local) context. *) + | GVar of Names.Id.t + (** An identifier that cannot be regarded as "GRef". + Bound variables are typically represented this way. *) + | GEvar of existential_name * (Names.Id.t * glob_constr) list + | GPatVar of Evar_kinds.matching_var_kind + | GApp of glob_constr * glob_constr list + | GLambda of Names.Name.t * Decl_kinds.binding_kind * glob_constr * glob_constr + | GProd of Names.Name.t * Decl_kinds.binding_kind * glob_constr * glob_constr + | GLetIn of Names.Name.t * glob_constr * glob_constr option * glob_constr + | GCases of Term.case_style * glob_constr option * tomatch_tuples * cases_clauses + | GLetTuple of Names.Name.t list * (Names.Name.t * glob_constr option) * glob_constr * glob_constr + | GIf of glob_constr * (Names.Name.t * glob_constr option) * glob_constr * glob_constr + | GRec of fix_kind * Names.Id.t array * glob_decl list array * + glob_constr array * glob_constr array + | GSort of Misctypes.glob_sort + | GHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option + | GCast of glob_constr * glob_constr Misctypes.cast_type + + and glob_constr = glob_constr_r CAst.t + + and glob_decl = Names.Name.t * Decl_kinds.binding_kind * glob_constr option * glob_constr + + and fix_recursion_order = Glob_term.fix_recursion_order = + | GStructRec + | GWfRec of glob_constr + | GMeasureRec of glob_constr * glob_constr option + + and fix_kind = Glob_term.fix_kind = + | GFix of ((int option * fix_recursion_order) array * int) + | GCoFix of int + + and predicate_pattern = + Names.Name.t * (Names.inductive * Names.Name.t list) Loc.located option + + and tomatch_tuple = (glob_constr * predicate_pattern) + + and tomatch_tuples = tomatch_tuple list + + and cases_clause = (Names.Id.t list * cases_pattern list * glob_constr) Loc.located + and cases_clauses = cases_clause list + + type closure = Glob_term.closure = + { idents:Names.Id.t Names.Id.Map.t; + typed: Pattern.constr_under_binders Names.Id.Map.t ; + untyped:closed_glob_constr Names.Id.Map.t } + and closed_glob_constr = Glob_term.closed_glob_constr = { + closure: closure; + term: glob_constr } + + type var_map = Pattern.constr_under_binders Names.Id.Map.t + type uconstr_var_map = Glob_term.closed_glob_constr Names.Id.Map.t + type unbound_ltac_var_map = Geninterp.Val.t Names.Id.Map.t + type ltac_var_map = Glob_term.ltac_var_map = { + ltac_constrs : var_map; + (** Ltac variables bound to constrs *) + ltac_uconstrs : uconstr_var_map; + (** Ltac variables bound to untyped constrs *) + ltac_idents: Names.Id.t Names.Id.Map.t; + (** Ltac variables bound to identifiers *) + ltac_genargs : unbound_ltac_var_map; + (** Ltac variables bound to other kinds of arguments *) + } +end + +module Libnames : +sig + type full_path = Libnames.full_path + val pr_path : Libnames.full_path -> Pp.std_ppcmds + val make_path : Names.DirPath.t -> Names.Id.t -> full_path + val eq_full_path : full_path -> full_path -> bool + val dirpath : full_path -> Names.DirPath.t + val path_of_string : string -> full_path + + type qualid = Libnames.qualid + val make_qualid : Names.DirPath.t -> Names.Id.t -> qualid + val qualid_eq : qualid -> qualid -> bool + val repr_qualid : qualid -> Names.DirPath.t * Names.Id.t + val pr_qualid : qualid -> Pp.std_ppcmds + val string_of_qualid : qualid -> string + val qualid_of_string : string -> qualid + val qualid_of_path : full_path -> qualid + val qualid_of_dirpath : Names.DirPath.t -> qualid + val qualid_of_ident : Names.Id.t -> qualid + + type reference = Prelude.reference = + | Qualid of Libnames.qualid Loc.located + | Ident of Names.Id.t Loc.located + val loc_of_reference : reference -> Loc.t option + val qualid_of_reference : reference -> qualid Loc.located + val pr_reference : reference -> Pp.std_ppcmds + + val is_dirpath_prefix_of : Names.DirPath.t -> Names.DirPath.t -> bool + val split_dirpath : Names.DirPath.t -> Names.DirPath.t * Names.Id.t + val dirpath_of_string : string -> Names.DirPath.t + val pr_dirpath : Names.DirPath.t -> Pp.std_ppcmds + + val string_of_path : full_path -> string + val basename : full_path -> Names.Id.t + + type object_name = Libnames.full_path * Names.KerName.t + type object_prefix = Names.DirPath.t * (Names.ModPath.t * Names.DirPath.t) + + module Dirset : module type of struct include Libnames.Dirset end + module Dirmap : module type of struct include Libnames.Dirmap end + module Spmap : module type of struct include Libnames.Spmap end +end + +module Libobject : +sig + type obj = Libobject.obj + type 'a substitutivity = 'a Libobject.substitutivity = + | Dispose + | Substitute of 'a + | Keep of 'a + | Anticipate of 'a + type 'a object_declaration = 'a Libobject.object_declaration = + { + object_name : string; + cache_function : Libnames.object_name * 'a -> unit; + load_function : int -> Libnames.object_name * 'a -> unit; + open_function : int -> Libnames.object_name * 'a -> unit; + classify_function : 'a -> 'a substitutivity; + subst_function : Mod_subst.substitution * 'a -> 'a; + discharge_function : Libnames.object_name * 'a -> 'a option; + rebuild_function : 'a -> 'a + } + val declare_object : 'a object_declaration -> ('a -> obj) + val default_object : string -> 'a object_declaration + val object_tag : obj -> string +end + +module Universes : +sig + type universe_binders = Universes.universe_binders + type universe_opt_subst = Universes.universe_opt_subst + val fresh_inductive_instance : Environ.env -> Names.inductive -> Term.pinductive Univ.in_universe_context_set + val new_Type : Names.DirPath.t -> Term.types + val unsafe_type_of_global : Globnames.global_reference -> Term.types + val constr_of_global : Prelude.global_reference -> Term.constr + val universes_of_constr : Term.constr -> Univ.LSet.t + val restrict_universe_context : Univ.ContextSet.t -> Univ.LSet.t -> Univ.ContextSet.t + val new_univ_level : Names.DirPath.t -> Univ.Level.t + val unsafe_constr_of_global : Globnames.global_reference -> Term.constr Univ.in_universe_context + val new_sort_in_family : Sorts.family -> Sorts.t + val pr_with_global_universes : Univ.Level.t -> Pp.std_ppcmds + val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds + type universe_constraint = Universes.universe_constraint + module Constraints : + sig + type t = Universes.Constraints.t + val pr : t -> Pp.std_ppcmds + end +end + +module Global : +sig + val env : unit -> Environ.env + val lookup_mind : Names.MutInd.t -> Declarations.mutual_inductive_body + val lookup_constant : Names.Constant.t -> Declarations.constant_body + val lookup_module : Names.ModPath.t -> Declarations.module_body + val lookup_modtype : Names.ModPath.t -> Declarations.module_type_body + val lookup_inductive : Names.inductive -> Declarations.mutual_inductive_body * Declarations.one_inductive_body + val constant_of_delta_kn : Names.KerName.t -> Names.Constant.t + val register : + Retroknowledge.field -> Term.constr -> Term.constr -> unit + val env_of_context : Environ.named_context_val -> Environ.env + val is_polymorphic : Globnames.global_reference -> bool + + val type_of_global_unsafe : Globnames.global_reference -> Term.types + + val current_dirpath : unit -> Names.DirPath.t + val body_of_constant_body : Declarations.constant_body -> Term.constr option + val body_of_constant : Names.Constant.t -> Term.constr option + val add_constraints : Univ.Constraint.t -> unit +end + +module Lib : sig + type is_type = bool + type export = bool option + type node = Lib.node = + | Leaf of Libobject.obj (* FIX: horrible hack (wrt. Enrico) *) + | CompilingLibrary of Libnames.object_prefix + | OpenedModule of is_type * export * Libnames.object_prefix * Summary.frozen + | ClosedModule of library_segment + | OpenedSection of Libnames.object_prefix * Summary.frozen + | ClosedSection of library_segment + + and library_segment = (Libnames.object_name * node) list + + val current_mp : unit -> Names.ModPath.t + val is_modtype : unit -> bool + val is_module : unit -> bool + val sections_are_opened : unit -> bool + val add_anonymous_leaf : ?cache_first:bool -> Libobject.obj -> unit + val contents : unit -> library_segment + val cwd : unit -> Names.DirPath.t + val add_leaf : Names.Id.t -> Libobject.obj -> Libnames.object_name + val make_kn : Names.Id.t -> Names.KerName.t + val make_path : Names.Id.t -> Libnames.full_path + val discharge_con : Names.Constant.t -> Names.Constant.t + val discharge_inductive : Names.inductive -> Names.inductive +end + +module Library : +sig + val library_is_loaded : Names.DirPath.t -> bool + val loaded_libraries : unit -> Names.DirPath.t list +end + +module Summary : +sig + type marshallable = Summary.marshallable + type 'a summary_declaration = 'a Summary.summary_declaration = + { freeze_function : marshallable -> 'a; + unfreeze_function : 'a -> unit; + init_function : unit -> unit; } + val ref : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref + val declare_summary : string -> 'a summary_declaration -> unit + module Local : + sig + type 'a local_ref = 'a Summary.Local.local_ref + val ref : ?freeze:('a -> 'a) -> name:string -> 'a -> 'a local_ref + val (:=) : 'a local_ref -> 'a -> unit + val (!) : 'a local_ref -> 'a + end +end + +module Declare : +sig + type internal_flag = Declare.internal_flag = + | UserAutomaticRequest + | InternalTacticRequest + | UserIndividualRequest + type constant_declaration = Safe_typing.private_constants Entries.constant_entry * Decl_kinds.logical_kind + type section_variable_entry = Declare.section_variable_entry = + | SectionLocalDef of Safe_typing.private_constants Entries.definition_entry + | SectionLocalAssum of Term.types Univ.in_universe_context_set * Decl_kinds.polymorphic * bool + type variable_declaration = Names.DirPath.t * section_variable_entry * Decl_kinds.logical_kind + val declare_constant : + ?internal:internal_flag -> ?local:bool -> Names.Id.t -> ?export_seff:bool -> constant_declaration -> Names.Constant.t + val declare_universe_context : Decl_kinds.polymorphic -> Univ.ContextSet.t -> unit + val declare_definition : + ?internal:internal_flag -> ?opaque:bool -> ?kind:Decl_kinds.definition_object_kind -> + ?local:bool -> ?poly:Decl_kinds.polymorphic -> Names.Id.t -> ?types:Term.constr -> + Term.constr Univ.in_universe_context_set -> Names.Constant.t + val definition_entry : ?fix_exn:Future.fix_exn -> + ?opaque:bool -> ?inline:bool -> ?types:Term.types -> + ?poly:Decl_kinds.polymorphic -> ?univs:Univ.UContext.t -> + ?eff:Safe_typing.private_constants -> Term.constr -> Safe_typing.private_constants Entries.definition_entry + val definition_message : Names.Id.t -> unit + val declare_variable : Names.Id.t -> variable_declaration -> Libnames.object_name +end + +module Reductionops : +sig + type local_reduction_function = Evd.evar_map -> EConstr.constr -> EConstr.constr + + type reduction_function = Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr + + type local_stack_reduction_function = + Evd.evar_map -> EConstr.constr -> EConstr.constr * EConstr.constr list + + type e_reduction_function = Environ.env -> Evd.evar_map -> EConstr.constr -> Evd.evar_map * EConstr.constr + type state = Reductionops.state + + val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function + val nf_beta : local_reduction_function + val nf_betaiota : local_reduction_function + val splay_prod : Environ.env -> Evd.evar_map -> EConstr.constr -> + (Names.Name.t * EConstr.constr) list * EConstr.constr + val splay_prod_n : Environ.env -> Evd.evar_map -> int -> EConstr.constr -> EConstr.rel_context * EConstr.constr + val whd_all : reduction_function + val whd_beta : local_reduction_function + + val whd_betaiotazeta : local_reduction_function + + val whd_betaiota_stack : local_stack_reduction_function + + val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function + val is_conv : ?reds:Names.transparent_state -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool + val beta_applist : Evd.evar_map -> EConstr.constr * EConstr.constr list -> EConstr.constr + val sort_of_arity : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.ESorts.t + val is_conv_leq : ?reds:Names.transparent_state -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool + val whd_betaiota : local_reduction_function + val is_arity : Environ.env -> Evd.evar_map -> EConstr.constr -> bool + val nf_evar : Evd.evar_map -> EConstr.constr -> EConstr.constr + val nf_meta : Evd.evar_map -> EConstr.constr -> EConstr.constr + val hnf_prod_appvect : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr array -> EConstr.constr + val pr_state : state -> Pp.std_ppcmds + module Stack : + sig + type 'a t = 'a Reductionops.Stack.t + val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds + end + module Cst_stack : + sig + type t = Reductionops.Cst_stack.t + val pr : t -> Pp.std_ppcmds + end +end + +module Inductiveops : +sig + type inductive_family = Inductiveops.inductive_family + type inductive_type = Inductiveops.inductive_type = + | IndType of inductive_family * EConstr.constr list + type constructor_summary = Inductiveops.constructor_summary = + { + cs_cstr : Term.pconstructor; + cs_params : Term.constr list; + cs_nargs : int; + cs_args : Context.Rel.t; + cs_concl_realargs : Term.constr array; + } + + val arities_of_constructors : Environ.env -> Term.pinductive -> Term.types array + val constructors_nrealargs_env : Environ.env -> Names.inductive -> int array + val constructor_nallargs_env : Environ.env -> Names.constructor -> int + + val inductive_nparams : Names.inductive -> int + + val inductive_nparamdecls : Names.inductive -> int + + val type_of_constructors : Environ.env -> Term.pinductive -> Term.types array + val find_mrectype : Environ.env -> Evd.evar_map -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.constr list + val mis_is_recursive : + Names.inductive * Declarations.mutual_inductive_body * Declarations.one_inductive_body -> bool + val nconstructors : Names.inductive -> int + val find_rectype : Environ.env -> Evd.evar_map -> EConstr.types -> inductive_type + val get_constructors : Environ.env -> inductive_family -> constructor_summary array + val dest_ind_family : inductive_family -> Names.inductive Term.puniverses * Term.constr list + val find_inductive : Environ.env -> Evd.evar_map -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * Term.constr list + val type_of_inductive : Environ.env -> Term.pinductive -> Term.types +end + +module Recordops : +sig + type cs_pattern = Recordops.cs_pattern = + | Const_cs of Globnames.global_reference + | Prod_cs + | Sort_cs of Sorts.family + | Default_cs + type obj_typ = Recordops.obj_typ = { + o_DEF : Term.constr; + o_CTX : Univ.ContextSet.t; + o_INJ : int option; (** position of trivial argument *) + o_TABS : Term.constr list; (** ordered *) + o_TPARAMS : Term.constr list; (** ordered *) + o_NPARAMS : int; + o_TCOMPS : Term.constr list } + val lookup_projections : Names.inductive -> Names.Constant.t option list + val lookup_canonical_conversion : (Globnames.global_reference * cs_pattern) -> Term.constr * obj_typ + val find_projection_nparams : Globnames.global_reference -> int +end + +module Retyping : (* reconstruct the type of a term knowing that it was already typechecked *) +sig + val get_type_of : ?polyprop:bool -> ?lax:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types + val get_sort_family_of : ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.family + val expand_projection : Environ.env -> Evd.evar_map -> Names.Projection.t -> EConstr.constr -> EConstr.constr list -> EConstr.constr + val get_sort_of : + ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.t +end + +module Typing : +sig + val e_sort_of : Environ.env -> Evd.evar_map ref -> EConstr.types -> Sorts.t + + val type_of : ?refresh:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> Evd.evar_map * EConstr.types + val e_solve_evars : Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.constr + + val unsafe_type_of : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types + + val e_check : Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.types -> unit + + val e_type_of : ?refresh:bool -> Environ.env -> Evd.evar_map ref -> EConstr.constr -> EConstr.types +end + +module Evarsolve : +sig + val refresh_universes : + ?status:Evd.rigid -> ?onlyalg:bool -> ?refreshset:bool -> bool option -> + Environ.env -> Evd.evar_map -> EConstr.types -> Evd.evar_map * EConstr.types +end + +module Constr_matching : +sig + val special_meta : Prelude.metavariable + + type binding_bound_vars = Names.Id.Set.t + type bound_ident_map = Names.Id.t Names.Id.Map.t + val is_matching : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> bool + val extended_matches : + Environ.env -> Evd.evar_map -> binding_bound_vars * Pattern.constr_pattern -> + EConstr.constr -> bound_ident_map * Pattern.extended_patvar_map + exception PatternMatchingFailure + type matching_result = + { m_sub : bound_ident_map * Pattern.patvar_map; + m_ctx : EConstr.constr } + val match_subterm_gen : Environ.env -> Evd.evar_map -> + bool -> + binding_bound_vars * Pattern.constr_pattern -> EConstr.constr -> + matching_result IStream.t + val matches : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> Pattern.patvar_map +end + +module Tactypes : +sig + type glob_constr_and_expr = Glob_term.glob_constr * Constrexpr.constr_expr option + type glob_constr_pattern_and_expr = Names.Id.Set.t * glob_constr_and_expr * Pattern.constr_pattern + type 'a delayed_open = Environ.env -> Evd.evar_map -> Evd.evar_map * 'a + type delayed_open_constr = EConstr.constr delayed_open + type delayed_open_constr_with_bindings = EConstr.constr Misctypes.with_bindings delayed_open + type intro_pattern = delayed_open_constr Misctypes.intro_pattern_expr Loc.located + type intro_patterns = delayed_open_constr Misctypes.intro_pattern_expr Loc.located list + type intro_pattern_naming = Misctypes.intro_pattern_naming_expr Loc.located + type or_and_intro_pattern = delayed_open_constr Misctypes.or_and_intro_pattern_expr Loc.located +end + +module Pretyping : +sig + type typing_constraint = Pretyping.typing_constraint = + | OfType of EConstr.types + | IsType + | WithoutTypeConstraint + + type inference_hook = Environ.env -> Evd.evar_map -> Evar.t -> Evd.evar_map * EConstr.constr + type inference_flags = Pretyping.inference_flags = { + use_typeclasses : bool; + solve_unification_constraints : bool; + use_hook : inference_hook option; + fail_evar : bool; + expand_evars : bool + } + + type pure_open_constr = Evd.evar_map * EConstr.constr + type glob_constr_ltac_closure = Glob_term.ltac_var_map * Glob_term.glob_constr + + val understand_ltac : inference_flags -> + Environ.env -> Evd.evar_map -> Glob_term.ltac_var_map -> + typing_constraint -> Glob_term.glob_constr -> pure_open_constr + val understand_tcc : ?flags:inference_flags -> Environ.env -> Evd.evar_map -> + ?expected_type:typing_constraint -> Glob_term.glob_constr -> Evd.evar_map * EConstr.constr + val type_uconstr : + ?flags:inference_flags -> + ?expected_type:typing_constraint -> + Geninterp.interp_sign -> Glob_term.closed_glob_constr -> EConstr.constr Tactypes.delayed_open + val understand : ?flags:inference_flags -> ?expected_type:typing_constraint -> + Environ.env -> Evd.evar_map -> Glob_term.glob_constr -> Term.constr Evd.in_evar_universe_context + val check_evars : Environ.env -> Evd.evar_map -> Evd.evar_map -> EConstr.constr -> unit + val interp_elimination_sort : Misctypes.glob_sort -> Sorts.family + val register_constr_interp0 : + ('r, 'g, 't) Genarg.genarg_type -> + (Glob_term.unbound_ltac_var_map -> Environ.env -> Evd.evar_map -> EConstr.types -> 'g -> EConstr.constr * Evd.evar_map) -> unit + val all_and_fail_flags : inference_flags + val ise_pretype_gen : + inference_flags -> Environ.env -> Evd.evar_map -> + Glob_term.ltac_var_map -> typing_constraint -> Glob_term.glob_constr -> Evd.evar_map * EConstr.constr +end + +module Evarconv : +sig + val e_conv : Environ.env -> ?ts:Names.transparent_state -> Evd.evar_map ref -> EConstr.constr -> EConstr.constr -> bool + val the_conv_x : Environ.env -> ?ts:Names.transparent_state -> EConstr.constr -> EConstr.constr -> Evd.evar_map -> Evd.evar_map + val the_conv_x_leq : Environ.env -> ?ts:Names.transparent_state -> EConstr.constr -> EConstr.constr -> Evd.evar_map -> Evd.evar_map + val solve_unif_constraints_with_heuristics : Environ.env -> ?ts:Names.transparent_state -> Evd.evar_map -> Evd.evar_map +end + +module Unification : +sig + type core_unify_flags = Unification.core_unify_flags = + { + modulo_conv_on_closed_terms : Names.transparent_state option; + use_metas_eagerly_in_conv_on_closed_terms : bool; + use_evars_eagerly_in_conv_on_closed_terms : bool; + modulo_delta : Names.transparent_state; + modulo_delta_types : Names.transparent_state; + check_applied_meta_types : bool; + use_pattern_unification : bool; + use_meta_bound_pattern_unification : bool; + frozen_evars : Evar.Set.t; + restrict_conv_on_strict_subterms : bool; + modulo_betaiota : bool; + modulo_eta : bool; + } + type unify_flags = Unification.unify_flags = + { + core_unify_flags : core_unify_flags; + merge_unify_flags : core_unify_flags; + subterm_unify_flags : core_unify_flags; + allow_K_in_toplevel_higher_order_unification : bool; + resolve_evars : bool + } + val default_no_delta_unify_flags : unit -> unify_flags + val w_unify : Environ.env -> Evd.evar_map -> Reduction.conv_pb -> ?flags:unify_flags -> EConstr.constr -> EConstr.constr -> Evd.evar_map + val elim_flags : unit -> unify_flags + val w_unify_to_subterm : + Environ.env -> Evd.evar_map -> ?flags:unify_flags -> EConstr.constr * EConstr.constr -> Evd.evar_map * EConstr.constr +end + +module Typeclasses : +sig + type typeclass = Typeclasses.typeclass = { + cl_impl : Globnames.global_reference; + cl_context : (Globnames.global_reference * bool) option list * Context.Rel.t; + cl_props : Context.Rel.t; + cl_projs : (Names.Name.t * (direction * Vernacexpr.hint_info_expr) option + * Names.Constant.t option) list; + cl_strict : bool; + cl_unique : bool; + } + and direction = Typeclasses.direction + type instance = Typeclasses.instance + type evar_filter = Evar.t -> Evar_kinds.t -> bool + val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool -> + ?split:bool -> ?fail:bool -> Environ.env -> Evd.evar_map -> Evd.evar_map + val set_resolvable : Evd.Store.t -> bool -> Evd.Store.t + val resolve_one_typeclass : ?unique:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Evd.evar_map * EConstr.constr + val class_info : Globnames.global_reference -> typeclass + val mark_resolvables : ?filter:evar_filter -> Evd.evar_map -> Evd.evar_map + val add_instance : instance -> unit + val new_instance : typeclass -> Vernacexpr.hint_info_expr -> bool -> Decl_kinds.polymorphic -> + Globnames.global_reference -> instance +end + +module Pretype_errors : +sig + type unification_error = Pretype_errors.unification_error + type subterm_unification_error = Pretype_errors.subterm_unification_error + type pretype_error = Pretype_errors.pretype_error = + | CantFindCaseType of EConstr.constr + | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error + | UnifOccurCheck of Evar.t * EConstr.constr + | UnsolvableImplicit of Evar.t * Evd.unsolvability_explanation option + | CannotUnify of EConstr.constr * EConstr.constr * unification_error option + | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr + | CannotUnifyBindingType of EConstr.constr * EConstr.constr + | CannotGeneralize of EConstr.constr + | NoOccurrenceFound of EConstr.constr * Names.Id.t option + | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (Environ.env * Pretype_errors.type_error) option + | WrongAbstractionType of Names.Name.t * EConstr.constr * EConstr.types * EConstr.types + | AbstractionOverMeta of Names.Name.t * Names.Name.t + | NonLinearUnification of Names.Name.t * EConstr.constr + | VarNotFound of Names.Id.t + | UnexpectedType of EConstr.constr * EConstr.constr + | NotProduct of EConstr.constr + | TypingError of Pretype_errors.type_error + | CannotUnifyOccurrences of subterm_unification_error + | UnsatisfiableConstraints of + (Evar.t * Evar_kinds.t) option * Evar.Set.t option + + exception PretypeError of Environ.env * Evd.evar_map * pretype_error + val error_var_not_found : ?loc:Loc.t -> Names.Id.t -> 'b + val precatchable_exception : exn -> bool +end + +module Smartlocate : +sig + val locate_global_with_alias : ?head:bool -> Libnames.qualid Loc.located -> Globnames.global_reference + val global_with_alias : ?head:bool -> Prelude.reference -> Globnames.global_reference + val global_of_extended_global : Globnames.extended_global_reference -> Globnames.global_reference + val loc_of_smart_reference : Prelude.reference Misctypes.or_by_notation -> Loc.t option + val smart_global : ?head:bool -> Prelude.reference Misctypes.or_by_notation -> Globnames.global_reference +end + +module Dumpglob : +sig + val add_glob : ?loc:Loc.t -> Globnames.global_reference -> unit + val pause : unit -> unit + val continue : unit -> unit +end + +module Stdarg : +sig + val loc_of_or_by_notation : ('a -> Loc.t option) -> 'a Misctypes.or_by_notation -> Loc.t option + val wit_unit : unit Genarg.uniform_genarg_type + val wit_int : int Genarg.uniform_genarg_type + val wit_var : (Names.Id.t Loc.located, Names.Id.t Loc.located, Names.Id.t) Genarg.genarg_type + val wit_bool : bool Genarg.uniform_genarg_type + val wit_string : string Genarg.uniform_genarg_type + val wit_pre_ident : string Genarg.uniform_genarg_type + val wit_global : (Prelude.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type + val wit_ident : Names.Id.t Genarg.uniform_genarg_type + val wit_integer : int Genarg.uniform_genarg_type + val wit_constr : (Constrexpr.constr_expr, Tactypes.glob_constr_and_expr, EConstr.constr) Genarg.genarg_type + val wit_open_constr : (Constrexpr.constr_expr, Tactypes.glob_constr_and_expr, EConstr.constr) Genarg.genarg_type + val wit_intro_pattern : (Constrexpr.constr_expr Misctypes.intro_pattern_expr Loc.located, Tactypes.glob_constr_and_expr Misctypes.intro_pattern_expr Loc.located, Tactypes.intro_pattern) Genarg.genarg_type + val wit_int_or_var : (int Misctypes.or_var, int Misctypes.or_var, int) Genarg.genarg_type + val wit_ref : (Prelude.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type + val wit_clause_dft_concl : (Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Locus.clause_expr) Genarg.genarg_type + val wit_uconstr : (Constrexpr.constr_expr , Tactypes.glob_constr_and_expr, Glob_term.closed_glob_constr) Genarg.genarg_type + val wit_red_expr : + ((Constrexpr.constr_expr,Prelude.reference Misctypes.or_by_notation,Constrexpr.constr_expr) Genredexpr.red_expr_gen, + (Tactypes.glob_constr_and_expr,Names.evaluable_global_reference Misctypes.and_short_name Misctypes.or_var,Tactypes.glob_constr_pattern_and_expr) Genredexpr.red_expr_gen, + (EConstr.constr,Names.evaluable_global_reference,Pattern.constr_pattern) Genredexpr.red_expr_gen) Genarg.genarg_type + val wit_quant_hyp : Misctypes.quantified_hypothesis Genarg.uniform_genarg_type + val wit_bindings : + (Constrexpr.constr_expr Misctypes.bindings, + Tactypes.glob_constr_and_expr Misctypes.bindings, + EConstr.constr Misctypes.bindings Tactypes.delayed_open) Genarg.genarg_type + val wit_constr_with_bindings : + (Constrexpr.constr_expr Misctypes.with_bindings, + Tactypes.glob_constr_and_expr Misctypes.with_bindings, + EConstr.constr Misctypes.with_bindings Tactypes.delayed_open) Genarg.genarg_type + val wit_intropattern : (Constrexpr.constr_expr Misctypes.intro_pattern_expr Loc.located, Tactypes.glob_constr_and_expr Misctypes.intro_pattern_expr Loc.located, Tactypes.intro_pattern) Genarg.genarg_type + val wit_quantified_hypothesis : Misctypes.quantified_hypothesis Genarg.uniform_genarg_type + val wit_clause : (Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Loc.located Locus.clause_expr,Names.Id.t Locus.clause_expr) Genarg.genarg_type + val wit_preident : string Genarg.uniform_genarg_type + val wit_reference : (Prelude.reference, Globnames.global_reference Loc.located Misctypes.or_var, Globnames.global_reference) Genarg.genarg_type + val wit_open_constr_with_bindings : + (Constrexpr.constr_expr Misctypes.with_bindings, + Tactypes.glob_constr_and_expr Misctypes.with_bindings, + EConstr.constr Misctypes.with_bindings Tactypes.delayed_open) Genarg.genarg_type +end + +module Coqlib : +sig + type coq_eq_data = Coqlib.coq_eq_data = { eq : Globnames.global_reference; + ind : Globnames.global_reference; + refl : Globnames.global_reference; + sym : Globnames.global_reference; + trans: Globnames.global_reference; + congr: Globnames.global_reference; + } + type coq_sigma_data = Coqlib.coq_sigma_data = { + proj1 : Globnames.global_reference; + proj2 : Globnames.global_reference; + elim : Globnames.global_reference; + intro : Globnames.global_reference; + typ : Globnames.global_reference } + val gen_reference : string -> string list -> string -> Globnames.global_reference + val find_reference : string -> string list -> string -> Globnames.global_reference + val check_required_library : string list -> unit + val logic_module_name : string list + val glob_true : Globnames.global_reference + val glob_false : Globnames.global_reference + val glob_O : Globnames.global_reference + val glob_S : Globnames.global_reference + val nat_path : Libnames.full_path + val datatypes_module_name : string list + val glob_eq : Globnames.global_reference + val build_coq_eq_sym : Globnames.global_reference Util.delayed + val build_coq_False : Globnames.global_reference Util.delayed + val build_coq_not : Globnames.global_reference Util.delayed + val build_coq_eq : Globnames.global_reference Util.delayed + val build_coq_eq_data : coq_eq_data Util.delayed + val path_of_O : Names.constructor + val path_of_S : Names.constructor + val build_prod : coq_sigma_data Util.delayed + val build_coq_True : Globnames.global_reference Util.delayed + val coq_iff_ref : Globnames.global_reference lazy_t + val build_coq_iff_left_proj : Globnames.global_reference Util.delayed + val build_coq_iff_right_proj : Globnames.global_reference Util.delayed + val init_modules : string list list + val build_coq_eq_refl : Globnames.global_reference Util.delayed + val arith_modules : string list list + val zarith_base_modules : string list list + val gen_reference_in_modules : string -> string list list-> string -> Globnames.global_reference + val jmeq_module_name : string list + val coq_eq_ref : Globnames.global_reference lazy_t + val coq_not_ref : Globnames.global_reference lazy_t + val coq_or_ref : Globnames.global_reference lazy_t + val build_coq_and : Globnames.global_reference Util.delayed + val build_coq_I : Globnames.global_reference Util.delayed + val coq_reference : string -> string list -> string -> Globnames.global_reference +end + +module Impargs : +sig + type implicit_status = Impargs.implicit_status + type implicit_side_condition = Impargs.implicit_side_condition + type implicits_list = implicit_side_condition * implicit_status list + type manual_explicitation = Constrexpr.explicitation * (bool * bool * bool) + type manual_implicits = manual_explicitation list + val is_status_implicit : implicit_status -> bool + val name_of_implicit : implicit_status -> Names.Id.t + val implicits_of_global : Globnames.global_reference -> implicits_list list + val declare_manual_implicits : bool -> Globnames.global_reference -> ?enriching:bool -> + manual_implicits list -> unit + val is_implicit_args : unit -> bool + val is_strict_implicit_args : unit -> bool + val is_contextual_implicit_args : unit -> bool + val make_implicit_args : bool -> unit + val make_strict_implicit_args : bool -> unit + val make_contextual_implicit_args : bool -> unit +end + +module Constrintern : +sig + type ltac_sign = Constrintern.ltac_sign = { + ltac_vars : Names.Id.Set.t; + ltac_bound : Names.Id.Set.t; + ltac_extra : Genintern.Store.t; + } + type var_internalization_data = Constrintern.var_internalization_data + type var_internalization_type = Constrintern.var_internalization_type = + | Inductive of Names.Id.t list * bool + | Recursive + | Method + | Variable + type internalization_env = var_internalization_data Names.Id.Map.t + + val interp_constr_evars : Environ.env -> Evd.evar_map ref -> + ?impls:internalization_env -> Constrexpr.constr_expr -> EConstr.constr + + val interp_type_evars : Environ.env -> Evd.evar_map ref -> + ?impls:internalization_env -> Constrexpr.constr_expr -> EConstr.types + + val empty_ltac_sign : ltac_sign + val intern_gen : Pretyping.typing_constraint -> Environ.env -> + ?impls:internalization_env -> ?pattern_mode:bool -> ?ltacvars:ltac_sign -> + Constrexpr.constr_expr -> Glob_term.glob_constr + val intern_constr_pattern : + Environ.env -> ?as_type:bool -> ?ltacvars:ltac_sign -> + Constrexpr.constr_pattern_expr -> Names.Id.t list * Pattern.constr_pattern + val intern_constr : Environ.env -> Constrexpr.constr_expr -> Glob_term.glob_constr + val for_grammar : ('a -> 'b) -> 'a -> 'b + val interp_reference : ltac_sign -> Prelude.reference -> Glob_term.glob_constr + val interp_constr : Environ.env -> Evd.evar_map -> ?impls:internalization_env -> + Constrexpr.constr_expr -> Term.constr Evd.in_evar_universe_context + val interp_open_constr : Environ.env -> Evd.evar_map -> Constrexpr.constr_expr -> Evd.evar_map * EConstr.constr + val locate_reference : Libnames.qualid -> Globnames.global_reference + val interp_type : Environ.env -> Evd.evar_map -> ?impls:internalization_env -> + Constrexpr.constr_expr -> Term.types Evd.in_evar_universe_context + val interp_context_evars : + ?global_level:bool -> ?impl_env:internalization_env -> ?shift:int -> + Environ.env -> Evd.evar_map ref -> Constrexpr.local_binder_expr list -> + internalization_env * ((Environ.env * EConstr.rel_context) * Impargs.manual_implicits) + val compute_internalization_data : Environ.env -> var_internalization_type -> + Term.types -> Impargs.manual_explicitation list -> var_internalization_data + val empty_internalization_env : internalization_env + val global_reference : Names.Id.t -> Globnames.global_reference +end + +module Notation_term : +sig + type scope_name = string + type notation_var_instance_type = Notation_term.notation_var_instance_type = + | NtnTypeConstr | NtnTypeOnlyBinder | NtnTypeConstrList | NtnTypeBinderList + type tmp_scope_name = Notation_term.tmp_scope_name + type subscopes = tmp_scope_name option * scope_name list + type notation_constr = Notation_term.notation_constr = + | NRef of Globnames.global_reference + | NVar of Names.Id.t + | NApp of notation_constr * notation_constr list + | NHole of Evar_kinds.t * Misctypes.intro_pattern_naming_expr * Genarg.glob_generic_argument option + | NList of Names.Id.t * Names.Id.t * notation_constr * notation_constr * bool + | NLambda of Names.Name.t * notation_constr * notation_constr + | NProd of Names.Name.t * notation_constr * notation_constr + | NBinderList of Names.Id.t * Names.Id.t * notation_constr * notation_constr + | NLetIn of Names.Name.t * notation_constr * notation_constr option * notation_constr + | NCases of Term.case_style * notation_constr option * + (notation_constr * (Names.Name.t * (Names.inductive * Names.Name.t list) option)) list * + (Glob_term.cases_pattern list * notation_constr) list + | NLetTuple of Names.Name.t list * (Names.Name.t * notation_constr option) * + notation_constr * notation_constr + | NIf of notation_constr * (Names.Name.t * notation_constr option) * + notation_constr * notation_constr + | NRec of Glob_term.fix_kind * Names.Id.t array * + (Names.Name.t * notation_constr option * notation_constr) list array * + notation_constr array * notation_constr array + | NSort of Misctypes.glob_sort + | NCast of notation_constr * notation_constr Misctypes.cast_type + type interpretation = (Names.Id.t * (subscopes * notation_var_instance_type)) list * + notation_constr +end + +module Notation : +sig + type cases_pattern_status = bool + type required_module = Libnames.full_path * string list + type 'a prim_token_interpreter = ?loc:Loc.t -> 'a -> Glob_term.glob_constr + type 'a prim_token_uninterpreter = Glob_term.glob_constr list * (Glob_term.glob_constr -> 'a option) * cases_pattern_status + type delimiters = string + type local_scopes = Notation_term.tmp_scope_name option * Notation_term.scope_name list + type notation_location = (Names.DirPath.t * Names.DirPath.t) * string + val declare_string_interpreter : Notation_term.scope_name -> required_module -> + string prim_token_interpreter -> string prim_token_uninterpreter -> unit + val declare_numeral_interpreter : Notation_term.scope_name -> required_module -> + Bigint.bigint prim_token_interpreter -> Bigint.bigint prim_token_uninterpreter -> unit + val interp_notation_as_global_reference : ?loc:Loc.t -> (Globnames.global_reference -> bool) -> + Constrexpr.notation -> delimiters option -> Globnames.global_reference + val locate_notation : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Constrexpr.notation -> + Notation_term.scope_name option -> Pp.std_ppcmds + val find_delimiters_scope : ?loc:Loc.t -> delimiters -> Notation_term.scope_name + val pr_scope : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Notation_term.scope_name -> Pp.std_ppcmds + val pr_scopes : (Glob_term.glob_constr -> Pp.std_ppcmds) -> Pp.std_ppcmds + val interp_notation : ?loc:Loc.t -> Constrexpr.notation -> local_scopes -> + Notation_term.interpretation * (notation_location * Notation_term.scope_name option) + val uninterp_prim_token : Glob_term.glob_constr -> Notation_term.scope_name * Constrexpr.prim_token +end + +module Mltop : +sig + val declare_cache_obj : (unit -> unit) -> string -> unit + val add_known_plugin : (unit -> unit) -> string -> unit + val add_known_module : string -> unit + val module_is_known : string -> bool +end + +(* All items in the Proof_type module are deprecated. *) +module Proof_type : +sig + type goal = Evar.t + type rule = Proof_type.prim_rule = + | Cut of bool * bool * Names.Id.t * Term.types + | Refine of Term.constr + + type tactic = goal Evd.sigma -> goal list Evd.sigma +end + +module Redexpr : +sig + type red_expr = + (EConstr.constr, Names.evaluable_global_reference, Pattern.constr_pattern) Genredexpr.red_expr_gen + val reduction_of_red_expr : + Environ.env -> red_expr -> Reductionops.e_reduction_function * Term.cast_kind + val declare_reduction : string -> Reductionops.reduction_function -> unit +end + +module Tacmach : +sig + type tactic = Proof_type.tactic + [@@ocaml.deprecated "alias for API.Proof_type.tactic"] + + type 'a sigma = 'a Evd.sigma + [@@ocaml.deprecated "alias of API.Evd.sigma"] + + val re_sig : 'a -> Evd.evar_map -> 'a Evd.sigma + + val pf_reduction_of_red_expr : Proof_type.goal Evd.sigma -> Redexpr.red_expr -> EConstr.constr -> Evd.evar_map * EConstr.constr + + val pf_unsafe_type_of : Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.types + + val pf_get_new_id : Names.Id.t -> Proof_type.goal Evd.sigma -> Names.Id.t + + val pf_env : Proof_type.goal Evd.sigma -> Environ.env + + val pf_concl : Proof_type.goal Evd.sigma -> EConstr.types + + val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> Proof_type.goal Evd.sigma -> 'a + + val pf_get_hyp : Proof_type.goal Evd.sigma -> Names.Id.t -> EConstr.named_declaration + val pf_get_hyp_typ : Proof_type.goal Evd.sigma -> Names.Id.t -> EConstr.types + val project : Proof_type.goal Evd.sigma -> Evd.evar_map + val refine : EConstr.constr -> Proof_type.tactic + val pf_type_of : Proof_type.goal Evd.sigma -> EConstr.constr -> Evd.evar_map * EConstr.types + + val pf_hyps : Proof_type.goal Evd.sigma -> EConstr.named_context + + val pf_ids_of_hyps : Proof_type.goal Evd.sigma -> Names.Id.t list + + val pf_reduce_to_atomic_ind : Proof_type.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types + + val pf_reduce_to_quantified_ind : Proof_type.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types + + val pf_eapply : (Environ.env -> Evd.evar_map -> 'a -> Evd.evar_map * 'b) -> + Evar.t Evd.sigma -> 'a -> Evar.t Evd.sigma * 'b + + val pf_unfoldn : (Locus.occurrences * Names.evaluable_global_reference) list + -> Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr + + val pf_reduce : (Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr) -> Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr + + val pf_conv_x : Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr -> bool + + val pf_is_matching : Proof_type.goal Evd.sigma -> Pattern.constr_pattern -> EConstr.constr -> bool + + val pf_hyps_types : Proof_type.goal Evd.sigma -> (Names.Id.t * EConstr.types) list + + val pr_gls : Proof_type.goal Evd.sigma -> Pp.std_ppcmds + + val pf_nf_betaiota : Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr + + val pf_last_hyp : Proof_type.goal Evd.sigma -> EConstr.named_declaration + + val pf_nth_hyp_id : Proof_type.goal Evd.sigma -> int -> Names.Id.t + + val sig_it : 'a Evd.sigma -> 'a + + module New : + sig + val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> 'b Proofview.Goal.t -> 'a + val project : 'a Proofview.Goal.t -> Evd.evar_map + val pf_unsafe_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types + val of_old : (Proof_type.goal Evd.sigma -> 'a) -> [ `NF ] Proofview.Goal.t -> 'a + + val pf_env : 'a Proofview.Goal.t -> Environ.env + val pf_ids_of_hyps : 'a Proofview.Goal.t -> Names.Id.t list + val pf_concl : 'a Proofview.Goal.t -> EConstr.types + val pf_get_new_id : Names.Id.t -> 'a Proofview.Goal.t -> Names.Id.t + val pf_get_hyp_typ : Names.Id.t -> 'a Proofview.Goal.t -> EConstr.types + val pf_get_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types + val pf_global : Names.Id.t -> 'a Proofview.Goal.t -> Globnames.global_reference + val pf_hyps_types : 'a Proofview.Goal.t -> (Names.Id.t * EConstr.types) list + end +end + +module Proof : +sig + type proof = Proof.proof + type 'a focus_kind = 'a Proof.focus_kind + val run_tactic : Environ.env -> + unit Proofview.tactic -> proof -> proof * (bool * Proofview_monad.Info.tree) + val unshelve : proof -> proof + val maximal_unfocus : 'a focus_kind -> proof -> proof + val pr_proof : proof -> Pp.std_ppcmds + module V82 : + sig + val grab_evars : proof -> proof + + val subgoals : proof -> Goal.goal list Evd.sigma + end +end + +module Proof_global : +sig + type proof_mode = Proof_global.proof_mode = { + name : string; + set : unit -> unit ; + reset : unit -> unit + } + type proof_universes = UState.t * Universes.universe_binders option + type proof_object = Proof_global.proof_object = { + id : Names.Id.t; + entries : Safe_typing.private_constants Entries.definition_entry list; + persistence : Decl_kinds.goal_kind; + universes: proof_universes; + } + type proof_ending = Proof_global.proof_ending = + | Admitted of Names.Id.t * Decl_kinds.goal_kind * Entries.parameter_entry * + proof_universes + | Proved of Vernacexpr.opacity_flag * + Vernacexpr.lident option * + proof_object + type proof_terminator = Proof_global.proof_terminator + type lemma_possible_guards = Proof_global.lemma_possible_guards + type universe_binders = Proof_global.universe_binders + type closed_proof = proof_object * proof_terminator + val make_terminator : (proof_ending -> unit) -> proof_terminator + val start_dependent_proof : + Names.Id.t -> ?pl:universe_binders -> Decl_kinds.goal_kind -> + Proofview.telescope -> proof_terminator -> unit + val with_current_proof : + (unit Proofview.tactic -> Proof.proof -> Proof.proof * 'a) -> 'a + val simple_with_current_proof : + (unit Proofview.tactic -> Proof.proof -> Proof.proof) -> unit + val compact_the_proof : unit -> unit + val register_proof_mode : proof_mode -> unit + val get_default_goal_selector : unit -> Vernacexpr.goal_selector + + exception NoCurrentProof + val give_me_the_proof : unit -> Proof.proof + (** @raise NoCurrentProof when outside proof mode. *) + + val discard_all : unit -> unit + val discard_current : unit -> unit + val get_current_proof_name : unit -> Names.Id.t +end + +module Nametab : +sig + exception GlobalizationError of Libnames.qualid + + type ltac_constant = Names.KerName.t + + val global_of_path : Libnames.full_path -> Globnames.global_reference + val shortest_qualid_of_global : Names.Id.Set.t -> Globnames.global_reference -> Libnames.qualid + val path_of_global : Globnames.global_reference -> Libnames.full_path + val locate_extended : Libnames.qualid -> Globnames.extended_global_reference + val full_name_module : Libnames.qualid -> Names.DirPath.t + val locate_tactic : Libnames.qualid -> Names.KerName.t + val pr_global_env : Names.Id.Set.t -> Globnames.global_reference -> Pp.std_ppcmds + val shortest_qualid_of_tactic : Names.KerName.t -> Libnames.qualid + val basename_of_global : Globnames.global_reference -> Names.Id.t + + type visibility = Nametab.visibility = + | Until of int + | Exactly of int + + val push_tactic : visibility -> Libnames.full_path -> Names.KerName.t -> unit + val error_global_not_found : ?loc:Loc.t -> Libnames.qualid -> 'a + val shortest_qualid_of_module : Names.ModPath.t -> Libnames.qualid + val dirpath_of_module : Names.ModPath.t -> Names.DirPath.t + val locate_module : Libnames.qualid -> Names.ModPath.t + val dirpath_of_global : Globnames.global_reference -> Names.DirPath.t + val locate : Libnames.qualid -> Globnames.global_reference + val locate_constant : Libnames.qualid -> Names.Constant.t +end + +module Ppextend : +sig + type precedence = int + type parenRelation = Ppextend.parenRelation = + | L | E | Any | Prec of precedence + type tolerability = precedence * parenRelation +end + +module Refiner : +sig + val project : 'a Evd.sigma -> Evd.evar_map + + val unpackage : 'a Evd.sigma -> Evd.evar_map ref * 'a + + val repackage : Evd.evar_map ref -> 'a -> 'a Evd.sigma + + val refiner : Proof_type.rule -> Proof_type.tactic + + val tclSHOWHYPS : Proof_type.tactic -> Proof_type.tactic + exception FailError of int * Pp.std_ppcmds Lazy.t + + val tclEVARS : Evd.evar_map -> Proof_type.tactic + val tclMAP : ('a -> Proof_type.tactic) -> 'a list -> Proof_type.tactic + val tclREPEAT : Proof_type.tactic -> Proof_type.tactic + val tclORELSE : Proof_type.tactic -> Proof_type.tactic -> Proof_type.tactic + val tclFAIL : int -> Pp.std_ppcmds -> Proof_type.tactic + val tclIDTAC : Proof_type.tactic + val tclTHEN : Proof_type.tactic -> Proof_type.tactic -> Proof_type.tactic + val tclTHENLIST : Proof_type.tactic list -> Proof_type.tactic + val tclTRY : Proof_type.tactic -> Proof_type.tactic + val tclAT_LEAST_ONCE : Proof_type.tactic -> Proof_type.tactic +end + +module Termops : +sig + val it_mkLambda_or_LetIn : Term.constr -> Context.Rel.t -> Term.constr + val local_occur_var : Evd.evar_map -> Names.Id.t -> EConstr.constr -> bool + val occur_var : Environ.env -> Evd.evar_map -> Names.Id.t -> EConstr.constr -> bool + val pr_evar_info : Evd.evar_info -> Pp.std_ppcmds + + val print_constr : EConstr.constr -> Pp.std_ppcmds + + (** [dependent m t] tests whether [m] is a subterm of [t] *) + val dependent : Prelude.evar_map -> EConstr.constr -> EConstr.constr -> bool + + (** [pop c] returns a copy of [c] with decremented De Bruijn indexes *) + val pop : EConstr.constr -> EConstr.constr + + (** Does a given term contain an existential variable? *) + val occur_existential : Prelude.evar_map -> EConstr.constr -> bool + + (** [map_constr_with_binders_left_to_right g f acc c] maps [f updated_acc] on all the immediate subterms of [c]. + {ul {- if a given immediate subterm of [c] is not below a binder, then [updated_acc] is the same as [acc].} + {- if a given immediate subterm of [c] is below a binder [b], then [updated_acc] is computed as [g b acc].}} *) + val map_constr_with_binders_left_to_right : + Prelude.evar_map -> (EConstr.rel_declaration -> 'a -> 'a) -> ('a -> EConstr.constr -> EConstr.constr) -> 'a -> EConstr.constr -> EConstr.constr + + (** Remove the outer-most {!Term.kind_of_term.Cast} from a given term. *) + val strip_outer_cast : Prelude.evar_map -> EConstr.constr -> EConstr.constr + + (** [nb_lam] ⟦[fun (x1:t1)...(xn:tn) => c]⟧ where [c] is not an abstraction gives [n]. + Casts are ignored. *) + val nb_lam : Prelude.evar_map -> EConstr.constr -> int + + (** [push_rel_assum env_assumtion env] adds a given {i env assumption} to the {i env context} of a given {i environment}. *) + val push_rel_assum : Names.Name.t * EConstr.types -> Environ.env -> Environ.env + + (** [push_rels_assum env_assumptions env] adds given {i env assumptions} to the {i env context} of a given {i environment}. *) + val push_rels_assum : (Names.Name.t * Term.types) list -> Environ.env -> Environ.env + + type meta_value_map = Prelude.meta_value_map + + val last_arg : Evd.evar_map -> EConstr.constr -> EConstr.constr + val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Names.Name.t * 't) list + val prod_applist : Evd.evar_map -> EConstr.constr -> EConstr.constr list -> EConstr.constr + val nb_prod : Evd.evar_map -> EConstr.constr -> int + val is_section_variable : Names.Id.t -> bool + val ids_of_rel_context : ('c, 't) Context.Rel.pt -> Names.Id.t list + val subst_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> EConstr.constr + val global_vars_set_of_decl : Environ.env -> Evd.evar_map -> EConstr.named_declaration -> Names.Id.Set.t + val vars_of_env: Environ.env -> Names.Id.Set.t + val ids_of_named_context : ('c, 't) Context.Named.pt -> Names.Id.t list + val ids_of_context : Environ.env -> Names.Id.t list + val global_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference * EConstr.EInstance.t + val print_named_context : Environ.env -> Pp.std_ppcmds + val print_constr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds + val clear_named_body : Names.Id.t -> Environ.env -> Environ.env + val is_Prop : Evd.evar_map -> EConstr.constr -> bool + val is_global : Evd.evar_map -> Globnames.global_reference -> EConstr.constr -> bool + + val eq_constr : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool + + val occur_var_in_decl : + Environ.env -> Evd.evar_map -> + Names.Id.t -> EConstr.named_declaration -> bool + + val subst_meta : Prelude.meta_value_map -> Term.constr -> Term.constr + + val free_rels : Evd.evar_map -> EConstr.constr -> Int.Set.t + + val occur_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool + [@@ocaml.deprecated "alias of API.Termops.dependent"] + + val replace_term : Evd.evar_map -> EConstr.constr -> EConstr.constr -> EConstr.constr -> EConstr.constr + val map_named_decl : ('a -> 'b) -> ('a, 'a) Context.Named.Declaration.pt -> ('b, 'b) Context.Named.Declaration.pt + val map_rel_decl : ('a -> 'b) -> ('a, 'a) Context.Rel.Declaration.pt -> ('b, 'b) Context.Rel.Declaration.pt + val pr_metaset : Evd.Metaset.t -> Pp.std_ppcmds + val pr_evar_map : ?with_univs:bool -> int option -> Evd.evar_map -> Pp.std_ppcmds + val pr_evar_universe_context : UState.t -> Pp.std_ppcmds +end + +module Locality : +sig + val make_section_locality : bool option -> bool + module LocalityFixme : sig + val consume : unit -> bool option + end + val make_module_locality : bool option -> bool +end + +module Search : +sig + type glob_search_about_item = Search.glob_search_about_item = + | GlobSearchSubPattern of Pattern.constr_pattern + | GlobSearchString of string + type filter_function = Globnames.global_reference -> Environ.env -> Term.constr -> bool + type display_function = Globnames.global_reference -> Environ.env -> Term.constr -> unit + val search_about_filter : glob_search_about_item -> filter_function + val module_filter : Names.DirPath.t list * bool -> filter_function + val generic_search : int option -> display_function -> unit +end + +module Notation_ops : +sig + val glob_constr_of_notation_constr : ?loc:Loc.t -> Notation_term.notation_constr -> Glob_term.glob_constr + val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t -> + ('a -> Names.Name.t -> 'a * Names.Name.t) -> + ('a -> Notation_term.notation_constr -> Glob_term.glob_constr) -> + 'a -> Notation_term.notation_constr -> Glob_term.glob_constr +end + +module Constrextern : +sig + val extern_glob_constr : Names.Id.Set.t -> Glob_term.glob_constr -> Constrexpr.constr_expr + val extern_glob_type : Names.Id.Set.t -> Glob_term.glob_constr -> Constrexpr.constr_expr + val extern_constr : ?lax:bool -> bool -> Environ.env -> Evd.evar_map -> Term.constr -> Constrexpr.constr_expr + val without_symbols : ('a -> 'b) -> 'a -> 'b + val print_universes : bool ref + val extern_type : bool -> Environ.env -> Evd.evar_map -> Term.types -> Constrexpr.constr_expr + val with_universes : ('a -> 'b) -> 'a -> 'b + val set_extern_reference : + (?loc:Loc.t -> Names.Id.Set.t -> Globnames.global_reference -> Libnames.reference) -> unit +end + +module Patternops : +sig + val pattern_of_glob_constr : Glob_term.glob_constr -> Names.Id.t list * Pattern.constr_pattern + val subst_pattern : Mod_subst.substitution -> Pattern.constr_pattern -> Pattern.constr_pattern + val pattern_of_constr : Environ.env -> Evd.evar_map -> Term.constr -> Pattern.constr_pattern + val instantiate_pattern : Environ.env -> + Evd.evar_map -> Pattern.extended_patvar_map -> + Pattern.constr_pattern -> Pattern.constr_pattern +end + +module Printer : +sig + val pr_named_context : Environ.env -> Evd.evar_map -> Context.Named.t -> Pp.std_ppcmds + val pr_rel_context : Environ.env -> Evd.evar_map -> Context.Rel.t -> Pp.std_ppcmds + val pr_goal : Proof_type.goal Evd.sigma -> Pp.std_ppcmds + + val pr_constr_env : Prelude.env -> Prelude.evar_map -> Term.constr -> Pp.std_ppcmds + val pr_lconstr_env : Prelude.env -> Prelude.evar_map -> Term.constr -> Pp.std_ppcmds + + val pr_constr : Term.constr -> Pp.std_ppcmds + + val pr_lconstr : Term.constr -> Pp.std_ppcmds + + val pr_econstr : EConstr.constr -> Pp.std_ppcmds + val pr_glob_constr : Glob_term.glob_constr -> Pp.std_ppcmds + val pr_constr_pattern : Pattern.constr_pattern -> Pp.std_ppcmds + val pr_glob_constr_env : Environ.env -> Glob_term.glob_constr -> Pp.std_ppcmds + val pr_lglob_constr_env : Environ.env -> Glob_term.glob_constr -> Pp.std_ppcmds + val pr_econstr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds + val pr_constr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.std_ppcmds + val pr_lconstr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.std_ppcmds + val pr_closed_glob : Glob_term.closed_glob_constr -> Pp.std_ppcmds + val pr_lglob_constr : Glob_term.glob_constr -> Pp.std_ppcmds + val pr_leconstr_env : Environ.env -> Evd.evar_map -> EConstr.constr -> Pp.std_ppcmds + val pr_leconstr : EConstr.constr -> Pp.std_ppcmds + val pr_global : Globnames.global_reference -> Pp.std_ppcmds + val pr_lconstr_under_binders : Pattern.constr_under_binders -> Pp.std_ppcmds + val pr_lconstr_under_binders_env : Environ.env -> Evd.evar_map -> Pattern.constr_under_binders -> Pp.std_ppcmds + + val pr_constr_under_binders_env : Environ.env -> Evd.evar_map -> Pattern.constr_under_binders -> Pp.std_ppcmds + val pr_closed_glob_env : Environ.env -> Evd.evar_map -> Glob_term.closed_glob_constr -> Pp.std_ppcmds + val pr_rel_context_of : Environ.env -> Evd.evar_map -> Pp.std_ppcmds + val pr_named_context_of : Environ.env -> Evd.evar_map -> Pp.std_ppcmds + val pr_ltype : Term.types -> Pp.std_ppcmds + val pr_ljudge : EConstr.unsafe_judgment -> Pp.std_ppcmds * Pp.std_ppcmds + val pr_idpred : Names.Id.Pred.t -> Pp.std_ppcmds + val pr_cpred : Names.Cpred.t -> Pp.std_ppcmds + val pr_transparent_state : Names.transparent_state -> Pp.std_ppcmds +end + +module Classes : +sig + val set_typeclass_transparency : Names.evaluable_global_reference -> bool -> bool -> unit + val new_instance : + ?abstract:bool -> + ?global:bool -> + ?refine:bool -> + Decl_kinds.polymorphic -> + Constrexpr.local_binder_expr list -> + Constrexpr.typeclass_constraint -> + (bool * Constrexpr.constr_expr) option -> + ?generalize:bool -> + ?tac:unit Proofview.tactic -> + ?hook:(Globnames.global_reference -> unit) -> + Vernacexpr.hint_info_expr -> + Names.Id.t +end + +module Classops : +sig + type coe_index = Classops.coe_index + type inheritance_path = coe_index list + type cl_index = Classops.cl_index + + val hide_coercion : Globnames.global_reference -> int option + val lookup_path_to_sort_from : Environ.env -> Evd.evar_map -> EConstr.types -> + EConstr.types * inheritance_path + val get_coercion_value : coe_index -> Constr.t + val coercions : unit -> coe_index list + val pr_cl_index : cl_index -> Pp.std_ppcmds +end + +module ExplainErr : +sig + val process_vernac_interp_error : ?allow_uncaught:bool -> Util.iexn -> Util.iexn + val register_additional_error_info : (Util.iexn -> Pp.std_ppcmds option Loc.located option) -> unit +end + +module Tacred : +sig + val try_red_product : Reductionops.reduction_function + val simpl : Reductionops.reduction_function + val unfoldn : + (Locus.occurrences * Names.evaluable_global_reference) list -> Reductionops.reduction_function + val hnf_constr : Reductionops.reduction_function + val red_product : Reductionops.reduction_function + val is_evaluable : Environ.env -> Names.evaluable_global_reference -> bool + val evaluable_of_global_reference : + Environ.env -> Globnames.global_reference -> Names.evaluable_global_reference + val error_not_evaluable : Globnames.global_reference -> 'a + val reduce_to_quantified_ref : + Environ.env -> Evd.evar_map -> Globnames.global_reference -> EConstr.types -> EConstr.types + val pattern_occs : (Locus.occurrences * EConstr.constr) list -> Reductionops.e_reduction_function + val cbv_norm_flags : CClosure.RedFlags.reds -> Reductionops.reduction_function +end + +module Detyping : +sig + val print_universes : bool ref + val print_evar_arguments : bool ref + val detype : ?lax:bool -> bool -> Names.Id.t list -> Environ.env -> Evd.evar_map -> EConstr.constr -> Glob_term.glob_constr + val subst_glob_constr : Mod_subst.substitution -> Glob_term.glob_constr -> Glob_term.glob_constr + val set_detype_anonymous : (?loc:Loc.t -> int -> Glob_term.glob_constr) -> unit +end + +module Constrexpr_ops : +sig + val mkIdentC : Names.Id.t -> Constrexpr.constr_expr + val mkAppC : Constrexpr.constr_expr * Constrexpr.constr_expr list -> Constrexpr.constr_expr + val names_of_local_assums : Constrexpr.local_binder_expr list -> Names.Name.t Loc.located list + val coerce_reference_to_id : Prelude.reference -> Names.Id.t + val coerce_to_id : Constrexpr.constr_expr -> Names.Id.t Loc.located + val constr_loc : Constrexpr.constr_expr -> Loc.t option + val mkRefC : Prelude.reference -> Constrexpr.constr_expr + val mkLambdaC : Names.Name.t Loc.located list * Constrexpr.binder_kind * Constrexpr.constr_expr * Constrexpr.constr_expr -> Constrexpr.constr_expr + val default_binder_kind : Constrexpr.binder_kind + val mkLetInC : Names.Name.t Loc.located * Constrexpr.constr_expr * Constrexpr.constr_expr option * Constrexpr.constr_expr -> Constrexpr.constr_expr + val mkCProdN : ?loc:Loc.t -> Constrexpr.local_binder_expr list -> Constrexpr.constr_expr -> Constrexpr.constr_expr +end + +module Glob_ops : +sig + val map_glob_constr_left_to_right : (Glob_term.glob_constr -> Glob_term.glob_constr) -> Glob_term.glob_constr -> Glob_term.glob_constr + val loc_of_glob_constr : Glob_term.glob_constr -> Loc.t option + val glob_constr_eq : Glob_term.glob_constr -> Glob_term.glob_constr -> bool + val bound_glob_vars : Glob_term.glob_constr -> Names.Id.Set.t + + (** Conversion from glob_constr to cases pattern, if possible + + Take the current alias as parameter, + @raise Not_found if translation is impossible *) + val cases_pattern_of_glob_constr : Names.Name.t -> Glob_term.glob_constr -> Glob_term.cases_pattern + val map_glob_constr : + (Glob_term.glob_constr -> Glob_term.glob_constr) -> Glob_term.glob_constr -> Glob_term.glob_constr + val empty_lvar : Glob_term.ltac_var_map +end + +module Indrec : +sig + type dep_flag = bool + val lookup_eliminator : Names.inductive -> Sorts.family -> Globnames.global_reference + val build_case_analysis_scheme : Environ.env -> Evd.evar_map -> Term.pinductive -> + dep_flag -> Sorts.family -> Evd.evar_map * Term.constr + val make_elimination_ident : Names.Id.t -> Sorts.family -> Names.Id.t + val build_mutual_induction_scheme : + Environ.env -> Evd.evar_map -> (Term.pinductive * dep_flag * Sorts.family) list -> Evd.evar_map * Term.constr list + val build_case_analysis_scheme_default : Environ.env -> Evd.evar_map -> Term.pinductive -> + Sorts.family -> Evd.evar_map * Term.constr +end + +module Logic : +sig + type refiner_error = Logic.refiner_error = + | BadType of Term.constr * Term.constr * Term.constr + | UnresolvedBindings of Names.Name.t list + | CannotApply of Term.constr * Term.constr + | NotWellTyped of Term.constr + | NonLinearProof of Term.constr + | MetaInType of EConstr.constr + | IntroNeedsProduct + | DoesNotOccurIn of Term.constr * Names.Id.t + | NoSuchHyp of Names.Id.t + exception RefinerError of refiner_error + val catchable_exception : exn -> bool +end + +module Himsg : +sig + val explain_refiner_error : Logic.refiner_error -> Pp.std_ppcmds + val explain_pretype_error : Environ.env -> Evd.evar_map -> Pretype_errors.pretype_error -> Pp.std_ppcmds +end + +module Extend : +sig + type ('self, 'a) symbol = ('self, 'a) Extend.symbol + type 'a user_symbol = 'a Extend.user_symbol = + | Ulist1 of 'a user_symbol + | Ulist1sep of 'a user_symbol * string + | Ulist0 of 'a user_symbol + | Ulist0sep of 'a user_symbol * string + | Uopt of 'a user_symbol + | Uentry of 'a + | Uentryl of 'a * int +end + +module Pputils : +sig + val pr_with_occurrences : ('a -> Pp.std_ppcmds) -> (string -> Pp.std_ppcmds) -> 'a Locus.with_occurrences -> Pp.std_ppcmds + val pr_red_expr : + ('a -> Pp.std_ppcmds) * ('a -> Pp.std_ppcmds) * ('b -> Pp.std_ppcmds) * ('c -> Pp.std_ppcmds) -> + (string -> Pp.std_ppcmds) -> + ('a,'b,'c) Genredexpr.red_expr_gen -> Pp.std_ppcmds + val pr_raw_generic : Environ.env -> Genarg.rlevel Genarg.generic_argument -> Pp.std_ppcmds + val pr_glb_generic : Environ.env -> Genarg.glevel Genarg.generic_argument -> Pp.std_ppcmds + val pr_or_var : ('a -> Pp.std_ppcmds) -> 'a Misctypes.or_var -> Pp.std_ppcmds + val pr_or_by_notation : ('a -> Pp.std_ppcmds) -> 'a Misctypes.or_by_notation -> Pp.std_ppcmds +end + +module Ppconstr : +sig + val pr_name : Names.Name.t -> Pp.std_ppcmds + [@@ocaml.deprecated "alias of API.Names.Name.print"] + + val pr_id : Names.Id.t -> Pp.std_ppcmds + val pr_or_var : ('a -> Pp.std_ppcmds) -> 'a Misctypes.or_var -> Pp.std_ppcmds + val pr_with_comments : ?loc:Loc.t -> Pp.std_ppcmds -> Pp.std_ppcmds + val pr_lident : Names.Id.t Loc.located -> Pp.std_ppcmds + val pr_lname : Names.Name.t Loc.located -> Pp.std_ppcmds + val prec_less : int -> int * Ppextend.parenRelation -> bool + val pr_constr_expr : Constrexpr.constr_expr -> Pp.std_ppcmds + val pr_lconstr_expr : Constrexpr.constr_expr -> Pp.std_ppcmds + val pr_constr_pattern_expr : Constrexpr.constr_pattern_expr -> Pp.std_ppcmds + val pr_lconstr_pattern_expr : Constrexpr.constr_pattern_expr -> Pp.std_ppcmds + val pr_binders : Constrexpr.local_binder_expr list -> Pp.std_ppcmds + val pr_glob_sort : Misctypes.glob_sort -> Pp.std_ppcmds +end + +module Genprint : +sig + type 'a printer = 'a -> Pp.std_ppcmds + val generic_top_print : Genarg.tlevel Genarg.generic_argument printer + val register_print0 : ('raw, 'glb, 'top) Genarg.genarg_type -> + 'raw printer -> 'glb printer -> 'top printer -> unit +end + +module Miscprint : +sig + val pr_or_and_intro_pattern : + ('a -> Pp.std_ppcmds) -> 'a Misctypes.or_and_intro_pattern_expr -> Pp.std_ppcmds + val pr_intro_pattern_naming : Misctypes.intro_pattern_naming_expr -> Pp.std_ppcmds + val pr_intro_pattern : + ('a -> Pp.std_ppcmds) -> 'a Misctypes.intro_pattern_expr Loc.located -> Pp.std_ppcmds + val pr_bindings : + ('a -> Pp.std_ppcmds) -> + ('a -> Pp.std_ppcmds) -> 'a Misctypes.bindings -> Pp.std_ppcmds + val pr_bindings_no_with : + ('a -> Pp.std_ppcmds) -> + ('a -> Pp.std_ppcmds) -> 'a Misctypes.bindings -> Pp.std_ppcmds + val pr_with_bindings : + ('a -> Pp.std_ppcmds) -> + ('a -> Pp.std_ppcmds) -> 'a * 'a Misctypes.bindings -> Pp.std_ppcmds +end + +module Miscops : +sig + val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) -> + ('a,'b,'c) Genredexpr.red_expr_gen -> ('d,'e,'f) Genredexpr.red_expr_gen + val map_cast_type : ('a -> 'b) -> 'a Misctypes.cast_type -> 'b Misctypes.cast_type +end + +module Stateid : +sig + type t = Stateid.t + module Self : module type of struct include Stateid.Self end +end + +module Stm : +sig + type state = Stm.state + val state_of_id : + Stateid.t -> [ `Valid of state option | `Expired | `Error of exn ] +end + +module Declaremods : +sig + val append_end_library_hook : (unit -> unit) -> unit +end + +module Pfedit : +sig + val solve_by_implicit_tactic : unit -> Pretyping.inference_hook option + val refine_by_tactic : Environ.env -> Evd.evar_map -> EConstr.types -> unit Proofview.tactic -> + Term.constr * Evd.evar_map + val declare_implicit_tactic : unit Proofview.tactic -> unit + val clear_implicit_tactic : unit -> unit + val by : unit Proofview.tactic -> bool + val solve : ?with_end_tac:unit Proofview.tactic -> + Vernacexpr.goal_selector -> int option -> unit Proofview.tactic -> + Proof.proof -> Proof.proof * bool + val cook_proof : + unit -> (Names.Id.t * (Safe_typing.private_constants Entries.definition_entry * Proof_global.proof_universes * Decl_kinds.goal_kind)) + + val get_current_context : unit -> Evd.evar_map * Environ.env + + (* Deprecated *) + val delete_current_proof : unit -> unit + [@@ocaml.deprecated "use Proof_global.discard_current"] + + val get_current_proof_name : unit -> Names.Id.t + [@@ocaml.deprecated "use Proof_global.get_current_proof_name"] + +end + +module Tactics : +sig + open Proofview + + type change_arg = Pattern.patvar_map -> Evd.evar_map -> Evd.evar_map * EConstr.constr + type tactic_reduction = Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr + type elim_scheme = Tactics.elim_scheme = + { + elimc: EConstr.constr Misctypes.with_bindings option; + elimt: EConstr.types; + indref: Globnames.global_reference option; + params: EConstr.rel_context; + nparams: int; + predicates: EConstr.rel_context; + npredicates: int; + branches: EConstr.rel_context; + nbranches: int; + args: EConstr.rel_context; + nargs: int; + indarg: EConstr.rel_declaration option; + concl: EConstr.types; + indarg_in_concl: bool; + farg_in_concl: bool; + } + + val unify : ?state:Names.transparent_state -> EConstr.constr -> EConstr.constr -> unit Proofview.tactic + val intro_then : (Names.Id.t -> unit Proofview.tactic) -> unit Proofview.tactic + val reflexivity : unit tactic + val change_concl : EConstr.constr -> unit tactic + val apply : EConstr.constr -> unit tactic + val normalise_vm_in_concl : unit tactic + val assert_before : Names.Name.t -> EConstr.types -> unit tactic + val exact_check : EConstr.constr -> unit tactic + val simplest_elim : EConstr.constr -> unit tactic + val introf : unit tactic + val cut : EConstr.types -> unit tactic + val convert_concl : ?check:bool -> EConstr.types -> Term.cast_kind -> unit tactic + val intro_using : Names.Id.t -> unit tactic + val intro : unit tactic + val fresh_id_in_env : Names.Id.t list -> Names.Id.t -> Environ.env -> Names.Id.t + val is_quantified_hypothesis : Names.Id.t -> 'a Goal.t -> bool + val tclABSTRACT : ?opaque:bool -> Names.Id.t option -> unit Proofview.tactic -> unit Proofview.tactic + val intro_patterns : bool -> Tactypes.intro_patterns -> unit Proofview.tactic + val apply_with_delayed_bindings_gen : + Misctypes.advanced_flag -> Misctypes.evars_flag -> (Misctypes.clear_flag * Tactypes.delayed_open_constr_with_bindings Loc.located) list -> unit Proofview.tactic + val apply_delayed_in : + Misctypes.advanced_flag -> Misctypes.evars_flag -> Names.Id.t -> + (Misctypes.clear_flag * Tactypes.delayed_open_constr_with_bindings Loc.located) list -> + Tactypes.intro_pattern option -> unit Proofview.tactic + val elim : + Misctypes.evars_flag -> Misctypes.clear_flag -> EConstr.constr Misctypes.with_bindings -> EConstr.constr Misctypes.with_bindings option -> unit Proofview.tactic + val general_case_analysis : Misctypes.evars_flag -> Misctypes.clear_flag -> EConstr.constr Misctypes.with_bindings -> unit Proofview.tactic + val mutual_fix : + Names.Id.t -> int -> (Names.Id.t * int * EConstr.constr) list -> int -> unit Proofview.tactic + val mutual_cofix : Names.Id.t -> (Names.Id.t * EConstr.constr) list -> int -> unit Proofview.tactic + val forward : bool -> unit Proofview.tactic option option -> + Tactypes.intro_pattern option -> EConstr.constr -> unit Proofview.tactic + val generalize_gen : (EConstr.constr Locus.with_occurrences * Names.Name.t) list -> unit Proofview.tactic + val letin_tac : (bool * Tactypes.intro_pattern_naming) option -> + Names.Name.t -> EConstr.constr -> EConstr.types option -> Locus.clause -> unit Proofview.tactic + val letin_pat_tac : Misctypes.evars_flag -> + (bool * Tactypes.intro_pattern_naming) option -> + Names.Name.t -> + Evd.evar_map * EConstr.constr -> + Locus.clause -> unit Proofview.tactic + val induction_destruct : Misctypes.rec_flag -> Misctypes.evars_flag -> + (Tactypes.delayed_open_constr_with_bindings Misctypes.destruction_arg + * (Tactypes.intro_pattern_naming option * Tactypes.or_and_intro_pattern option) + * Locus.clause option) list * + EConstr.constr Misctypes.with_bindings option -> unit Proofview.tactic + val reduce : Redexpr.red_expr -> Locus.clause -> unit Proofview.tactic + val change : Pattern.constr_pattern option -> change_arg -> Locus.clause -> unit Proofview.tactic + val intros_reflexivity : unit Proofview.tactic + val exact_no_check : EConstr.constr -> unit Proofview.tactic + val assumption : unit Proofview.tactic + val intros_transitivity : EConstr.constr option -> unit Proofview.tactic + val vm_cast_no_check : EConstr.constr -> unit Proofview.tactic + val native_cast_no_check : EConstr.constr -> unit Proofview.tactic + val case_type : EConstr.types -> unit Proofview.tactic + val elim_type : EConstr.types -> unit Proofview.tactic + val cut_and_apply : EConstr.constr -> unit Proofview.tactic + val left_with_bindings : Misctypes.evars_flag -> EConstr.constr Misctypes.bindings -> unit Proofview.tactic + val right_with_bindings : Misctypes.evars_flag -> EConstr.constr Misctypes.bindings -> unit Proofview.tactic + val any_constructor : Misctypes.evars_flag -> unit Proofview.tactic option -> unit Proofview.tactic + val constructor_tac : Misctypes.evars_flag -> int option -> int -> + EConstr.constr Misctypes.bindings -> unit Proofview.tactic + val specialize : EConstr.constr Misctypes.with_bindings -> Tactypes.intro_pattern option -> unit Proofview.tactic + val intros_symmetry : Locus.clause -> unit Proofview.tactic + val split_with_bindings : Misctypes.evars_flag -> EConstr.constr Misctypes.bindings list -> unit Proofview.tactic + val intros_until : Misctypes.quantified_hypothesis -> unit Proofview.tactic + val intro_move : Names.Id.t option -> Names.Id.t Misctypes.move_location -> unit Proofview.tactic + val move_hyp : Names.Id.t -> Names.Id.t Misctypes.move_location -> unit Proofview.tactic + val rename_hyp : (Names.Id.t * Names.Id.t) list -> unit Proofview.tactic + val revert : Names.Id.t list -> unit Proofview.tactic + val simple_induct : Misctypes.quantified_hypothesis -> unit Proofview.tactic + val simple_destruct : Misctypes.quantified_hypothesis -> unit Proofview.tactic + val fix : Names.Id.t option -> int -> unit Proofview.tactic + val cofix : Names.Id.t option -> unit Proofview.tactic + val keep : Names.Id.t list -> unit Proofview.tactic + val clear : Names.Id.t list -> unit Proofview.tactic + val clear_body : Names.Id.t list -> unit Proofview.tactic + val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> EConstr.constr -> unit Proofview.tactic + val force_destruction_arg : Misctypes.evars_flag -> Environ.env -> Evd.evar_map -> + Tactypes.delayed_open_constr_with_bindings Misctypes.destruction_arg -> + Evd.evar_map * EConstr.constr Misctypes.with_bindings Misctypes.destruction_arg + val apply_with_bindings : EConstr.constr Misctypes.with_bindings -> unit Proofview.tactic + val abstract_generalize : ?generalize_vars:bool -> ?force_dep:bool -> Names.Id.t -> unit Proofview.tactic + val specialize_eqs : Names.Id.t -> unit Proofview.tactic + val generalize : EConstr.constr list -> unit Proofview.tactic + val simplest_case : EConstr.constr -> unit Proofview.tactic + val introduction : ?check:bool -> Names.Id.t -> unit Proofview.tactic + val convert_concl_no_check : EConstr.types -> Term.cast_kind -> unit Proofview.tactic + val reduct_in_concl : tactic_reduction * Term.cast_kind -> unit Proofview.tactic + val reduct_in_hyp : ?check:bool -> tactic_reduction -> Locus.hyp_location -> unit Proofview.tactic + val convert_hyp_no_check : EConstr.named_declaration -> unit Proofview.tactic + val reflexivity_red : bool -> unit Proofview.tactic + val symmetry_red : bool -> unit Proofview.tactic + val eapply : EConstr.constr -> unit Proofview.tactic + val transitivity_red : bool -> EConstr.constr option -> unit Proofview.tactic + val assert_after_replacing : Names.Id.t -> EConstr.types -> unit Proofview.tactic + val intros : unit Proofview.tactic + val setoid_reflexivity : unit Proofview.tactic Hook.t + val setoid_symmetry : unit Proofview.tactic Hook.t + val setoid_symmetry_in : (Names.Id.t -> unit Proofview.tactic) Hook.t + val setoid_transitivity : (EConstr.constr option -> unit Proofview.tactic) Hook.t + val unfold_in_concl : + (Locus.occurrences * Names.evaluable_global_reference) list -> unit Proofview.tactic + val intros_using : Names.Id.t list -> unit Proofview.tactic + val simpl_in_concl : unit Proofview.tactic + val reduct_option : ?check:bool -> tactic_reduction * Term.cast_kind -> Locus.goal_location -> unit Proofview.tactic + val simplest_split : unit Proofview.tactic + val unfold_in_hyp : + (Locus.occurrences * Names.evaluable_global_reference) list -> Locus.hyp_location -> unit Proofview.tactic + val split : EConstr.constr Misctypes.bindings -> unit Proofview.tactic + val red_in_concl : unit Proofview.tactic + val change_in_concl : (Locus.occurrences * Pattern.constr_pattern) option -> change_arg -> unit Proofview.tactic + val eapply_with_bindings : EConstr.constr Misctypes.with_bindings -> unit Proofview.tactic + val assert_by : Names.Name.t -> EConstr.types -> unit Proofview.tactic -> + unit Proofview.tactic + val intro_avoiding : Names.Id.t list -> unit Proofview.tactic + val pose_proof : Names.Name.t -> EConstr.constr -> unit Proofview.tactic + val pattern_option : (Locus.occurrences * EConstr.constr) list -> Locus.goal_location -> unit Proofview.tactic + val compute_elim_sig : Evd.evar_map -> ?elimc:EConstr.constr Misctypes.with_bindings -> EConstr.types -> elim_scheme + val try_intros_until : + (Names.Id.t -> unit Proofview.tactic) -> Misctypes.quantified_hypothesis -> unit Proofview.tactic + val cache_term_by_tactic_then : + opaque:bool -> ?goal_type:(EConstr.constr option) -> Names.Id.t -> + Decl_kinds.goal_kind -> unit Proofview.tactic -> (EConstr.constr -> EConstr.constr list -> unit Proofview.tactic) -> unit Proofview.tactic + val apply_type : EConstr.constr -> EConstr.constr list -> unit Proofview.tactic + val hnf_in_concl : unit Proofview.tactic + val intro_mustbe_force : Names.Id.t -> unit Proofview.tactic + + module New : + sig + val refine : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.constr) -> unit Proofview.tactic + val reduce_after_refine : unit Proofview.tactic + end + module Simple : + sig + val intro : Names.Id.t -> unit Proofview.tactic + val apply : EConstr.constr -> unit Proofview.tactic + val case : EConstr.constr -> unit Proofview.tactic + end +end + +module Tacticals : +sig + open Proof_type + val tclORELSE : tactic -> tactic -> tactic + val tclDO : int -> tactic -> tactic + val tclIDTAC : tactic + val tclFAIL : int -> Pp.std_ppcmds -> tactic + val tclTHEN : tactic -> tactic -> tactic + val tclTHENLIST : tactic list -> tactic + val pf_constr_of_global : + Globnames.global_reference -> (EConstr.constr -> Proof_type.tactic) -> Proof_type.tactic + val tclMAP : ('a -> tactic) -> 'a list -> tactic + val tclTRY : tactic -> tactic + val tclCOMPLETE : tactic -> tactic + val tclTHENS : tactic -> tactic list -> tactic + val tclFIRST : tactic list -> tactic + val tclTHENFIRST : tactic -> tactic -> tactic + val tclTHENLAST : tactic -> tactic -> tactic + val tclTHENSFIRSTn : tactic -> tactic array -> tactic -> tactic + val tclTHENSLASTn : tactic -> tactic -> tactic array -> tactic + val tclSOLVE : tactic list -> tactic + + val onClause : (Names.Id.t option -> tactic) -> Locus.clause -> tactic + val onAllHypsAndConcl : (Names.Id.t option -> tactic) -> tactic + val onLastHypId : (Names.Id.t -> tactic) -> tactic + val onNthHypId : int -> (Names.Id.t -> tactic) -> tactic + val onNLastHypsId : int -> (Names.Id.t list -> tactic) -> tactic + + val tclTHENSEQ : tactic list -> tactic + [@@ocaml.deprecated "alias of API.Tacticals.tclTHENLIST"] + + val nLastDecls : int -> Proof_type.goal Evd.sigma -> EConstr.named_context + + val tclTHEN_i : tactic -> (int -> tactic) -> tactic + + val tclPROGRESS : tactic -> tactic + + val elimination_sort_of_goal : Proof_type.goal Evd.sigma -> Sorts.family + + module New : + sig + open Proofview + val tclORELSE0 : unit tactic -> unit tactic -> unit tactic + val tclFAIL : int -> Pp.std_ppcmds -> 'a tactic + val pf_constr_of_global : Globnames.global_reference -> EConstr.constr tactic + val tclTHEN : unit tactic -> unit tactic -> unit tactic + val tclTHENS : unit tactic -> unit tactic list -> unit tactic + val tclFIRST : unit tactic list -> unit tactic + val tclZEROMSG : ?loc:Loc.t -> Pp.std_ppcmds -> 'a tactic + val tclORELSE : unit tactic -> unit tactic -> unit tactic + val tclREPEAT : unit tactic -> unit tactic + val tclTRY : unit tactic -> unit tactic + val tclTHENFIRST : unit tactic -> unit tactic -> unit tactic + val tclPROGRESS : unit Proofview.tactic -> unit Proofview.tactic + val tclTHENS3PARTS : unit tactic -> unit tactic array -> unit tactic -> unit tactic array -> unit tactic + val tclDO : int -> unit tactic -> unit tactic + val tclTIMEOUT : int -> unit tactic -> unit tactic + val tclTIME : string option -> 'a tactic -> 'a tactic + val tclOR : unit tactic -> unit tactic -> unit tactic + val tclONCE : unit tactic -> unit tactic + val tclEXACTLY_ONCE : unit tactic -> unit tactic + val tclIFCATCH : + unit tactic -> + (unit -> unit tactic) -> + (unit -> unit tactic) -> unit tactic + val tclSOLVE : unit tactic list -> unit tactic + val tclCOMPLETE : 'a tactic -> 'a tactic + val tclSELECT : Vernacexpr.goal_selector -> 'a tactic -> 'a tactic + val tclWITHHOLES : bool -> 'a tactic -> Evd.evar_map -> 'a tactic + val tclDELAYEDWITHHOLES : bool -> 'a Tactypes.delayed_open -> ('a -> unit tactic) -> unit tactic + val tclTHENLIST : unit tactic list -> unit tactic + val tclTHENLAST : unit tactic -> unit tactic -> unit tactic + val tclMAP : ('a -> unit tactic) -> 'a list -> unit tactic + val tclIDTAC : unit tactic + val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic + val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic + end +end + +module Equality : +sig + type orientation = bool + type freeze_evars_flag = bool + type dep_proof_flag = bool + type conditions = + | Naive + | FirstSolved + | AllMatches + + val build_selector : + Environ.env -> Evd.evar_map -> int -> EConstr.constr -> EConstr.types -> + EConstr.constr -> EConstr.constr -> Evd.evar_map * EConstr.constr + val replace : EConstr.constr -> EConstr.constr -> unit Proofview.tactic + val general_rewrite : + orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag -> + ?tac:(unit Proofview.tactic * conditions) -> EConstr.constr -> unit Proofview.tactic + val inj : Tactypes.intro_patterns option -> Misctypes.evars_flag -> + Misctypes.clear_flag -> EConstr.constr Misctypes.with_bindings -> unit Proofview.tactic + val general_multi_rewrite : + Misctypes.evars_flag -> (bool * Misctypes.multi * Misctypes.clear_flag * Tactypes.delayed_open_constr_with_bindings) list -> + Locus.clause -> (unit Proofview.tactic * conditions) option -> unit Proofview.tactic + val replace_in_clause_maybe_by : EConstr.constr -> EConstr.constr -> Locus.clause -> unit Proofview.tactic option -> unit Proofview.tactic + val replace_term : bool option -> EConstr.constr -> Locus.clause -> unit Proofview.tactic + val dEq : Misctypes.evars_flag -> EConstr.constr Misctypes.with_bindings Misctypes.destruction_arg option -> unit Proofview.tactic + val discr_tac : Misctypes.evars_flag -> + EConstr.constr Misctypes.with_bindings Misctypes.destruction_arg option -> unit Proofview.tactic + val injClause : Tactypes.intro_patterns option -> Misctypes.evars_flag -> + EConstr.constr Misctypes.with_bindings Misctypes.destruction_arg option -> unit Proofview.tactic + + val simpleInjClause : Misctypes.evars_flag -> + EConstr.constr Misctypes.with_bindings Misctypes.destruction_arg option -> + unit Proofview.tactic + val rewriteInConcl : bool -> EConstr.constr -> unit Proofview.tactic + val rewriteInHyp : bool -> EConstr.constr -> Names.Id.t -> unit Proofview.tactic + val cutRewriteInConcl : bool -> EConstr.constr -> unit Proofview.tactic + val cutRewriteInHyp : bool -> EConstr.types -> Names.Id.t -> unit Proofview.tactic + val general_rewrite_ebindings_clause : Names.Id.t option -> + orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag -> + ?tac:(unit Proofview.tactic * conditions) -> EConstr.constr Misctypes.with_bindings -> Misctypes.evars_flag -> unit Proofview.tactic + val subst : Names.Id.t list -> unit Proofview.tactic + type subst_tactic_flags = Equality.subst_tactic_flags = { + only_leibniz : bool; + rewrite_dependent_proof : bool + } + val subst_all : ?flags:subst_tactic_flags -> unit -> unit Proofview.tactic + + val general_rewrite_in : + orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag -> + ?tac:(unit Proofview.tactic * conditions) -> Names.Id.t -> EConstr.constr -> Misctypes.evars_flag -> unit Proofview.tactic + + val general_setoid_rewrite_clause : + (Names.Id.t option -> orientation -> Locus.occurrences -> EConstr.constr Misctypes.with_bindings -> + new_goals:EConstr.constr list -> unit Proofview.tactic) Hook.t + + val discrConcl : unit Proofview.tactic + val rewriteLR : ?tac:(unit Proofview.tactic * conditions) -> EConstr.constr -> unit Proofview.tactic + val rewriteRL : ?tac:(unit Proofview.tactic * conditions) -> EConstr.constr -> unit Proofview.tactic + val general_rewrite_bindings : + orientation -> Locus.occurrences -> freeze_evars_flag -> dep_proof_flag -> + ?tac:(unit Proofview.tactic * conditions) -> EConstr.constr Misctypes.with_bindings -> Misctypes.evars_flag -> unit Proofview.tactic + val discriminable : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool + val discrHyp : Names.Id.t -> unit Proofview.tactic + val injectable : Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool + val injHyp : Misctypes.clear_flag -> Names.Id.t -> unit Proofview.tactic + val subst_gen : bool -> Names.Id.t list -> unit Proofview.tactic +end + +module Contradiction : +sig + val contradiction : EConstr.constr Misctypes.with_bindings option -> unit Proofview.tactic + val absurd : EConstr.constr -> unit Proofview.tactic +end + +module Clenv : +sig + type hole = Clenv.hole = { + hole_evar : EConstr.constr; + hole_type : EConstr.types; + hole_deps : bool; + hole_name : Names.Name.t; + } + type clause = Clenv.clause = { + cl_holes : hole list; + cl_concl : EConstr.types; + } + val make_evar_clause : Environ.env -> Evd.evar_map -> ?len:int -> EConstr.types -> + (Evd.evar_map * clause) + val solve_evar_clause : Environ.env -> Evd.evar_map -> bool -> clause -> EConstr.constr Misctypes.bindings -> + Evd.evar_map + type clausenv = Clenv.clausenv + val pr_clenv : Clenv.clausenv -> Pp.std_ppcmds +end + +module Hints : +sig + type hint = Hints.hint + type debug = Hints.debug = + | Debug | Info | Off + type 'a hints_path_atom_gen = 'a Hints.hints_path_atom_gen = + | PathHints of 'a list + | PathAny + type hint_term = Hints.hint_term = + | IsGlobRef of Globnames.global_reference + | IsConstr of EConstr.constr * Univ.ContextSet.t + type hint_db_name = string + type hint_info = (Names.Id.t list * Pattern.constr_pattern) Vernacexpr.hint_info_gen + type hnf = bool + type hints_path_atom = Globnames.global_reference hints_path_atom_gen + + type 'a hints_path_gen = 'a Hints.hints_path_gen = + | PathAtom of 'a hints_path_atom_gen + | PathStar of 'a hints_path_gen + | PathSeq of 'a hints_path_gen * 'a hints_path_gen + | PathOr of 'a hints_path_gen * 'a hints_path_gen + | PathEmpty + | PathEpsilon + + type hints_path = Globnames.global_reference hints_path_gen + + type hints_entry = Hints.hints_entry = + | HintsResolveEntry of (hint_info * Decl_kinds.polymorphic * hnf * hints_path_atom * hint_term) list + | HintsImmediateEntry of (hints_path_atom * Decl_kinds.polymorphic * hint_term) list + | HintsCutEntry of hints_path + | HintsUnfoldEntry of Names.evaluable_global_reference list + | HintsTransparencyEntry of Names.evaluable_global_reference list * bool + | HintsModeEntry of Globnames.global_reference * Vernacexpr.hint_mode list + | HintsExternEntry of hint_info * Genarg.glob_generic_argument + + type 'a hint_ast = 'a Hints.hint_ast = + | Res_pf of 'a + | ERes_pf of 'a + | Give_exact of 'a + | Res_pf_THEN_trivial_fail of 'a + | Unfold_nth of Names.evaluable_global_reference + | Extern of Genarg.glob_generic_argument + type raw_hint = EConstr.constr * EConstr.types * Univ.ContextSet.t + type 'a with_metadata = 'a Hints.with_metadata = private { + pri : int; + poly : Decl_kinds.polymorphic; + pat : Pattern.constr_pattern option; + name : hints_path_atom; + db : string option; + secvars : Names.Id.Pred.t; + code : 'a; + } + type full_hint = hint with_metadata + + module Hint_db : + sig + type t = Hints.Hint_db.t + val empty : ?name:hint_db_name -> Names.transparent_state -> bool -> t + val transparent_state : t -> Names.transparent_state + val iter : (Globnames.global_reference option -> + Vernacexpr.hint_mode array list -> full_hint list -> unit) -> t -> unit + end + type hint_db = Hint_db.t + + val add_hints : Vernacexpr.locality_flag -> hint_db_name list -> hints_entry -> unit + val searchtable_map : hint_db_name -> hint_db + val pp_hints_path_atom : ('a -> Pp.std_ppcmds) -> 'a hints_path_atom_gen -> Pp.std_ppcmds + val pp_hints_path_gen : ('a -> Pp.std_ppcmds) -> 'a hints_path_gen -> Pp.std_ppcmds + val glob_hints_path_atom : + Prelude.reference hints_path_atom_gen -> Globnames.global_reference hints_path_atom_gen + val pp_hints_path : hints_path -> Pp.std_ppcmds + val glob_hints_path : + Prelude.reference hints_path_gen -> Globnames.global_reference hints_path_gen + val typeclasses_db : hint_db_name + val add_hints_init : (unit -> unit) -> unit + val create_hint_db : bool -> hint_db_name -> Names.transparent_state -> bool -> unit + val empty_hint_info : 'a Vernacexpr.hint_info_gen + val repr_hint : hint -> (raw_hint * Clenv.clausenv) hint_ast + val pr_hint_db : Hint_db.t -> Pp.std_ppcmds +end + +module Auto : +sig + val default_auto : unit Proofview.tactic + val full_trivial : ?debug:Hints.debug -> + Tactypes.delayed_open_constr list -> unit Proofview.tactic + val h_auto : ?debug:Hints.debug -> + int option -> Tactypes.delayed_open_constr list -> Hints.hint_db_name list option -> unit Proofview.tactic + val h_trivial : ?debug:Hints.debug -> + Tactypes.delayed_open_constr list -> Hints.hint_db_name list option -> unit Proofview.tactic + val new_full_auto : ?debug:Hints.debug -> + int -> Tactypes.delayed_open_constr list -> unit Proofview.tactic + val full_auto : ?debug:Hints.debug -> + int -> Tactypes.delayed_open_constr list -> unit Proofview.tactic + val new_auto : ?debug:Hints.debug -> + int -> Tactypes.delayed_open_constr list -> Hints.hint_db_name list -> unit Proofview.tactic + val default_full_auto : unit Proofview.tactic +end + +module Hipattern : +sig + exception NoEquationFound + type 'a matching_function = Evd.evar_map -> EConstr.constr -> 'a option + type testing_function = Evd.evar_map -> EConstr.constr -> bool + val is_disjunction : ?strict:bool -> ?onlybinary:bool -> testing_function + val match_with_disjunction : ?strict:bool -> ?onlybinary:bool -> (EConstr.constr * EConstr.constr list) matching_function + val match_with_equality_type : (EConstr.constr * EConstr.constr list) matching_function + val is_empty_type : testing_function + val is_unit_type : testing_function + val is_unit_or_eq_type : testing_function + val is_conjunction : ?strict:bool -> ?onlybinary:bool -> testing_function + val match_with_conjunction : ?strict:bool -> ?onlybinary:bool -> (EConstr.constr * EConstr.constr list) matching_function + val match_with_imp_term : (EConstr.constr * EConstr.constr) matching_function + val match_with_forall_term : (Names.Name.t * EConstr.constr * EConstr.constr) matching_function + val match_with_nodep_ind : (EConstr.constr * EConstr.constr list * int) matching_function + val match_with_sigma_type : (EConstr.constr * EConstr.constr list) matching_function +end + +module Inv : +sig + val dinv : + Misctypes.inversion_kind -> EConstr.constr option -> + Tactypes.or_and_intro_pattern option -> Misctypes.quantified_hypothesis -> unit Proofview.tactic + val inv_clause : + Misctypes.inversion_kind -> Tactypes.or_and_intro_pattern option -> Names.Id.t list -> + Misctypes.quantified_hypothesis -> unit Proofview.tactic + val inv_clear_tac : Names.Id.t -> unit Proofview.tactic + val inv_tac : Names.Id.t -> unit Proofview.tactic + val dinv_tac : Names.Id.t -> unit Proofview.tactic + val dinv_clear_tac : Names.Id.t -> unit Proofview.tactic + val inv : Misctypes.inversion_kind -> Tactypes.or_and_intro_pattern option -> + Misctypes.quantified_hypothesis -> unit Proofview.tactic +end + +module Leminv : +sig + val lemInv_clause : + Misctypes.quantified_hypothesis -> EConstr.constr -> Names.Id.t list -> unit Proofview.tactic + val add_inversion_lemma_exn : + Names.Id.t -> Constrexpr.constr_expr -> Misctypes.glob_sort -> bool -> (Names.Id.t -> unit Proofview.tactic) -> + unit +end + +module Vernacentries : +sig + val dump_global : Prelude.reference Misctypes.or_by_notation -> unit + val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr -> + Evd.evar_map * Redexpr.red_expr) Hook.t + val command_focus : unit Proof.focus_kind +end + +module Evar_refiner : +sig + val w_refine : Evar.t * Evd.evar_info -> + Pretyping.glob_constr_ltac_closure -> Evd.evar_map -> Evd.evar_map +end + +module Obligations : +sig + val default_tactic : unit Proofview.tactic ref + val obligation : int * Names.Id.t option * Constrexpr.constr_expr option -> + Genarg.glob_generic_argument option -> unit + val next_obligation : Names.Id.t option -> Genarg.glob_generic_argument option -> unit + val try_solve_obligation : int -> Names.Id.t option -> unit Proofview.tactic option -> unit + val try_solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> unit + val solve_all_obligations : unit Proofview.tactic option -> unit + val admit_obligations : Names.Id.t option -> unit + val show_obligations : ?msg:bool -> Names.Id.t option -> unit + val show_term : Names.Id.t option -> Pp.std_ppcmds +end + +module Elim : +sig + val h_decompose : Names.inductive list -> EConstr.constr -> unit Proofview.tactic + val h_double_induction : Misctypes.quantified_hypothesis -> Misctypes.quantified_hypothesis-> unit Proofview.tactic + val h_decompose_or : EConstr.constr -> unit Proofview.tactic + val h_decompose_and : EConstr.constr -> unit Proofview.tactic +end + +module Redops : +sig + val all_flags : 'a Genredexpr.glob_red_flag + val make_red_flag : 'a Genredexpr.red_atom list -> 'a Genredexpr.glob_red_flag +end + +module Autorewrite : +sig + type rew_rule = { rew_lemma: Term.constr; + rew_type: Term.types; + rew_pat: Term.constr; + rew_ctx: Univ.ContextSet.t; + rew_l2r: bool; + rew_tac: Genarg.glob_generic_argument option } + type raw_rew_rule = (Term.constr Univ.in_universe_context_set * bool * + Genarg.raw_generic_argument option) + Loc.located + val auto_multi_rewrite : ?conds:Equality.conditions -> string list -> Locus.clause -> unit Proofview.tactic + val auto_multi_rewrite_with : ?conds:Equality.conditions -> unit Proofview.tactic -> string list -> Locus.clause -> unit Proofview.tactic + val add_rew_rules : string -> raw_rew_rule list -> unit + val find_rewrites : string -> rew_rule list + val find_matches : string -> Term.constr -> rew_rule list + val print_rewrite_hintdb : string -> Pp.std_ppcmds +end + +module Refine : +sig + val refine : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit Proofview.tactic + val solve_constraints : unit Proofview.tactic +end + +module Find_subterm : +sig + val error_invalid_occurrence : int list -> 'a +end + +module Vernac_classifier : +sig + val declare_vernac_classifier : + Vernacexpr.extend_name -> (Genarg.raw_generic_argument list -> unit -> Vernacexpr.vernac_classification) -> unit + val classify_as_proofstep : Vernacexpr.vernac_classification + val classify_as_query : Vernacexpr.vernac_classification + val classify_as_sideeff : Vernacexpr.vernac_classification + val classify_vernac : Vernacexpr.vernac_expr -> Vernacexpr.vernac_classification +end + +module Keys : +sig + type key = Keys.key + val constr_key : ('a -> ('a, 't, 'u, 'i) Term.kind_of_term) -> 'a -> key option + val declare_equiv_keys : key -> key -> unit + val pr_keys : (Globnames.global_reference -> Pp.std_ppcmds) -> Pp.std_ppcmds +end + +module Eauto : +sig + val e_assumption : unit Proofview.tactic + val e_give_exact : ?flags:Unification.unify_flags -> EConstr.constr -> unit Proofview.tactic + val prolog_tac : Tactypes.delayed_open_constr list -> int -> unit Proofview.tactic + val make_dimension : int option -> int option -> bool * int + val gen_eauto : ?debug:Hints.debug -> bool * int -> Tactypes.delayed_open_constr list -> + Hints.hint_db_name list option -> unit Proofview.tactic + val autounfold_tac : Hints.hint_db_name list option -> Locus.clause -> unit Proofview.tactic + val autounfold_one : Hints.hint_db_name list -> Locus.hyp_location option -> unit Proofview.tactic + val eauto_with_bases : + ?debug:Hints.debug -> bool * int -> Tactypes.delayed_open_constr list -> Hints.hint_db list -> Proof_type.tactic +end + +module Class_tactics : +sig + type search_strategy = Class_tactics.search_strategy = + | Dfs + | Bfs + val set_typeclasses_debug : bool -> unit + val set_typeclasses_strategy : search_strategy -> unit + val set_typeclasses_depth : int option -> unit + val typeclasses_eauto : ?only_classes:bool -> ?st:Names.transparent_state -> ?strategy:search_strategy -> + depth:(Int.t option) -> + Hints.hint_db_name list -> unit Proofview.tactic + val head_of_constr : Names.Id.t -> EConstr.constr -> unit Proofview.tactic + val not_evar : EConstr.constr -> unit Proofview.tactic + val is_ground : EConstr.constr -> unit Proofview.tactic + val autoapply : EConstr.constr -> Hints.hint_db_name -> unit Proofview.tactic + val catchable : exn -> bool +end + +module Ind_tables : +sig + type individual = Ind_tables.individual + type 'a scheme_kind = 'a Ind_tables.scheme_kind + + val check_scheme : 'a scheme_kind -> Names.inductive -> bool + val find_scheme : ?mode:Declare.internal_flag -> 'a scheme_kind -> Names.inductive -> Names.Constant.t * Safe_typing.private_constants + val pr_scheme_kind : 'a scheme_kind -> Pp.std_ppcmds +end + +module Elimschemes : +sig + val case_scheme_kind_from_prop : Ind_tables.individual Ind_tables.scheme_kind + val case_dep_scheme_kind_from_type_in_prop : Ind_tables.individual Ind_tables.scheme_kind + val case_scheme_kind_from_type : Ind_tables.individual Ind_tables.scheme_kind + val case_dep_scheme_kind_from_type : Ind_tables.individual Ind_tables.scheme_kind + val case_dep_scheme_kind_from_prop : Ind_tables.individual Ind_tables.scheme_kind +end + +module Lemmas : +sig + type 'a declaration_hook = 'a Lemmas.declaration_hook + val mk_hook : + (Decl_kinds.locality -> Globnames.global_reference -> 'a) -> 'a declaration_hook + val start_proof : Names.Id.t -> ?pl:Proof_global.universe_binders -> Decl_kinds.goal_kind -> Evd.evar_map -> + ?terminator:(Proof_global.lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) -> + ?sign:Environ.named_context_val -> EConstr.types -> + ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards -> + unit declaration_hook -> unit + val call_hook : + Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a + val save_proof : ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit + val get_current_context : unit -> Evd.evar_map * Environ.env +end + +module Eqdecide : +sig + val compare : EConstr.constr -> EConstr.constr -> unit Proofview.tactic + val decideEqualityGoal : unit Proofview.tactic +end + +module Locusops : +sig + val clause_with_generic_occurrences : 'a Locus.clause_expr -> bool + val nowhere : 'a Locus.clause_expr + val allHypsAndConcl : 'a Locus.clause_expr + val is_nowhere : 'a Locus.clause_expr -> bool + val occurrences_map : + ('a list -> 'b list) -> 'a Locus.occurrences_gen -> 'b Locus.occurrences_gen + val convert_occs : Locus.occurrences -> bool * int list + val onConcl : 'a Locus.clause_expr + val onHyp : 'a -> 'a Locus.clause_expr +end + +module Topfmt : +sig + val std_ft : Format.formatter ref + val with_output_to : out_channel -> Format.formatter + val get_margin : unit -> int option +end + +module Nameops : +sig + val atompart_of_id : Names.Id.t -> string + + val pr_id : Names.Id.t -> Pp.std_ppcmds + [@@ocaml.deprecated "alias of API.Names.Id.print"] + + val pr_name : Names.Name.t -> Pp.std_ppcmds + [@@ocaml.deprecated "alias of API.Names.Name.print"] + + val name_fold : (Names.Id.t -> 'a -> 'a) -> Names.Name.t -> 'a -> 'a + val name_app : (Names.Id.t -> Names.Id.t) -> Names.Name.t -> Names.Name.t + val add_suffix : Names.Id.t -> string -> Names.Id.t + val increment_subscript : Names.Id.t -> Names.Id.t + val make_ident : string -> int option -> Names.Id.t + val out_name : Names.Name.t -> Names.Id.t + val pr_lab : Names.Label.t -> Pp.std_ppcmds + module Name : + sig + include module type of struct include Names.Name end + val get_id : t -> Names.Id.t + val fold_right : (Names.Id.t -> 'a -> 'a) -> t -> 'a -> 'a + end +end + +module Declareops : +sig + val constant_has_body : Declarations.constant_body -> bool + val is_opaque : Declarations.constant_body -> bool + val eq_recarg : Declarations.recarg -> Declarations.recarg -> bool + val body_of_constant : + Opaqueproof.opaquetab -> Declarations.constant_body -> Term.constr option +end + +module Constr : +sig + type t = Term.constr + [@@ocaml.deprecated "alias of API.Term.constr"] + + type constr = Term.constr + [@@ocaml.deprecated "alias of API.Term.constr"] + + type types = Term.constr + [@@ocaml.deprecated "alias of API.Term.types"] + + type cast_kind = Term.cast_kind = + | VMcast + | NATIVEcast + | DEFAULTcast + | REVERTcast + type ('constr, 'types, 'sort, 'univs) kind_of_term = ('constr, 'types, 'sort, 'univs) Term.kind_of_term = + | Rel of int + | Var of Names.Id.t + | Meta of Term.metavariable + | Evar of 'constr Term.pexistential + | Sort of 'sort + | Cast of 'constr * cast_kind * 'types + | Prod of Names.Name.t * 'types * 'types + | Lambda of Names.Name.t * 'types * 'constr + | LetIn of Names.Name.t * 'constr * 'types * 'constr + | App of 'constr * 'constr array + | Const of (Names.Constant.t * 'univs) + | Ind of (Names.inductive * 'univs) + | Construct of (Names.constructor * 'univs) + | Case of Term.case_info * 'constr * 'constr * 'constr array + | Fix of ('constr, 'types) Term.pfixpoint + | CoFix of ('constr, 'types) Term.pcofixpoint + | Proj of Names.Projection.t * 'constr + [@@ocaml.deprecated "alias of API.Term.cast_kind"] + + val equal : Term.constr -> Term.constr -> bool + [@@ocaml.deprecated "alias of API.Term.eq_constr"] + + val mkIndU : Term.pinductive -> Term.constr + [@@ocaml.deprecated "alias of API.Term.mkIndU"] + + val mkConstU : Term.pconstant -> Term.constr + [@@ocaml.deprecated "alias of API.Term.mkConstU"] + + val mkConst : Names.Constant.t -> Term.constr + [@@ocaml.deprecated "alias of API.Term.mkConst"] + + val mkVar : Names.Id.t -> Term.constr + [@@ocaml.deprecated "alias of API.Term.mkVar"] + + val compare : Term.constr -> Term.constr -> int + [@@ocaml.deprecated "alias of API.Term.constr_ord"] + + val mkApp : Term.constr * Term.constr array -> Term.constr + [@@ocaml.deprecated "alias of API.Term.mkApp"] +end +[@@ocaml.deprecated "alias of API.Term"] + +module Coq_config : +sig + val exec_extension : string +end + +module Kindops : +sig + val logical_kind_of_goal_kind : Decl_kinds.goal_object_kind -> Decl_kinds.logical_kind +end + +module States : +sig + val with_state_protection_on_exception : ('a -> 'b) -> 'a -> 'b + val with_state_protection : ('a -> 'b) -> 'a -> 'b +end + +module Command : +sig + type structured_fixpoint_expr = Command.structured_fixpoint_expr + type recursive_preentry = Names.Id.t list * Term.constr option list * Term.types list + type structured_inductive_expr = Command.structured_inductive_expr + type one_inductive_impls = Command.one_inductive_impls + + val do_mutual_inductive : + (Vernacexpr.one_inductive_expr * Vernacexpr.decl_notation list) list -> Decl_kinds.polymorphic -> + Decl_kinds.private_flag -> Decl_kinds.recursivity_kind -> unit + + val do_definition : Names.Id.t -> Decl_kinds.definition_kind -> Vernacexpr.lident list option -> + Constrexpr.local_binder_expr list -> Redexpr.red_expr option -> Constrexpr.constr_expr -> + Constrexpr.constr_expr option -> unit Lemmas.declaration_hook -> unit + + val do_fixpoint : + Decl_kinds.locality -> Decl_kinds.polymorphic -> (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list -> unit + + val extract_fixpoint_components : bool -> + (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list -> + structured_fixpoint_expr list * Vernacexpr.decl_notation list + + val interp_fixpoint : + structured_fixpoint_expr list -> Vernacexpr.decl_notation list -> + recursive_preentry * Vernacexpr.lident list option * UState.t * + (EConstr.rel_context * Impargs.manual_implicits * int option) list + + val extract_mutual_inductive_declaration_components : + (Vernacexpr.one_inductive_expr * Vernacexpr.decl_notation list) list -> + structured_inductive_expr * Libnames.qualid list * Vernacexpr.decl_notation list + + val interp_mutual_inductive : + structured_inductive_expr -> Vernacexpr.decl_notation list -> Decl_kinds.polymorphic -> + Decl_kinds.private_flag -> Decl_kinds.recursivity_kind -> + Entries.mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list + + val declare_mutual_inductive_with_eliminations : + Entries.mutual_inductive_entry -> Universes.universe_binders -> one_inductive_impls list -> + Names.MutInd.t +end + +module Ppvernac : +sig + val pr_vernac : Vernacexpr.vernac_expr -> Pp.std_ppcmds + val pr_rec_definition : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) -> Pp.std_ppcmds +end + +module Topconstr : +sig + val replace_vars_constr_expr : + Names.Id.t Names.Id.Map.t -> Constrexpr.constr_expr -> Constrexpr.constr_expr +end diff --git a/API/API.mllib b/API/API.mllib new file mode 100644 index 000000000..f4bdf83db --- /dev/null +++ b/API/API.mllib @@ -0,0 +1,2 @@ +API +Grammar_API diff --git a/API/PROPERTIES b/API/PROPERTIES new file mode 100644 index 000000000..cd942e202 --- /dev/null +++ b/API/PROPERTIES @@ -0,0 +1,8 @@ +0 : All API elements, i.e.: + - modules + - module types + - functions & values + - types + are present if and only if are needed for implementing Coq plugins. + +1 : Individual API elements are not aliased. diff --git a/API/grammar_API.ml b/API/grammar_API.ml new file mode 100644 index 000000000..2f3da8d98 --- /dev/null +++ b/API/grammar_API.ml @@ -0,0 +1,63 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +module G_proofs = G_proofs +module Metasyntax = Metasyntax +module Egramcoq = Egramcoq +module G_vernac = G_vernac +module Pcoq = Pcoq +module Tok = Tok +module CLexer = CLexer +module Egramml = Egramml +module Mltop = Mltop +module Vernacinterp = Vernacinterp +module Genintern = Genintern + +module Extend = + struct + type 'a entry = 'a Extend.entry + type ('self, 'a) symbol = ('self, 'a) Extend.symbol = + | Atoken : Tok.t -> ('self, string) symbol + | Alist1 : ('self, 'a) symbol -> ('self, 'a list) symbol + | Alist1sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol + | Alist0 : ('self, 'a) symbol -> ('self, 'a list) symbol + | Alist0sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol + | Aopt : ('self, 'a) symbol -> ('self, 'a option) symbol + | Aself : ('self, 'self) symbol + | Anext : ('self, 'self) symbol + | Aentry : 'a entry -> ('self, 'a) symbol + | Aentryl : 'a entry * int -> ('self, 'a) symbol + | Arules : 'a rules list -> ('self, 'a) symbol + and ('self, 'a, 'r) rule = ('self, 'a, 'r) Extend.rule = + | Stop : ('self, 'r, 'r) rule + | Next : ('self, 'a, 'r) rule * ('self, 'b) symbol -> ('self, 'b -> 'a, 'r) rule + and ('a, 'r) norec_rule = ('a, 'r) Extend.norec_rule = + { norec_rule : 's. ('s, 'a, 'r) rule } + and 'a rules = 'a Extend.rules = + | Rules : ('act, Loc.t -> 'a) norec_rule * 'act -> 'a rules + type gram_assoc = Extend.gram_assoc = NonA | RightA | LeftA + type 'a production_rule = 'a Extend.production_rule = + | Rule : ('a, 'act, Loc.t -> 'a) rule * 'act -> 'a production_rule + type 'a single_extend_statment = string option * gram_assoc option * 'a production_rule list + type gram_position = Extend.gram_position = + | First + | Last + | Before of string + | After of string + | Level of string + type 'a extend_statment = Extend.gram_position option * 'a single_extend_statment list + + type 'a user_symbol = 'a Extend.user_symbol = + | Ulist1 of 'a user_symbol + | Ulist1sep of 'a user_symbol * string + | Ulist0 of 'a user_symbol + | Ulist0sep of 'a user_symbol * string + | Uopt of 'a user_symbol + | Uentry of 'a + | Uentryl of 'a * int + end diff --git a/API/grammar_API.mli b/API/grammar_API.mli new file mode 100644 index 000000000..c643f0908 --- /dev/null +++ b/API/grammar_API.mli @@ -0,0 +1,249 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +module Extend : +sig + type 'a entry = 'a Pcoq.Gram.Entry.e + type ('self, 'a) symbol = ('self, 'a) Extend.symbol = + | Atoken : Tok.t -> ('self, string) symbol + | Alist1 : ('self, 'a) symbol -> ('self, 'a list) symbol + | Alist1sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol + | Alist0 : ('self, 'a) symbol -> ('self, 'a list) symbol + | Alist0sep : ('self, 'a) symbol * ('self, _) symbol -> ('self, 'a list) symbol + | Aopt : ('self, 'a) symbol -> ('self, 'a option) symbol + | Aself : ('self, 'self) symbol + | Anext : ('self, 'self) symbol + | Aentry : 'a entry -> ('self, 'a) symbol + | Aentryl : 'a entry * int -> ('self, 'a) symbol + | Arules : 'a rules list -> ('self, 'a) symbol + and ('self, 'a, 'r) rule = ('self, 'a, 'r) Extend.rule = + | Stop : ('self, 'r, 'r) rule + | Next : ('self, 'a, 'r) rule * ('self, 'b) symbol -> ('self, 'b -> 'a, 'r) rule + and ('a, 'r) norec_rule = ('a, 'r) Extend.norec_rule = + { norec_rule : 's. ('s, 'a, 'r) rule } + and 'a rules = 'a Extend.rules = + | Rules : ('act, Loc.t -> 'a) norec_rule * 'act -> 'a rules + type gram_assoc = Extend.gram_assoc = NonA | RightA | LeftA + type 'a production_rule = 'a Extend.production_rule = + | Rule : ('a, 'act, Loc.t -> 'a) rule * 'act -> 'a production_rule + type 'a single_extend_statment = string option * gram_assoc option * 'a production_rule list + type gram_position = Extend.gram_position = + | First + | Last + | Before of string + | After of string + | Level of string + type 'a extend_statment = gram_position option * 'a single_extend_statment list + type 'a user_symbol = 'a Extend.user_symbol = + | Ulist1 of 'a user_symbol + | Ulist1sep of 'a user_symbol * string + | Ulist0 of 'a user_symbol + | Ulist0sep of 'a user_symbol * string + | Uopt of 'a user_symbol + | Uentry of 'a + | Uentryl of 'a * int +end + +module Genintern : +sig + open API + module Store : module type of struct include Genintern.Store end + type glob_sign = Genintern.glob_sign = + { ltacvars : Names.Id.Set.t; + genv : Environ.env; + extra : Store.t } + type ('raw, 'glb) intern_fun = glob_sign -> 'raw -> glob_sign * 'glb + type 'glb subst_fun = Mod_subst.substitution -> 'glb -> 'glb + type 'glb ntn_subst_fun = Tactypes.glob_constr_and_expr Names.Id.Map.t -> 'glb -> 'glb + val empty_glob_sign : Environ.env -> glob_sign + val register_intern0 : ('raw, 'glb, 'top) Genarg.genarg_type -> + ('raw, 'glb) intern_fun -> unit + val register_subst0 : ('raw, 'glb, 'top) Genarg.genarg_type -> + 'glb subst_fun -> unit + val register_ntn_subst0 : ('raw, 'glb, 'top) Genarg.genarg_type -> + 'glb ntn_subst_fun -> unit + val generic_substitute : Genarg.glob_generic_argument subst_fun + val generic_intern : (Genarg.raw_generic_argument, Genarg.glob_generic_argument) intern_fun +end + +module Tok : +sig + type t = Tok.t = + | KEYWORD of string + | PATTERNIDENT of string + | IDENT of string + | FIELD of string + | INT of string + | STRING of string + | LEFTQMARK + | BULLET of string + | EOI +end + +module Pcoq : +sig + type gram_universe = Pcoq.gram_universe + module Gram : + sig + type te = Tok.t + module Entry : + sig + type 'a e = 'a Extend.entry + val of_parser : string -> (te Stream.t -> 'a) -> 'a e + val obj : 'a e -> te Gramext.g_entry + val create : string -> 'a e + end + type 'a entry = 'a Entry.e + val extend : 'a Pcoq.Gram.Entry.e -> Gramext.position option -> + (string option * Gramext.g_assoc option * + (Tok.t Gramext.g_symbol list * Gramext.g_action) list) list -> unit + val entry_create : string -> 'a Entry.e + end + module Prim : sig + open Names + open Loc + val preident : string Gram.Entry.e + val ident : Names.Id.t Gram.Entry.e + val name : Name.t located Gram.Entry.e + val identref : Names.Id.t located Gram.Entry.e + val pidentref : (Names.Id.t located * (Names.Id.t located list) option) Gram.Entry.e + val pattern_ident : Names.Id.t Gram.Entry.e + val pattern_identref : Names.Id.t located Gram.Entry.e + val base_ident : Names.Id.t Gram.Entry.e + val natural : int Gram.Entry.e + val bigint : Constrexpr.raw_natural_number Gram.Entry.e + val integer : int Gram.Entry.e + val string : string Gram.Entry.e + val qualid : API.Libnames.qualid located Gram.Entry.e + val fullyqualid : Names.Id.t list located Gram.Entry.e + val reference : API.Libnames.reference Gram.Entry.e + val by_notation : (string * string option) Loc.located Gram.entry + val smart_global : API.Libnames.reference API.Misctypes.or_by_notation Gram.Entry.e + val dirpath : DirPath.t Gram.Entry.e + val ne_string : string Gram.Entry.e + val ne_lstring : string located Gram.Entry.e + val var : Names.Id.t located Gram.Entry.e + end + + val eoi_entry : 'a Gram.Entry.e -> 'a Gram.Entry.e + val create_generic_entry : gram_universe -> string -> + ('a, Genarg.rlevel) Genarg.abstract_argument_type -> 'a Gram.Entry.e + val utactic : gram_universe + type gram_reinit = Extend.gram_assoc * Extend.gram_position + val grammar_extend : 'a Gram.Entry.e -> gram_reinit option -> + 'a Extend.extend_statment -> unit + val genarg_grammar : ('raw, 'glb, 'top) Genarg.genarg_type -> 'raw Gram.Entry.e + val register_grammar : ('raw, 'glb, 'top) Genarg.genarg_type -> 'raw Gram.Entry.e -> unit + module Constr : + sig + val sort : API.Misctypes.glob_sort Gram.Entry.e + val lconstr : API.Constrexpr.constr_expr Gram.Entry.e + val lconstr_pattern : API.Constrexpr.constr_expr Gram.Entry.e + val ident : API.Names.Id.t Gram.Entry.e + val constr : API.Constrexpr.constr_expr Gram.Entry.e + val closed_binder : API.Constrexpr.local_binder_expr list Gram.Entry.e + val constr_pattern : API.Constrexpr.constr_expr Gram.Entry.e + val global : API.Libnames.reference Gram.Entry.e + val binder_constr : API.Constrexpr.constr_expr Gram.Entry.e + val operconstr : API.Constrexpr.constr_expr Gram.Entry.e + val pattern : API.Constrexpr.cases_pattern_expr Gram.Entry.e + val binders : API.Constrexpr.local_binder_expr list Gram.Entry.e + end + module Vernac_ : + sig + val gallina : API.Vernacexpr.vernac_expr Gram.Entry.e + val gallina_ext : API.Vernacexpr.vernac_expr Gram.Entry.e + val red_expr : API.Genredexpr.raw_red_expr Gram.Entry.e + val noedit_mode : API.Vernacexpr.vernac_expr Gram.Entry.e + val command : API.Vernacexpr.vernac_expr Gram.Entry.e + val rec_definition : (API.Vernacexpr.fixpoint_expr * API.Vernacexpr.decl_notation list) Gram.Entry.e + val vernac : API.Vernacexpr.vernac_expr Gram.Entry.e + end + + type extend_rule = + | ExtendRule : 'a Gram.Entry.e * gram_reinit option * 'a Extend.extend_statment -> extend_rule + + module GramState : module type of struct include Pcoq.GramState end + type 'a grammar_command = 'a Pcoq.grammar_command + type 'a grammar_extension = 'a -> GramState.t -> extend_rule list * GramState.t + val create_grammar_command : string -> 'a grammar_extension -> 'a grammar_command + val extend_grammar_command : 'a grammar_command -> 'a -> unit + val epsilon_value : ('a -> 'self) -> ('self, 'a) Extend.symbol -> 'self option + val parse_string : 'a Gram.Entry.e -> string -> 'a + val (!@) : Ploc.t -> Loc.t + val set_command_entry : API.Vernacexpr.vernac_expr Gram.Entry.e -> unit + val to_coqloc : Ploc.t -> Loc.t +end + +module CLexer : +sig + type keyword_state = CLexer.keyword_state + val terminal : string -> Tok.t + val add_keyword : string -> unit + val is_keyword : string -> bool + val check_ident : string -> unit + val get_keyword_state : unit -> keyword_state + val set_keyword_state : keyword_state -> unit +end + +module Egramml : +sig + type 's grammar_prod_item = 's Egramml.grammar_prod_item = + | GramTerminal of string + | GramNonTerminal : ('a Genarg.raw_abstract_argument_type option * + ('s, 'a) Extend.symbol) Loc.located -> 's grammar_prod_item + + + val extend_vernac_command_grammar : + API.Vernacexpr.extend_name -> Vernacexpr.vernac_expr Pcoq.Gram.Entry.e option -> + Vernacexpr.vernac_expr grammar_prod_item list -> unit + + val make_rule : + (Loc.t -> Genarg.raw_generic_argument list -> 'a) -> + 'a grammar_prod_item list -> 'a Extend.production_rule +end + +module Mltop : +sig + val add_known_module : string -> unit + val module_is_known : string -> bool + val declare_cache_obj : (unit -> unit) -> string -> unit +end +module Vernacinterp : +sig + type deprecation = bool + type vernac_command = Genarg.raw_generic_argument list -> unit -> unit + val vinterp_add : deprecation -> API.Vernacexpr.extend_name -> + vernac_command -> unit +end + +module G_vernac : +sig + val def_body : API.Vernacexpr.definition_expr Pcoq.Gram.Entry.e + val section_subset_expr : API.Vernacexpr.section_subset_expr Pcoq.Gram.Entry.e + val query_command : (Vernacexpr.goal_selector option -> Vernacexpr.vernac_expr) + Pcoq.Gram.Entry.e +end + +module G_proofs : +sig + val hint : Vernacexpr.hints_expr Pcoq.Gram.Entry.e + val hint_proof_using : 'a Pcoq.Gram.Entry.e -> 'a option -> 'a option +end + +module Egramcoq : +sig +end + +module Metasyntax : +sig + type any_entry = Metasyntax.any_entry = + | AnyEntry : 'a Pcoq.Gram.Entry.e -> any_entry + val register_grammar : string -> any_entry list -> unit + val add_token_obj : string -> unit +end @@ -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 @@ -81,6 +85,15 @@ Tools warnings when a deprecated feature is used. Please upgrade your _CoqProject accordingly. +Build Infrastructure + +- Note that 'make world' does not build the bytecode binaries anymore. + For that, you can use 'make byte' (and 'make install-byte' afterwards). + Warning: native and byte compilations should *not* be mixed in the same + instance of 'make -j', otherwise both ocamlc and ocamlopt might race for + access to the same .cmi files. In short, use "make -j && make -j byte" + instead of "make -j world byte". + Changes from V8.6beta1 to V8.6 ============================== @@ -55,8 +55,6 @@ QUICK INSTALLATION PROCEDURE. 1. ./configure 2. make 3. make install (you may need superuser rights) -4. make clean - INSTALLATION PROCEDURE IN DETAILS (NORMAL USERS). ================================================= @@ -131,10 +129,13 @@ INSTALLATION PROCEDURE IN DETAILS (NORMAL USERS). make - to compile Coq in Objective Caml bytecode (and native-code if supported). + to compile Coq in the best OCaml mode available (native-code if supported, + bytecode otherwise). This will compile the entire system. This phase can take more or less time, - depending on your architecture and is fairly verbose. + depending on your architecture and is fairly verbose. On a multi-core machine, + it is recommended to compile in parallel, via make -jN where N is your number + of cores. 6- You can now install the Coq system. Executables, libraries, manual pages and emacs mode are copied in some standard places of your system, defined at @@ -150,7 +151,19 @@ INSTALLATION PROCEDURE IN DETAILS (NORMAL USERS). (setq auto-mode-alist (cons '("\\.v$" . coq-mode) auto-mode-alist)) (autoload 'coq-mode "gallina" "Major mode for editing Coq vernacular." t) -7- You can now clean all the sources. (You can even erase them.) +7- Optionally, you could build the bytecode version of Coq via: + + make byte + + and install it via + + make install-byte + + This version is quite slower than the native code version of Coq, but could + be helpful for debugging purposes. In particular, coqtop.byte embeds an OCaml + toplevel accessible via the Drop command. + +8- You can now clean all the sources. (You can even erase them.) make clean @@ -182,11 +195,14 @@ THE AVAILABLE COMMANDS. coqtop The Coq toplevel coqc The Coq compiler - Under architecture where ocamlopt is available, there are actually two - binaries for the interactive system, coqtop.byte and coqtop (respectively - bytecode and native code versions of Coq). coqtop is a link to coqtop.byte - otherwise. coqc also invokes the fastest version of Coq. Options -opt and - -byte to coqtop and coqc selects a particular binary. + Under architecture where ocamlopt is available, coqtop is the native code + version of Coq. On such architecture, you could additionally request + the build of the bytecode version of Coq via 'make byte' and install it via + 'make install-byte'. This will create an extra binary named coqtop.byte, + that could be used for debugging purpose. If native code isn't available, + coqtop.byte is directly built by 'make', and coqtop is a link to coqtop.byte. + coqc also invokes the fastest version of Coq. Options -opt and -byte to coqtop + and coqc selects a particular binary. * `coqtop' launches Coq in the interactive mode. By default it loads basic logical definitions and tactics from the Init directory. @@ -291,3 +291,16 @@ package "ltac" ( archive(native) = "ltac_plugin.cmx" ) + +package "API" ( + + description = "Coq API" + version = "8.7" + + requires = "coq.toplevel" + directory = "API" + + archive(byte) = "API.cma" + archive(native) = "API.cmxa" + +) @@ -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. @@ -116,16 +115,19 @@ NOARG: world .PHONY: NOARG help noconfig submake help: - @echo "Please use either" + @echo "Please use either:" @echo " ./configure" @echo " make world" @echo " make install" @echo " make clean" @echo "or make archclean" - @echo @echo "For make to be verbose, add VERBOSE=1" + @echo "If you want camlp5 to generate human-readable files, add READABLE_ML4=1" @echo - @echo "If you want camlp{4,5} to generate human-readable files, add READABLE_ML4=1" + @echo "Bytecode compilation is now a separate target:" + @echo " make byte" + @echo " make install-byte" + @echo "Please do not mix bytecode and native targets in the same make -j" UNSAVED_FILES:=$(shell find . -name '.\#*v' -o -name '.\#*.ml' -o -name '.\#*.ml?') ifdef UNSAVED_FILES @@ -188,6 +190,7 @@ indepclean: rm -f test-suite/check.log rm -f glob.dump rm -f config/revision.ml revision + rm -f plugins/micromega/.micromega.ml.generated $(MAKE) -C test-suite clean docclean: diff --git a/Makefile.build b/Makefile.build index da736345c..0dafde997 100644 --- a/Makefile.build +++ b/Makefile.build @@ -51,9 +51,16 @@ COQ_XML ?= world: coq coqide documentation revision -coq: coqlib coqbinaries tools printers +coq: coqlib coqbinaries tools -.PHONY: world coq +# Note: 'world' does not build the bytecode binaries anymore. +# For that, you can use the 'byte' rule. Native and byte compilations +# shouldn't be done in a same make -j... run, otherwise both ocamlc and +# ocamlopt might race for access to the same .cmi files. + +byte: coqbyte coqide-byte pluginsbyte printers + +.PHONY: world coq byte ########################################################################### # Includes @@ -71,27 +78,6 @@ include Makefile.install include Makefile.dev ## provides the 'printers' and 'revision' rules ########################################################################### -# Adding missing pieces of information not discovered by ocamldep -# due to the fact that: -# - plugins/micromega/micromega_plugin.ml -# - plugins/micromega/micromega_plugin.mli -# are generated (and not yet present when we run "ocamldep"). -########################################################################### - -plugins/micromega/micromega_plugin.cmo : plugins/micromega/micromega.cmo -plugins/micromega/micromega_plugin.cmx : plugins/micromega/micromega.cmx - -plugins/micromega/certificate.cmo plugins/micromega/coq_micromega.cmo plugins/micromega/csdpcert.cmo plugins/micromega/mfourier.cmo plugins/micromega/mutils.cmo plugins/micromega/polynomial.cmo : plugins/micromega/micromega.cmo - -plugins/micromega/certificate.cmx plugins/micromega/coq_micromega.cmx plugins/micromega/csdpcert.cmx plugins/micromega/mfourier.cmx plugins/micromega/mutils.cmx plugins/micromega/polynomial.cmx : plugins/micromega/micromega.cmx - -plugins/micromega/micromega.cmx plugins/micromega/micromega.cmo : plugins/micromega/micromega.cmi -plugins/micromega/micromega.cmi : plugins/micromega/micromega.mli - -plugins/micromega/micromega.mli plugins/micromega/micromega.ml : plugins/micromega/MExtraction.vo - @: - -########################################################################### # This include below will lauch the build of all .d. # The - at front is for disabling warnings about currently missing ones. @@ -103,8 +89,6 @@ DEPENDENCIES := \ -include $(DEPENDENCIES) -plugins/micromega/micromega_FORPACK:= -for-pack Micromega_plugin - # All dependency includes must be declared secondary, otherwise make will # delete them if it decided to build them by dependency instead of because # of include, and they will then be automatically deleted, leading to an @@ -126,9 +110,9 @@ TIMER=$(if $(TIMED), $(STDTIME), $(TIMECMD)) # TIME="%C (%U user, %S sys, %e total, %M maxres)" COQOPTS=$(COQ_XML) $(NATIVECOMPUTE) -BOOTCOQC=$(TIMER) $(COQTOPEXE) -boot $(COQOPTS) -compile +BOOTCOQC=$(TIMER) $(COQTOPBEST) -boot $(COQOPTS) -compile -LOCALINCLUDES=$(addprefix -I , $(SRCDIRS) ) +LOCALINCLUDES=$(if $(filter plugins/%,$<),-I lib -I API $(addprefix -I plugins/,$(PLUGINDIRS)),$(addprefix -I ,$(SRCDIRS))) MLINCLUDES=$(LOCALINCLUDES) -I $(MYCAMLP4LIB) OCAMLC := $(OCAMLFIND) ocamlc $(CAMLFLAGS) @@ -136,7 +120,7 @@ OCAMLOPT := $(OCAMLFIND) opt $(CAMLFLAGS) BYTEFLAGS=-thread $(CAMLDEBUG) $(USERFLAGS) OPTFLAGS=-thread $(CAMLDEBUGOPT) $(CAMLTIMEPROF) $(USERFLAGS) -DEPFLAGS= $(LOCALINCLUDES) -I ide -I ide/utils +DEPFLAGS=$(LOCALINCLUDES)$(if $(filter plugins/%,$<),, -I ide -I ide/utils) # On MacOS, the binaries are signed, except our private ones ifeq ($(shell which codesign > /dev/null 2>&1 && echo $(ARCH)),Darwin) @@ -208,7 +192,7 @@ ifndef ORDER_ONLY_SEP $(error This Makefile needs GNU Make 3.81 or later (that is a version that supports the order-only dependency feature without major bugs.)) endif -VO_TOOLS_DEP := $(COQTOPEXE) +VO_TOOLS_DEP := $(COQTOPBEST) ifdef COQ_XML VO_TOOLS_DEP += $(COQDOC) endif @@ -315,11 +299,11 @@ grammar/%.cmi: grammar/%.mli # Main targets (coqmktop, coqtop.opt, coqtop.byte) ########################################################################### -.PHONY: coqbinaries +.PHONY: coqbinaries coqbyte -coqbinaries: $(COQMKTOP) $(COQTOPEXE) $(COQTOPBYTE) \ - $(CHICKEN) $(CHICKENBYTE) $(CSDPCERT) $(FAKEIDE) +coqbinaries: $(COQMKTOP) $(COQTOPEXE) $(CHICKEN) $(CSDPCERT) $(FAKEIDE) +coqbyte: $(COQTOPBYTE) $(CHICKENBYTE) ifeq ($(BEST),opt) $(COQTOPEXE): $(COQMKTOP) $(LINKCMX) $(LIBCOQRUN) $(TOPLOOPCMA:.cma=.cmxs) @@ -510,18 +494,13 @@ kernel/kernel.cma: kernel/kernel.mllib # For plugin packs -# Note: both ocamlc -pack and ocamlopt -pack will create the same .cmi, and there's -# apparently no way to avoid that (no -intf-suffix hack as below). -# We at least ensure that these two commands won't run at the same time, by a fake -# dependency from the packed .cmx to the packed .cmo. - %.cmo: %.mlpack $(SHOW)'OCAMLC -pack -o $@' $(HIDE)$(OCAMLC) $(MLINCLUDES) $(BYTEFLAGS) -pack -o $@ $(filter-out %.mlpack, $^) -%.cmx: %.mlpack %.cmo +%.cmx: %.mlpack $(SHOW)'OCAMLOPT -pack -o $@' - $(HIDE)$(OCAMLOPT) $(MLINCLUDES) $(OPTFLAGS) -pack -o $@ $(filter-out %.mlpack %.cmo, $^) + $(HIDE)$(OCAMLOPT) $(MLINCLUDES) $(OPTFLAGS) -pack -o $@ $(filter-out %.mlpack, $^) COND_BYTEFLAGS= \ $(if $(filter tools/fake_ide% tools/coq_makefile%,$<), -I ide,) $(MLINCLUDES) $(BYTEFLAGS) @@ -629,11 +608,6 @@ plugins: $(PLUGINSVO) .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) @@ -641,6 +615,26 @@ theories/Init/%.vo theories/Init/%.glob: theories/Init/%.v $(VO_TOOLS_DEP) $(HIDE)rm -f theories/Init/$*.glob $(HIDE)$(BOOTCOQC) $< $(COQ_XML) -noinit -R theories Coq +# MExtraction.v generates the ml core file of the micromega tactic. +# We check that this generated code is still in sync with the version +# of micromega.ml in the archive. + +# Note: we now dump to stdout there via "Recursive Extraction" for better +# control on the name of the generated file, and avoid a .ml that +# would end in our $(MLFILES). The "sed" below is to kill the final +# blank line printed by Recursive Extraction (unlike Extraction "foo"). + +MICROMEGAV:=plugins/micromega/MExtraction.v +MICROMEGAML:=plugins/micromega/micromega.ml +MICROMEGAGEN:=plugins/micromega/.micromega.ml.generated + +$(MICROMEGAV:.v=.vo) $(MICROMEGAV:.v=.glob) : $(MICROMEGAV) theories/Init/Prelude.vo $(VO_TOOLS_DEP) + $(SHOW)'COQC $<' + $(HIDE)rm -f $*.glob + $(HIDE)$(BOOTCOQC) $< | sed -e '$$d' > $(MICROMEGAGEN) + $(HIDE)cmp -s $(MICROMEGAML) $(MICROMEGAGEN) || \ + echo "Warning: $(MICROMEGAML) and the code generated by $(MICROMEGAV) differ !" + # The general rule for building .vo files : %.vo %.glob: %.v theories/Init/Prelude.vo $(VO_TOOLS_DEP) @@ -655,9 +649,9 @@ 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 $(DEPNATDYN) "$<" $(TOTARGET) + $(HIDE)$(COQDEPBOOT) -boot $(DYNDEP) "$<" $(TOTARGET) ########################################################################### diff --git a/Makefile.checker b/Makefile.checker index 3ea0baced..435d8e8f6 100644 --- a/Makefile.checker +++ b/Makefile.checker @@ -71,7 +71,7 @@ checker/%.cmo: checker/%.ml checker/%.cmx: checker/%.ml $(SHOW)'OCAMLOPT $<' - $(HIDE)$(OCAMLOPT) $(CHKLIBS) $(OPTFLAGS) $(HACKMLI) -c $< + $(HIDE)$(OCAMLOPT) $(CHKLIBS) $(OPTFLAGS) -c $< md5chk: $(SHOW)'MD5SUM cic.mli' diff --git a/Makefile.ci b/Makefile.ci index e4c63af9d..3be90c0a3 100644 --- a/Makefile.ci +++ b/Makefile.ci @@ -1,8 +1,8 @@ CI_TARGETS=ci-all \ - ci-bedrock-facade \ - ci-bedrock-src \ + ci-bignums \ ci-color \ ci-compcert \ + ci-coq-dpdgraph \ ci-coquelicot \ ci-cpdt \ ci-fiat-crypto \ diff --git a/Makefile.common b/Makefile.common index 4545fad05..ec5e6ac85 100644 --- a/Makefile.common +++ b/Makefile.common @@ -41,10 +41,26 @@ CSDPCERT:=plugins/micromega/csdpcert$(EXE) # Object and Source files ########################################################################### -ifeq ($(HASNATDYNLINK)-$(BEST),true-opt) - DEPNATDYN:= +ifeq ($(HASNATDYNLINK)-$(BEST),false-opt) + # static link of plugins, do not mention them in .v.d + DYNDEP:=-dyndep no +else + DYNDEP:=-dyndep var +endif + +# Which coqtop do we use to build .vo file ? The best ;-) +# Note: $(BEST) could be overridden by the user if a byte build is desired +# Note: coqdep -dyndep var will use $(DYNOBJ) and $(DYNLIB) extensions +# for Declare ML Module files. + +ifeq ($(BEST),opt) +COQTOPBEST:=$(COQTOPEXE) +DYNOBJ:=.cmxs +DYNLIB:=.cmxs else - DEPNATDYN:=-natdynlink no +COQTOPBEST:=$(COQTOPBYTE) +DYNOBJ:=.cmo +DYNLIB:=.cma endif INSTALLBIN:=install @@ -55,7 +71,7 @@ MKDIR:=install -d CORESRCDIRS:=\ config lib kernel intf kernel/byterun library \ engine pretyping interp proofs parsing printing \ - tactics vernac stm toplevel + tactics vernac stm toplevel API PLUGINDIRS:=\ omega romega micromega quote \ @@ -80,10 +96,10 @@ BYTERUN:=$(addprefix kernel/byterun/, \ # respecting this order is useful for developers that want to load or link # the libraries directly -CORECMA:=lib/clib.cma lib/lib.cma kernel/kernel.cma library/library.cma \ +CORECMA:=lib/clib.cma lib/lib.cma kernel/kernel.cma intf/intf.cma library/library.cma \ engine/engine.cma pretyping/pretyping.cma interp/interp.cma proofs/proofs.cma \ parsing/parsing.cma printing/printing.cma tactics/tactics.cma vernac/vernac.cma \ - stm/stm.cma toplevel/toplevel.cma parsing/highparsing.cma + stm/stm.cma toplevel/toplevel.cma parsing/highparsing.cma API/API.cma TOPLOOPCMA:=stm/proofworkertop.cma stm/tacworkertop.cma stm/queryworkertop.cma @@ -113,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 @@ -145,16 +161,8 @@ LINKCMX:=$(CORECMA:.cma=.cmxa) $(STATICPLUGINS:.cmo=.cmx) # vo files ########################################################################### -## we now retrieve the names of .vo file to compile in */vo.itarget 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/Makefile.dev b/Makefile.dev index fde92ec94..b0299bd16 100644 --- a/Makefile.dev +++ b/Makefile.dev @@ -120,7 +120,7 @@ highparsing: parsing/highparsing.cma stm: stm/stm.cma toplevel: toplevel/toplevel.cma -.PHONY: lib kernel byterun library proofs tactics interp parsing pretyping +.PHONY: lib kernel byterun library proofs tactics interp parsing pretyping API .PHONY: engine highparsing stm toplevel ###################### @@ -186,7 +186,7 @@ omega: $(OMEGAVO) $(OMEGACMO) $(ROMEGAVO) $(ROMEGACMO) micromega: $(MICROMEGAVO) $(MICROMEGACMO) $(CSDPCERT) setoid_ring: $(RINGVO) $(RINGCMO) nsatz: $(NSATZVO) $(NSATZCMO) -extraction: $(EXTRACTIONCMO) +extraction: $(EXTRACTIONCMO) $(EXTRACTIONVO) fourier: $(FOURIERVO) $(FOURIERCMO) funind: $(FUNINDCMO) $(FUNINDVO) cc: $(CCVO) $(CCCMO) diff --git a/Makefile.doc b/Makefile.doc index c31d81c8b..6a81b292e 100644 --- a/Makefile.doc +++ b/Makefile.doc @@ -415,7 +415,7 @@ OCAMLDOCDIR=dev/ocamldoc DOCMLIS=$(wildcard ./lib/*.mli ./intf/*.mli ./kernel/*.mli ./library/*.mli \ ./engine/*.mli ./pretyping/*.mli ./interp/*.mli printing/*.mli \ - ./parsing/*.mli ./proofs/*.mli \ + ./parsing/*.mli ./proofs/*.mli API/API.mli \ ./tactics/*.mli ./stm/*.mli ./toplevel/*.mli ./ltac/*.mli) # Defining options to generate dependencies graphs diff --git a/Makefile.ide b/Makefile.ide index 48a269ab7..0cfbdeb4e 100644 --- a/Makefile.ide +++ b/Makefile.ide @@ -61,23 +61,30 @@ GTKLIBS=$(shell pkg-config --variable=libdir gtk+-2.0) # CoqIde special targets ########################################################################### -.PHONY: coqide coqide-binaries coqide-no coqide-byte coqide-opt coqide-files -.PHONY: ide-toploop +.PHONY: coqide coqide-opt coqide-byte coqide-files +.PHONY: ide-toploop ide-byteloop ide-optloop # target to build CoqIde -coqide: coqide-files coqide-binaries theories/Init/Prelude.vo +coqide: coqide-files coqide-opt theories/Init/Prelude.vo -coqide-binaries: coqide-$(HASCOQIDE) ide-toploop -coqide-no: -coqide-byte: $(COQIDEBYTE) $(COQIDE) -coqide-opt: $(COQIDEBYTE) $(COQIDE) -coqide-files: $(IDEFILES) -ifeq ($(BEST),opt) -ide-toploop: $(IDETOPLOOPCMA) $(IDETOPLOOPCMA:.cma=.cmxs) +ifeq ($(HASCOQIDE),opt) +coqide-opt: $(COQIDE) ide-toploop else -ide-toploop: $(IDETOPLOOPCMA) +coqide-opt: ide-toploop endif +ifeq ($(HASCOQIDE),no) +coqide-byte: ide-byteloop +else +coqide-byte: $(COQIDEBYTE) ide-byteloop +endif + +coqide-files: $(IDEFILES) + +ide-byteloop: $(IDETOPLOOPCMA) +ide-optloop: $(IDETOPLOOPCMA:.cma=.cmxs) +ide-toploop: ide-$(BEST)loop + ifeq ($(HASCOQIDE),opt) $(COQIDE): $(LINKIDEOPT) $(SHOW)'OCAMLOPT -o $@' @@ -109,14 +116,14 @@ ide/%.cmo: ide/%.ml ide/%.cmx: ide/%.ml $(SHOW)'OCAMLOPT $<' - $(HIDE)$(OCAMLOPT) $(COQIDEFLAGS) $(OPTFLAGS) $(HACKMLI) -c $< + $(HIDE)$(OCAMLOPT) $(COQIDEFLAGS) $(OPTFLAGS) -c $< #################### ## Install targets #################### -.PHONY: install-coqide install-ide-bin install-ide-toploop install-ide-files install-ide-info install-ide-devfiles +.PHONY: install-coqide install-ide-bin install-ide-toploop install-ide-files install-ide-info install-ide-devfiles install-ide-byte ifeq ($(HASCOQIDE),no) install-coqide: install-ide-toploop @@ -124,20 +131,26 @@ else install-coqide: install-ide-bin install-ide-toploop install-ide-files install-ide-info install-ide-devfiles endif +# Apparently, coqide.byte is not meant to be installed + +install-ide-byte: + $(MKDIR) $(FULLCOQLIB) + $(INSTALLSH) $(FULLCOQLIB) $(IDECMA) + $(MKDIR) $(FULLCOQLIB)/toploop + $(INSTALLBIN) $(IDETOPLOOPCMA) $(FULLCOQLIB)/toploop/ + install-ide-bin: $(MKDIR) $(FULLBINDIR) $(INSTALLBIN) $(COQIDE) $(FULLBINDIR) install-ide-toploop: - $(MKDIR) $(FULLCOQLIB)/toploop - $(INSTALLBIN) $(IDETOPLOOPCMA) $(FULLCOQLIB)/toploop/ ifeq ($(BEST),opt) $(INSTALLBIN) $(IDETOPLOOPCMA:.cma=.cmxs) $(FULLCOQLIB)/toploop/ endif install-ide-devfiles: $(MKDIR) $(FULLCOQLIB) - $(INSTALLSH) $(FULLCOQLIB) $(IDECMA) \ + $(INSTALLSH) $(FULLCOQLIB) \ $(foreach lib,$(IDECMA:.cma=_MLLIB_DEPENDENCIES),$(addsuffix .cmi,$($(lib)))) ifeq ($(BEST),opt) $(INSTALLSH) $(FULLCOQLIB) $(IDECMA:.cma=.cmxa) $(IDECMA:.cma=.a) diff --git a/Makefile.install b/Makefile.install index 33f881c11..4a3227620 100644 --- a/Makefile.install +++ b/Makefile.install @@ -62,15 +62,26 @@ endif install-coq: install-binaries install-library install-coq-info install-devfiles +ifeq ($(BEST),byte) +install-coq: install-byte +endif + install-binaries: install-tools $(MKDIR) $(FULLBINDIR) - $(INSTALLBIN) $(COQC) $(COQTOPBYTE) $(COQTOPEXE) $(CHICKEN) $(FULLBINDIR) + $(INSTALLBIN) $(COQC) $(COQTOPEXE) $(CHICKEN) $(FULLBINDIR) $(MKDIR) $(FULLCOQLIB)/toploop - $(INSTALLBIN) $(TOPLOOPCMA) $(FULLCOQLIB)/toploop/ ifeq ($(BEST),opt) $(INSTALLBIN) $(TOPLOOPCMA:.cma=.cmxs) $(FULLCOQLIB)/toploop/ endif +install-byte: install-ide-byte + $(MKDIR) $(FULLBINDIR) + $(INSTALLBIN) $(COQTOPBYTE) $(FULLBINDIR) + $(INSTALLBIN) $(TOPLOOPCMA) $(FULLCOQLIB)/toploop/ + $(INSTALLSH) $(FULLCOQLIB) $(LINKCMO) $(PLUGINS) +ifndef CUSTOM + $(INSTALLLIB) $(DLLCOQRUN) $(FULLCOQLIB) +endif install-tools: $(MKDIR) $(FULLBINDIR) @@ -94,7 +105,7 @@ install-devfiles: $(MKDIR) $(FULLBINDIR) $(INSTALLBIN) $(COQMKTOP) $(FULLBINDIR) $(MKDIR) $(FULLCOQLIB) - $(INSTALLSH) $(FULLCOQLIB) $(LINKCMO) $(GRAMMARCMA) + $(INSTALLSH) $(FULLCOQLIB) $(GRAMMARCMA) $(INSTALLSH) $(FULLCOQLIB) $(INSTALLCMI) $(INSTALLSH) $(FULLCOQLIB) tools/CoqMakefile.in ifeq ($(BEST),opt) @@ -103,7 +114,7 @@ endif install-library: $(MKDIR) $(FULLCOQLIB) - $(INSTALLSH) $(FULLCOQLIB) $(LIBFILES) $(PLUGINS) + $(INSTALLSH) $(FULLCOQLIB) $(LIBFILES) $(MKDIR) $(FULLCOQLIB)/user-contrib $(MKDIR) $(FULLCOQLIB)/kernel/byterun ifndef CUSTOM diff --git a/checker/indtypes.ml b/checker/indtypes.ml index c9ee326cb..6c38f38e2 100644 --- a/checker/indtypes.ml +++ b/checker/indtypes.ml @@ -530,7 +530,7 @@ let check_positivity env_ar mind params nrecp inds = let check_inductive env kn mib = Flags.if_verbose Feedback.msg_notice (str " checking ind: " ++ MutInd.print kn); (* check mind_constraints: should be consistent with env *) - let env = add_constraints (Univ.UContext.constraints mib.mind_universes) env in + let env = Environ.push_context (Univ.instantiate_univ_context mib.mind_universes) env in (* check mind_record : TODO ? check #constructor = 1 ? *) (* check mind_finite : always OK *) (* check mind_ntypes *) diff --git a/config/coq_config.mli b/config/coq_config.mli index 2b3bc2c25..3f7b65c39 100644 --- a/config/coq_config.mli +++ b/config/coq_config.mli @@ -53,7 +53,10 @@ val compile_date : string (* compile date *) val vo_magic_number : int val state_magic_number : int +val core_src_dirs : string list +val api_dirs : string list val plugins_dirs : string list +val all_src_dirs : string list val exec_extension : string (* "" under Unix, ".exe" under MS-windows *) val with_geoproof : bool ref (* to (de)activate functions specific to Geoproof with Coqide *) diff --git a/configure.ml b/configure.ml index a5204d5b5..549b32772 100644 --- a/configure.ml +++ b/configure.ml @@ -301,33 +301,37 @@ let args_options = Arg.align [ "-emacslib", arg_string_option Prefs.emacslib, "<dir> Where to install emacs files"; "-emacs", Arg.String (fun s -> - printf "Warning: obsolete -emacs option\n"; + prerr_endline "Warning: -emacs option is deprecated. Use -emacslib instead."; Prefs.emacslib := Some s), - "<dir> Obsolete: same as -emacslib"; + "<dir> Deprecated: same as -emacslib"; "-coqdocdir", arg_string_option Prefs.coqdocdir, "<dir> Where to install Coqdoc style files"; "-ocamlfind", arg_string_option Prefs.ocamlfindcmd, "<dir> Specifies the ocamlfind command to use"; "-lablgtkdir", arg_string_option Prefs.lablgtkdir, "<dir> Specifies the path to the Lablgtk library"; - "-usecamlp5", Arg.Unit (fun () -> ()), - "Deprecated"; + "-usecamlp5", Arg.Unit (fun () -> + prerr_endline "Warning: -usecamlp5 option is deprecated. Camlp5 is already a required dependency."), + " Deprecated: Camlp5 is a required dependency (Camlp4 is not supported anymore)"; "-camlp5dir", Arg.String (fun s -> Prefs.camlp5dir:=Some s), "<dir> Specifies where is the Camlp5 library and tells to use it"; "-arch", arg_string_option Prefs.arch, "<arch> Specifies the architecture"; - "-opt", Arg.Unit (fun () -> printf "Warning: obsolete -opt option\n"), - " Obsolete: native OCaml executables detected automatically"; + "-opt", Arg.Unit (fun () -> + prerr_endline "Warning: -opt option is deprecated. Native OCaml executables are detected automatically."), + " Deprecated: native OCaml executables detected automatically"; "-natdynlink", arg_bool Prefs.natdynlink, "(yes|no) Use dynamic loading of native code or not"; "-coqide", Arg.String (fun s -> Prefs.coqide := Some (get_ide s)), - "(opt|byte|no) Specifies whether or not to compile Coqide"; + "(opt|byte|no) Specifies whether or not to compile CoqIDE"; "-nomacintegration", Arg.Clear Prefs.macintegration, - " Do not try to build coqide mac integration"; + " Do not try to build CoqIDE MacOS integration"; "-browser", arg_string_option Prefs.browser, "<command> Use <command> to open URL %s"; - "-nodoc", Arg.Clear Prefs.withdoc, + "-nodoc", Arg.Unit (fun () -> + prerr_endline "Warning: -nodoc option is deprecated. Use -with-doc no instead."; + Prefs.withdoc := false), " Deprecated: use -with-doc no instead"; "-with-doc", arg_bool Prefs.withdoc, "(yes|no) Compile the documentation or not"; @@ -335,18 +339,23 @@ let args_options = Arg.align [ "(yes|no) Use Geoproof binding or not"; "-byte-only", Arg.Set Prefs.byteonly, " Compiles only bytecode version of Coq"; - "-byteonly", Arg.Set Prefs.byteonly, - " Compiles only bytecode version of Coq"; - "-debug", Arg.Set Prefs.debug, - " Deprecated"; + "-byteonly", Arg.Unit (fun () -> + prerr_endline "Warning: -byteonly option is deprecated. Use -byte-only instead."; + Prefs.byteonly := true), + " Deprecated: use -byte-only instead"; + "-debug", Arg.Unit (fun () -> + prerr_endline "Warning: -debug option is deprecated. Coq is compiled in debug mode by default."; + Prefs.debug := true), + " Deprecated: Coq is compiled in debug mode by default"; "-nodebug", Arg.Clear Prefs.debug, " Do not add debugging information in the Coq executables"; "-profile", Arg.Set Prefs.profile, " Add profiling information in the Coq executables"; "-annotate", Arg.Set Prefs.annotate, " Dumps ml annotation files while compiling Coq"; - "-makecmd", Arg.String (fun _ -> printf "Warning: obsolete -makecmd option\n"), - "<command> Obsolete: name of GNU Make command"; + "-makecmd", Arg.String (fun _ -> + prerr_endline "Warning: -makecmd option is deprecated and doesn't have any effect."), + "<command> Deprecated"; "-native-compiler", arg_bool Prefs.nativecompiler, "(yes|no) Compilation to native code for conversion and normalization"; "-coqwebsite", Arg.Set_string Prefs.coqwebsite, @@ -1088,7 +1097,19 @@ let write_configml f = pr_s "wwwstdlib" (!Prefs.coqwebsite ^ "distrib/" ^ coq_version ^ "/stdlib/"); pr_s "localwwwrefman" ("file:/" ^ docdir ^ "/html/refman"); pr_b "no_native_compiler" (not !Prefs.nativecompiler); + + let core_src_dirs = [ "config"; "dev"; "kernel"; "library"; + "engine"; "pretyping"; "interp"; "parsing"; "proofs"; + "tactics"; "toplevel"; "printing"; "intf"; + "grammar"; "ide"; "stm"; "vernac" ] in + let core_src_dirs = List.fold_left (fun acc core_src_subdir -> acc ^ " \"" ^ core_src_subdir ^ "\";\n") + "" + core_src_dirs in + + pr "\nlet core_src_dirs = [\n%s]\n" core_src_dirs; + pr "\nlet api_dirs = [\"API\"; \"lib\"]\n"; pr "\nlet plugins_dirs = [\n"; + let plugins = Sys.readdir "plugins" in Array.sort compare plugins; Array.iter @@ -1097,6 +1118,9 @@ let write_configml f = if Sys.is_directory f' && f.[0] <> '.' then pr " %S;\n" f') plugins; pr "]\n"; + + pr "\nlet all_src_dirs = core_src_dirs @ api_dirs @ plugins_dirs\n"; + close_out o; Unix.chmod f 0o444 diff --git a/dev/base_include b/dev/base_include index 608624d06..f9af0696b 100644 --- a/dev/base_include +++ b/dev/base_include @@ -18,10 +18,12 @@ #directory "intf";; #directory "stm";; #directory "vernac";; +#directory "../API";; #directory "+camlp4";; (* lazy solution: add both of camlp4/5 so that *) #directory "+camlp5";; (* Gramext is found in top_printers.ml *) +#load "API.cma";; #use "top_printers.ml";; #use "vm_printers.ml";; @@ -56,6 +58,8 @@ (* Open main files *) +open API +open Grammar_API open Names open Term open Vars @@ -229,7 +233,7 @@ let pf_e gl s = let _ = Flags.in_debugger := false let _ = Flags.in_toplevel := true let _ = Constrextern.set_extern_reference - (fun loc _ r -> Libnames.Qualid (loc,Nametab.shortest_qualid_of_global Idset.empty r));; + (fun ?loc _ r -> Libnames.Qualid (loc,Nametab.shortest_qualid_of_global Idset.empty r));; open Coqloop let go = loop diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh index a6972c950..99ec43e41 100644 --- a/dev/ci/ci-basic-overlay.sh +++ b/dev/ci/ci-basic-overlay.sh @@ -13,8 +13,8 @@ ######################################################################## # UniMath ######################################################################## -: ${UniMath_CI_BRANCH:=coq_makefile2-fix} -: ${UniMath_CI_GITURL:=https://github.com/maximedenes/UniMath.git} +: ${UniMath_CI_BRANCH:=master} +: ${UniMath_CI_GITURL:=https://github.com/UniMath/UniMath.git} ######################################################################## # Unicoq + Metacoq @@ -28,11 +28,11 @@ ######################################################################## # Mathclasses + Corn ######################################################################## -: ${math_classes_CI_BRANCH:=v8.6} -: ${math_classes_CI_GITURL:=https://github.com/math-classes/math-classes.git} +: ${math_classes_CI_BRANCH:=external-bignums} +: ${math_classes_CI_GITURL:=https://github.com/letouzey/math-classes.git} -: ${Corn_CI_BRANCH:=v8.6} -: ${Corn_CI_GITURL:=https://github.com/c-corn/corn.git} +: ${Corn_CI_BRANCH:=external-bignums} +: ${Corn_CI_GITURL:=https://github.com/letouzey/corn.git} ######################################################################## # Iris @@ -46,8 +46,8 @@ ######################################################################## # HoTT ######################################################################## -# Temporal overlay -: ${HoTT_CI_BRANCH:=mz-8.7} +# Temporary overlay +: ${HoTT_CI_BRANCH:=ocaml.4.02.3} : ${HoTT_CI_GITURL:=https://github.com/ejgallego/HoTT.git} # : ${HoTT_CI_BRANCH:=master} # : ${HoTT_CI_GITURL:=https://github.com/HoTT/HoTT.git} @@ -73,38 +73,26 @@ ######################################################################## # CompCert ######################################################################## -: ${CompCert_CI_BRANCH:=master} -: ${CompCert_CI_GITURL:=https://github.com/AbsInt/CompCert.git} +: ${CompCert_CI_BRANCH:=less_init_plugins} +: ${CompCert_CI_GITURL:=https://github.com/letouzey/CompCert.git} ######################################################################## # VST ######################################################################## -: ${VST_CI_BRANCH:=master} -: ${VST_CI_GITURL:=https://github.com/PrincetonUniversity/VST.git} +: ${VST_CI_BRANCH:=less_init_plugins} +: ${VST_CI_GITURL:=https://github.com/letouzey/VST.git} ######################################################################## # fiat_parsers ######################################################################## -: ${fiat_parsers_CI_BRANCH:=master} -: ${fiat_parsers_CI_GITURL:=https://github.com/mit-plv/fiat.git} +: ${fiat_parsers_CI_BRANCH:=trunk__API} +: ${fiat_parsers_CI_GITURL:=https://github.com/matejkosik/fiat.git} ######################################################################## # fiat_crypto ######################################################################## -: ${fiat_crypto_CI_BRANCH:=master} -: ${fiat_crypto_CI_GITURL:=https://github.com/mit-plv/fiat-crypto.git} - -######################################################################## -# bedrock_src -######################################################################## -: ${bedrock_src_CI_BRANCH:=master} -: ${bedrock_src_CI_GITURL:=https://github.com/mit-plv/bedrock.git} - -######################################################################## -# bedrock_facade -######################################################################## -: ${bedrock_facade_CI_BRANCH:=master} -: ${bedrock_facade_CI_GITURL:=https://github.com/mit-plv/bedrock.git} +: ${fiat_crypto_CI_BRANCH:=less_init_plugins} +: ${fiat_crypto_CI_GITURL:=https://github.com/letouzey/fiat-crypto.git} ######################################################################## # formal-topology @@ -113,6 +101,12 @@ : ${formal_topology_CI_GITURL:=https://github.com/bmsherman/topology.git} ######################################################################## +# coq-dpdgraph +######################################################################## +: ${coq_dpdgraph_CI_BRANCH:=coq-trunk} +: ${coq_dpdgraph_CI_GITURL:=https://github.com/Karmaki/coq-dpdgraph.git} + +######################################################################## # CoLoR ######################################################################## : ${Color_CI_SVNURL:=https://scm.gforge.inria.fr/anonscm/svn/color/trunk/color} @@ -127,3 +121,9 @@ ######################################################################## : ${tlc_CI_BRANCH:=master} : ${tlc_CI_GITURL:=https://gforge.inria.fr/git/tlc/tlc.git} + +######################################################################## +# Bignums +######################################################################## +: ${bignums_CI_BRANCH:=master} +: ${bignums_CI_GITURL:=https://github.com/coq/bignums.git} diff --git a/dev/ci/ci-bedrock-facade.sh b/dev/ci/ci-bedrock-facade.sh deleted file mode 100755 index 95cfa3073..000000000 --- a/dev/ci/ci-bedrock-facade.sh +++ /dev/null @@ -1,10 +0,0 @@ -#!/usr/bin/env bash - -ci_dir="$(dirname "$0")" -source ${ci_dir}/ci-common.sh - -bedrock_facade_CI_DIR=${CI_BUILD_DIR}/bedrock-facade - -git_checkout ${bedrock_facade_CI_BRANCH} ${bedrock_facade_CI_GITURL} ${bedrock_facade_CI_DIR} - -( cd ${bedrock_facade_CI_DIR} && make -j ${NJOBS} facade ) diff --git a/dev/ci/ci-bedrock-src.sh b/dev/ci/ci-bedrock-src.sh deleted file mode 100755 index 532611d4b..000000000 --- a/dev/ci/ci-bedrock-src.sh +++ /dev/null @@ -1,10 +0,0 @@ -#!/usr/bin/env bash - -ci_dir="$(dirname "$0")" -source ${ci_dir}/ci-common.sh - -bedrock_src_CI_DIR=${CI_BUILD_DIR}/bedrock-src - -git_checkout ${bedrock_src_CI_BRANCH} ${bedrock_src_CI_GITURL} ${bedrock_src_CI_DIR} - -( cd ${bedrock_src_CI_DIR} && make -j ${NJOBS} src ) diff --git a/dev/ci/ci-bignums.sh b/dev/ci/ci-bignums.sh new file mode 100755 index 000000000..ff5935d4c --- /dev/null +++ b/dev/ci/ci-bignums.sh @@ -0,0 +1,16 @@ +#!/usr/bin/env bash + +ci_dir="$(dirname "$0")" + +# This script could be included inside other ones +# Let's avoid to source ci-common twice in this case +if [ -z "${CI_BUILD_DIR}"]; +then + source ${ci_dir}/ci-common.sh +fi + +bignums_CI_DIR=${CI_BUILD_DIR}/Bignums + +git_checkout ${bignums_CI_BRANCH} ${bignums_CI_GITURL} ${bignums_CI_DIR} + +( cd ${bignums_CI_DIR} && make -j ${NJOBS} && make install) diff --git a/dev/ci/ci-color.sh b/dev/ci/ci-color.sh index 3f0716511..a0a4e0749 100755 --- a/dev/ci/ci-color.sh +++ b/dev/ci/ci-color.sh @@ -5,6 +5,31 @@ source ${ci_dir}/ci-common.sh Color_CI_DIR=${CI_BUILD_DIR}/color +# Setup Bignums + +source ${ci_dir}/ci-bignums.sh + +# Compiles CoLoR + svn checkout ${Color_CI_SVNURL} ${Color_CI_DIR} +sed -i -e "s/From Coq Require Import BigN/From Bignums Require Import BigN/" ${Color_CI_DIR}/Util/*/*.v +sed -i -e "s/From Coq Require Export BigN/From Bignums Require Export BigN/" ${Color_CI_DIR}/Util/*/*.v +sed -i -e "s/From Coq Require Import BigZ/From Bignums Require Import BigZ/" ${Color_CI_DIR}/Util/*/*.v +sed -i -e "s/From Coq Require Export BigZ/From Bignums Require Export BigZ/" ${Color_CI_DIR}/Util/*/*.v + +# Adapt to PR #220 (FunInd not loaded in Prelude anymore) +sed -i -e "15i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/basis/ordered_set.v +sed -i -e "8i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/examples/cime_trace/equational_extension.v +sed -i -e "6i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/examples/cime_trace/more_list_extention.v +sed -i -e "6i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/examples/cime_trace/ring_extention.v +sed -i -e "27i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/list_extensions/dickson.v +sed -i -e "26i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/list_extensions/list_permut.v +sed -i -e "23i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/list_extensions/list_set.v +sed -i -e "25i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/list_extensions/list_sort.v +sed -i -e "21i From Coq Require Import FunInd." ${Color_CI_DIR}/Coccinelle/list_extensions/more_list.v +sed -i -e "21i From Coq Require Import FunInd." ${Color_CI_DIR}/Util/List/ListUtil.v +sed -i -e "17i From Coq Require Import FunInd." ${Color_CI_DIR}/Util/Multiset/MultisetOrder.v +sed -i -e "13i From Coq Require Import FunInd." ${Color_CI_DIR}/Util/Set/SetUtil.v + ( cd ${Color_CI_DIR} && make -j ${NJOBS} ) diff --git a/dev/ci/ci-coq-dpdgraph.sh b/dev/ci/ci-coq-dpdgraph.sh new file mode 100755 index 000000000..e8018158b --- /dev/null +++ b/dev/ci/ci-coq-dpdgraph.sh @@ -0,0 +1,10 @@ +#!/usr/bin/env bash + +ci_dir="$(dirname "$0")" +source ${ci_dir}/ci-common.sh + +coq_dpdgraph_CI_DIR=${CI_BUILD_DIR}/coq-dpdgraph + +git_checkout ${coq_dpdgraph_CI_BRANCH} ${coq_dpdgraph_CI_GITURL} ${coq_dpdgraph_CI_DIR} + +( cd ${coq_dpdgraph_CI_DIR} && autoconf && ./configure && make -j ${NJOBS} && make tests && (make tests | tee tmp.log) && (if grep DIFFERENCES tmp.log ; then exit 1 ; else exit 0 ; fi) ) diff --git a/dev/ci/ci-fiat-parsers.sh b/dev/ci/ci-fiat-parsers.sh index a0cb008a3..2095245eb 100755 --- a/dev/ci/ci-fiat-parsers.sh +++ b/dev/ci/ci-fiat-parsers.sh @@ -7,4 +7,4 @@ fiat_parsers_CI_DIR=${CI_BUILD_DIR}/fiat git_checkout ${fiat_parsers_CI_BRANCH} ${fiat_parsers_CI_GITURL} ${fiat_parsers_CI_DIR} -( cd ${fiat_parsers_CI_DIR} && make -j ${NJOBS} parsers && make -j ${NJOBS} fiat-core ) +( cd ${fiat_parsers_CI_DIR} && make -j ${NJOBS} parsers parsers-examples && make -j ${NJOBS} fiat-core ) diff --git a/dev/ci/ci-formal-topology.sh b/dev/ci/ci-formal-topology.sh index ecb36349f..64b78c039 100755 --- a/dev/ci/ci-formal-topology.sh +++ b/dev/ci/ci-formal-topology.sh @@ -9,6 +9,10 @@ Corn_CI_DIR=${CI_BUILD_DIR}/corn formal_topology_CI_DIR=${CI_BUILD_DIR}/formal-topology +# Setup Bignums + +source ${ci_dir}/ci-bignums.sh + # Setup Math-Classes git_checkout ${math_classes_CI_BRANCH} ${math_classes_CI_GITURL} ${math_classes_CI_DIR} diff --git a/dev/ci/ci-math-classes.sh b/dev/ci/ci-math-classes.sh index beb75773b..46581c638 100755 --- a/dev/ci/ci-math-classes.sh +++ b/dev/ci/ci-math-classes.sh @@ -7,6 +7,10 @@ math_classes_CI_DIR=${CI_BUILD_DIR}/math-classes Corn_CI_DIR=${CI_BUILD_DIR}/corn +# Setup Bignums + +source ${ci_dir}/ci-bignums.sh + # Setup Math-Classes git_checkout ${math_classes_CI_BRANCH} ${math_classes_CI_GITURL} ${math_classes_CI_DIR} diff --git a/dev/ci/ci-sf.sh b/dev/ci/ci-sf.sh index 7d23ccad9..23ef41d2d 100755 --- a/dev/ci/ci-sf.sh +++ b/dev/ci/ci-sf.sh @@ -7,6 +7,8 @@ source ${ci_dir}/ci-common.sh wget ${sf_CI_TARURL} tar xvfz sf.tgz +sed -i.bak '15i From Coq Require Extraction.' sf/Extraction.v + ( cd sf && sed -i.bak 's/(K,N)/((K,N))/' LibTactics.v && make clean && make -j ${NJOBS} ) diff --git a/dev/ci/ci-user-overlay.sh b/dev/ci/ci-user-overlay.sh index 195ede6d0..b242ce3bd 100644 --- a/dev/ci/ci-user-overlay.sh +++ b/dev/ci/ci-user-overlay.sh @@ -30,3 +30,25 @@ if [ $TRAVIS_PULL_REQUEST == "669" ] || [ $TRAVIS_BRANCH == "ssr-merge" ]; then mathcomp_CI_GITURL=https://github.com/maximedenes/math-comp.git fi +echo "DEBUG: ci-user-overlay.sh 0" +if [ $TRAVIS_PULL_REQUEST = "707" ] || [ $TRAVIS_BRANCH == "trunk__API__coq_makefile" ]; then + echo "DEBUG: ci-user-overlay.sh 1" + fiat_parsers_CI_BRANCH=trunk__API + fiat_parsers_CI_GITURL=https://github.com/matejkosik/fiat.git +fi + +if [ $TRAVIS_PULL_REQUEST == "498" ] || [ $TRAVIS_BRANCH == "outsource-bignums" ]; then + math_classes_CI_BRANCH=external-bignums + math_classes_CI_GITURL=https://github.com/letouzey/math-classes.git + Corn_CI_BRANCH=external-bignums + Corn_CI_GITURL=https://github.com/letouzey/corn.git +fi + +if [ $TRAVIS_PULL_REQUEST == "220" ] || [ $TRAVIS_BRANCH == "less_init_plugins" ]; then + CompCert_CI_BRANCH=less_init_plugins + CompCert_CI_GITURL=https://github.com/letouzey/CompCert.git + VST_CI_BRANCH=less_init_plugins + VST_CI_GITURL=https://github.com/letouzey/VST.git + fiat_crypto_CI_BRANCH=less_init_plugins + fiat_crypto_CI_GITURL=https://github.com/letouzey/fiat-crypto.git +fi diff --git a/dev/core.dbg b/dev/core.dbg index 6acdd0152..71d06cdb0 100644 --- a/dev/core.dbg +++ b/dev/core.dbg @@ -17,4 +17,6 @@ load_printer vernac.cma load_printer stm.cma load_printer toplevel.cma load_printer highparsing.cma +load_printer intf.cma +load_printer API.cma load_printer ltac_plugin.cmo diff --git a/dev/doc/changes.txt b/dev/doc/changes.txt index bcda4ff50..0728608f3 100644 --- a/dev/doc/changes.txt +++ b/dev/doc/changes.txt @@ -10,6 +10,16 @@ will fail to compile now. They should switch to `Bytes.t` * ML API * +Added two functions for declaring hooks to be executed in reduction +functions when some given constants are traversed: + + declare_reduction_effect: to declare a hook to be applied when some + constant are visited during the execution of some reduction functions + (primarily cbv). + + set_reduction_effect: to declare a constant on which a given effect + hook should be called. + We renamed the following functions: Context.Rel.Declaration.fold -> Context.Rel.Declaration.fold_constr @@ -144,6 +154,9 @@ In Coqlib / reference location: - The tclWEAK_PROGRESS and tclNOTSAMEGOAL tacticals were removed. Their usecase was very specific. Use tclPROGRESS instead. +- The unsafe flag of the Refine.refine function and its variants has been + renamed and dualized into typecheck and has been made mandatory. + ** Ltac API ** Many Ltac specific API has been moved in its own ltac/ folder. Amongst other diff --git a/dev/doc/proof-engine.md b/dev/doc/proof-engine.md index db69b08a2..8f96ac223 100644 --- a/dev/doc/proof-engine.md +++ b/dev/doc/proof-engine.md @@ -42,14 +42,13 @@ goal holes thanks to the `Refine` module, and in particular to the `Refine.refine` primitive. ```ocaml -val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic -(** In [refine ?unsafe t], [t] is a term with holes under some +val refine : typecheck:bool -> Constr.t Sigma.run -> unit tactic +(** In [refine typecheck t], [t] is a term with holes under some [evar_map] context. The term [t] is used as a partial solution for the current goal (refine is a goal-dependent tactic), the new holes created by [t] become the new subgoals. Exceptions raised during the interpretation of [t] are caught and result in - tactic failures. If [unsafe] is [false] (default is [true]) [t] is - type-checked beforehand. *) + tactic failures. If [typecheck] is [true] [t] is type-checked beforehand. *) ``` In a first approximation, we can think of `'a Sigma.run` as diff --git a/dev/doc/setup.txt b/dev/doc/setup.txt index 1b016a4e2..0c6d3ee80 100644 --- a/dev/doc/setup.txt +++ b/dev/doc/setup.txt @@ -12,7 +12,7 @@ How to compile Coq Getting build dependencies: - sudo apt-get install make opam git mercurial darcs + sudo apt-get install make opam git opam init --comp 4.02.3 # Then follow the advice displayed at the end as how to update your ~/.bashrc and ~/.ocamlinit files. @@ -41,7 +41,7 @@ Building coqtop: cd ~/git/coq git checkout trunk make distclean - ./configure -annotate -with-doc no -local -debug -usecamlp5 + ./configure -annotate -local make clean make -j4 coqide printers @@ -49,8 +49,6 @@ The "-annotate" option is essential when one wants to use Merlin. The "-local" option is useful if one wants to run the coqtop and coqide binaries without running make install -The "-debug" option is essential if one wants to use ocamldebug with the coqtop binary. - Then check if - bin/coqtop - bin/coqide @@ -60,7 +58,7 @@ behave as expected. A note about rlwrap ------------------- -Running "coqtop" under "rlwrap" is possible, but there is a catch. If you try: +Running "coqtop" under "rlwrap" is possible, but (on Debian) there is a catch. If you try: cd ~/git/coq rlwrap bin/coqtop diff --git a/dev/ocamldebug-coq.run b/dev/ocamldebug-coq.run index 3850c05fd..f4799f7b2 100644 --- a/dev/ocamldebug-coq.run +++ b/dev/ocamldebug-coq.run @@ -23,6 +23,7 @@ exec $OCAMLDEBUG \ -I $COQTOP/pretyping -I $COQTOP/parsing -I $COQTOP/vernac \ -I $COQTOP/interp -I $COQTOP/proofs -I $COQTOP/tactics -I $COQTOP/stm \ -I $COQTOP/toplevel -I $COQTOP/dev -I $COQTOP/config -I $COQTOP/ltac \ + -I $COQTOP/API \ -I $COQTOP/plugins/cc -I $COQTOP/plugins/dp \ -I $COQTOP/plugins/extraction -I $COQTOP/plugins/field \ -I $COQTOP/plugins/firstorder -I $COQTOP/plugins/fourier \ diff --git a/dev/tools/Makefile.devel b/dev/tools/Makefile.devel index 8dcc70cf7..ffdb1bdca 100644 --- a/dev/tools/Makefile.devel +++ b/dev/tools/Makefile.devel @@ -5,7 +5,7 @@ TOPDIR=. BASEDIR= -SOURCEDIRS=lib kernel library pretyping parsing proofs tactics toplevel +SOURCEDIRS=lib kernel library pretyping parsing proofs tactics toplevel API default: usage noargument diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 07a47c8b7..6ae5125f6 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -8,6 +8,7 @@ (* Printers for the ocaml toplevel. *) +open API open Util open Pp open Names @@ -36,7 +37,7 @@ let pp x = Pp.pp_with !Topfmt.std_ft x let ppfuture kx = pp (Future.print (fun _ -> str "_") kx) (* name printers *) -let ppid id = pp (pr_id id) +let ppid id = pp (Id.print id) let pplab l = pp (pr_lab l) let ppmbid mbid = pp (str (MBId.debug_to_string mbid)) let ppdir dir = pp (pr_dirpath dir) @@ -78,12 +79,12 @@ let ppbigint n = pp (str (Bigint.to_string n));; let prset pr l = str "[" ++ hov 0 (prlist_with_sep spc pr l) ++ str "]" let ppintset l = pp (prset int (Int.Set.elements l)) -let ppidset l = pp (prset pr_id (Id.Set.elements l)) +let ppidset l = pp (prset Id.print (Id.Set.elements l)) let prset' pr l = str "[" ++ hov 0 (prlist_with_sep pr_comma pr l) ++ str "]" let pridmap pr l = - let pr (id,b) = pr_id id ++ str "=>" ++ pr id b in + let pr (id,b) = Id.print id ++ str "=>" ++ pr id b in prset' pr (Id.Map.fold (fun a b l -> (a,b)::l) l []) let ppidmap pr l = pp (pridmap pr l) @@ -94,10 +95,10 @@ let ppevarsubst = ppidmap (fun id0 -> prset (fun (c,copt,id) -> (match copt with None -> mt () | Some c -> spc () ++ str "<expanded: " ++ Termops.print_constr (EConstr.of_constr c) ++ str">") ++ (if id = id0 then mt () - else spc () ++ str "<canonical: " ++ pr_id id ++ str ">")))) + else spc () ++ str "<canonical: " ++ Id.print id ++ str ">")))) -let prididmap = pridmap (fun _ -> pr_id) -let ppididmap = ppidmap (fun _ -> pr_id) +let prididmap = pridmap (fun _ -> Id.print) +let ppididmap = ppidmap (fun _ -> Id.print) let prconstrunderbindersidmap = pridmap (fun _ (l,c) -> hov 1 (str"[" ++ prlist_with_sep spc Id.print l ++ str"]") @@ -131,15 +132,15 @@ let safe_pr_global = function int i ++ str ")") | ConstructRef ((kn,i),j) -> pp (str "INDREF(" ++ debug_pr_mind kn ++ str "," ++ int i ++ str "," ++ int j ++ str ")") - | VarRef id -> pp (str "VARREF(" ++ pr_id id ++ str ")") + | VarRef id -> pp (str "VARREF(" ++ Id.print id ++ str ")") let ppglobal x = try pp(pr_global x) with _ -> safe_pr_global x let ppconst (sp,j) = - pp (str"#" ++ pr_kn sp ++ str"=" ++ pr_lconstr j.uj_val) + pp (str"#" ++ KerName.print sp ++ str"=" ++ pr_lconstr j.uj_val) let ppvar ((id,a)) = - pp (str"#" ++ pr_id id ++ str":" ++ pr_lconstr a) + pp (str"#" ++ Id.print id ++ str":" ++ pr_lconstr a) let genppj f j = let (c,t) = f j in (c ++ str " : " ++ t) @@ -492,6 +493,7 @@ VERNAC COMMAND EXTEND PrintConstr END *) +open Grammar_API open Genarg open Stdarg open Egramml @@ -536,21 +538,21 @@ let encode_path ?loc prefix mpdir suffix id = let dir = match mpdir with | None -> [] | Some (mp,dir) -> - (DirPath.repr (dirpath_of_string (string_of_mp mp))@ + (DirPath.repr (dirpath_of_string (ModPath.to_string mp))@ DirPath.repr dir) in Qualid (Loc.tag ?loc @@ make_qualid (DirPath.make (List.rev (Id.of_string prefix::dir@suffix))) id) let raw_string_of_ref ?loc _ = function | ConstRef cst -> - let (mp,dir,id) = repr_con cst in + let (mp,dir,id) = Constant.repr3 cst in encode_path ?loc "CST" (Some (mp,dir)) [] (Label.to_id id) | IndRef (kn,i) -> - let (mp,dir,id) = repr_mind kn in + let (mp,dir,id) = MutInd.repr3 kn in encode_path ?loc "IND" (Some (mp,dir)) [Label.to_id id] (Id.of_string ("_"^string_of_int i)) | ConstructRef ((kn,i),j) -> - let (mp,dir,id) = repr_mind kn in + let (mp,dir,id) = MutInd.repr3 kn in encode_path ?loc "CSTR" (Some (mp,dir)) [Label.to_id id;Id.of_string ("_"^string_of_int i)] (Id.of_string ("_"^string_of_int j)) @@ -559,14 +561,14 @@ let raw_string_of_ref ?loc _ = function let short_string_of_ref ?loc _ = function | VarRef id -> Ident (Loc.tag ?loc id) - | ConstRef cst -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (repr_con cst))) - | IndRef (kn,0) -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (repr_mind kn))) + | ConstRef cst -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (Constant.repr3 cst))) + | IndRef (kn,0) -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (MutInd.repr3 kn))) | IndRef (kn,i) -> - encode_path ?loc "IND" None [Label.to_id (pi3 (repr_mind kn))] + encode_path ?loc "IND" None [Label.to_id (pi3 (MutInd.repr3 kn))] (Id.of_string ("_"^string_of_int i)) | ConstructRef ((kn,i),j) -> encode_path ?loc "CSTR" None - [Label.to_id (pi3 (repr_mind kn));Id.of_string ("_"^string_of_int i)] + [Label.to_id (pi3 (MutInd.repr3 kn));Id.of_string ("_"^string_of_int i)] (Id.of_string ("_"^string_of_int j)) (* Anticipate that printers can be used from ocamldebug and that diff --git a/dev/vm_printers.ml b/dev/vm_printers.ml index afa94a63e..be6b914b6 100644 --- a/dev/vm_printers.ml +++ b/dev/vm_printers.ml @@ -1,3 +1,4 @@ +open API open Format open Term open Names diff --git a/doc/refman/Extraction.tex b/doc/refman/Extraction.tex index 01dbcfb1c..fa3d61b1c 100644 --- a/doc/refman/Extraction.tex +++ b/doc/refman/Extraction.tex @@ -19,6 +19,12 @@ be used (abusively) to refer to any of the three. %% the one in previous versions of \Coq: there is no more %% an explicit toplevel for the language (formerly called \textsc{Fml}). +Before using any of the commands or options described in this chapter, +the extraction framework should first be loaded explicitly +via {\tt Require Extraction}. Note that in earlier versions of Coq, these +commands and options were directly available without any preliminary +{\tt Require}. + \asection{Generating ML code} \comindex{Extraction} \comindex{Recursive Extraction} @@ -501,6 +507,7 @@ We can now extract this program to \ocaml: Reset Initial. \end{coq_eval} \begin{coq_example} +Require Extraction. Require Import Euclid Wf_nat. Extraction Inline gt_wf_rec lt_wf_rec induction_ltof2. Recursive Extraction eucl_dev. diff --git a/doc/refman/RefMan-ext.tex b/doc/refman/RefMan-ext.tex index 6dd0ddf81..939fc87a6 100644 --- a/doc/refman/RefMan-ext.tex +++ b/doc/refman/RefMan-ext.tex @@ -721,18 +721,20 @@ a given type. See Section~\ref{Show}. \section{Advanced recursive functions} -The \emph{experimental} command +The following \emph{experimental} command is available +when the {\tt FunInd} library has been loaded via {\tt Require Import FunInd}: \begin{center} \texttt{Function {\ident} {\binder$_1$}\ldots{\binder$_n$} \{decrease\_annot\} : type$_0$ := \term$_0$} \comindex{Function} \label{Function} \end{center} -can be seen as a generalization of {\tt Fixpoint}. It is actually a -wrapper for several ways of defining a function \emph{and other useful +This command can be seen as a generalization of {\tt Fixpoint}. It is actually +a wrapper for several ways of defining a function \emph{and other useful related objects}, namely: an induction principle that reflects the recursive structure of the function (see \ref{FunInduction}), and its -fixpoint equality. The meaning of this +fixpoint equality. + The meaning of this declaration is to define a function {\it ident}, similarly to {\tt Fixpoint}. Like in {\tt Fixpoint}, the decreasing argument must be given (unless the function is not recursive), but it must not diff --git a/doc/refman/RefMan-pro.tex b/doc/refman/RefMan-pro.tex index 0760d716e..b66659dc8 100644 --- a/doc/refman/RefMan-pro.tex +++ b/doc/refman/RefMan-pro.tex @@ -427,22 +427,6 @@ This command displays the current goals. This tactics script may contain some holes (subgoals not yet proved). They are printed under the form \verb!<Your Tactic Text here>!. -%% \item {\tt Show Tree.}\comindex{Show Tree}\\ -%% This command can be seen as a more structured way of -%% displaying the state of the proof than that -%% provided by {\tt Show Script}. Instead of just giving -%% the list of tactics that have been applied, it -%% shows the derivation tree constructed by then. -%% Each node of the tree contains the conclusion -%% of the corresponding sub-derivation (i.e. a -%% goal with its corresponding local context) and -%% the tactic that has generated all the -%% sub-derivations. The leaves of this tree are -%% the goals which still remain to be proved. - -%\item {\tt Show Node}\comindex{Show Node}\\ -% Not yet documented - \item {\tt Show Proof.}\comindex{Show Proof}\\ It displays the proof term generated by the tactics that have been applied. diff --git a/doc/refman/RefMan-sch.tex b/doc/refman/RefMan-sch.tex index 53aa6b86a..d3719bed4 100644 --- a/doc/refman/RefMan-sch.tex +++ b/doc/refman/RefMan-sch.tex @@ -196,8 +196,10 @@ Check tree_forest_mutind. The {\tt Functional Scheme} command is a high-level experimental tool for generating automatically induction principles -corresponding to (possibly mutually recursive) functions. Its -syntax follows the schema: +corresponding to (possibly mutually recursive) functions. +First, it must be made available via {\tt Require Import FunInd}. + Its +syntax then follows the schema: \begin{quote} {\tt Functional Scheme {\ident$_1$} := Induction for \ident'$_1$ Sort {\sort$_1$} \\ with\\ @@ -319,6 +321,7 @@ of a tree or a forest. Note that we use \texttt{Function} which generally produces better principles. \begin{coq_example*} +Require Import FunInd. Function tree_size (t:tree) : nat := match t with | node A f => S (forest_size f) diff --git a/doc/refman/RefMan-tac.tex b/doc/refman/RefMan-tac.tex index 253eb7f01..2bab04e90 100644 --- a/doc/refman/RefMan-tac.tex +++ b/doc/refman/RefMan-tac.tex @@ -2113,13 +2113,15 @@ The tactic \texttt{functional induction} performs case analysis and induction following the definition of a function. It makes use of a principle generated by \texttt{Function} (see Section~\ref{Function}) or \texttt{Functional Scheme} -(see Section~\ref{FunScheme}). +(see Section~\ref{FunScheme}). Note that this tactic is only available +after a {\tt Require Import FunInd}. \begin{coq_eval} Reset Initial. Import Nat. \end{coq_eval} \begin{coq_example} +Require Import FunInd. Functional Scheme minus_ind := Induction for minus Sort Prop. Check minus_ind. Lemma le_minus (n m:nat) : n - m <= n. @@ -4797,6 +4799,7 @@ that performs inversion on hypothesis {\ident} of the form \texttt{\qualid\ \term$_1$\dots\term$_n$\ = \term} or \texttt{\term\ = \qualid\ \term$_1$\dots\term$_n$} where \qualid\ must have been defined using \texttt{Function} (see Section~\ref{Function}). +Note that this tactic is only available after a {\tt Require Import FunInd}. \begin{ErrMsgs} \item \errindex{Hypothesis {\ident} must contain at least one Function} diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template index aeb0de48a..48f82f2d9 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> 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> 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> @@ -618,7 +589,6 @@ through the <tt>Require Import</tt> command.</p> </dt> <dd> theories/Compat/AdmitAxiom.v - theories/Compat/Coq84.v theories/Compat/Coq85.v theories/Compat/Coq86.v </dd> diff --git a/engine/namegen.ml b/engine/namegen.ml index 5bd62273c..783085654 100644 --- a/engine/namegen.ml +++ b/engine/namegen.ml @@ -412,13 +412,12 @@ let rename_bound_vars_as_displayed sigma avoid env c = let h_based_elimination_names = ref false -let use_h_based_elimination_names () = - !h_based_elimination_names && Flags.version_strictly_greater Flags.V8_4 +let use_h_based_elimination_names () = !h_based_elimination_names open Goptions let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "use of \"H\"-based proposition names in elimination tactics"; optkey = ["Standard";"Proposition";"Elimination";"Names"]; optread = (fun () -> !h_based_elimination_names); diff --git a/grammar/argextend.mlp b/grammar/argextend.mlp index 5068ba8a6..36b9d612a 100644 --- a/grammar/argextend.mlp +++ b/grammar/argextend.mlp @@ -46,17 +46,17 @@ let make_act loc act pil = make (List.rev pil) let make_prod_item = function - | ExtTerminal s -> <:expr< Extend.Atoken (CLexer.terminal $mlexpr_of_string s$) >> + | ExtTerminal s -> <:expr< Grammar_API.Extend.Atoken (Grammar_API.CLexer.terminal $mlexpr_of_string s$) >> | ExtNonTerminal (g, _) -> let base s = <:expr< $lid:s$ >> in mlexpr_of_prod_entry_key base g let rec make_prod = function -| [] -> <:expr< Extend.Stop >> -| item :: prods -> <:expr< Extend.Next $make_prod prods$ $make_prod_item item$ >> +| [] -> <:expr< Grammar_API.Extend.Stop >> +| item :: prods -> <:expr< Grammar_API.Extend.Next $make_prod prods$ $make_prod_item item$ >> let make_rule loc (prods,act) = - <:expr< Extend.Rule $make_prod (List.rev prods)$ $make_act loc act prods$ >> + <:expr< Grammar_API.Extend.Rule $make_prod (List.rev prods)$ $make_act loc act prods$ >> let is_ident x = function | <:expr< $lid:s$ >> -> (s : string) = x @@ -67,7 +67,7 @@ let make_extend loc s cl wit = match cl with (** Special handling of identity arguments by not redeclaring an entry *) <:str_item< value $lid:s$ = - let () = Pcoq.register_grammar $wit$ $lid:e$ in + let () = Grammar_API.Pcoq.register_grammar $wit$ $lid:e$ in $lid:e$ >> | _ -> @@ -75,8 +75,8 @@ let make_extend loc s cl wit = match cl with let rules = mlexpr_of_list (make_rule loc) (List.rev cl) in <:str_item< value $lid:s$ = - let $lid:s$ = Pcoq.create_generic_entry Pcoq.utactic $se$ (Genarg.rawwit $wit$) in - let () = Pcoq.grammar_extend $lid:s$ None (None, [(None, None, $rules$)]) in + let $lid:s$ = Grammar_API.Pcoq.create_generic_entry Grammar_API.Pcoq.utactic $se$ (Genarg.rawwit $wit$) in + let () = Grammar_API.Pcoq.grammar_extend $lid:s$ None (None, [(None, None, $rules$)]) in $lid:s$ >> let warning_redundant prefix s = @@ -127,7 +127,7 @@ let declare_tactic_argument loc s (typ, f, g, h) cl = begin match globtyp with | None -> let typ = match globtyp with None -> ExtraArgType s | Some typ -> typ in - <:expr< fun ist v -> Ftactic.return (Geninterp.Val.inject (Geninterp.val_tag $make_topwit loc typ$) v) >> + <:expr< fun ist v -> API.Ftactic.return (API.Geninterp.Val.inject (API.Geninterp.val_tag $make_topwit loc typ$) v) >> | Some globtyp -> <:expr< fun ist x -> Tacinterp.interp_genarg ist (Genarg.in_gen $make_globwit loc globtyp$ x) >> @@ -137,10 +137,10 @@ let declare_tactic_argument loc s (typ, f, g, h) cl = let typ = match globtyp with None -> ExtraArgType s | Some typ -> typ in <:expr< let f = $lid:f$ in - fun ist v -> Ftactic.nf_enter (fun gl -> - let (sigma, v) = Tacmach.New.of_old (fun gl -> f ist gl v) gl in - let v = Geninterp.Val.inject (Geninterp.val_tag $make_topwit loc typ$) v in - Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Ftactic.return v) + fun ist v -> API.Ftactic.nf_enter (fun gl -> + let (sigma, v) = API.Tacmach.New.of_old (fun gl -> f ist gl v) gl in + let v = API.Geninterp.Val.inject (API.Geninterp.val_tag $make_topwit loc typ$) v in + API.Proofview.tclTHEN (API.Proofview.Unsafe.tclEVARS sigma) (API.Ftactic.return v) ) >> in let subst = match h with @@ -156,15 +156,15 @@ let declare_tactic_argument loc s (typ, f, g, h) cl = | Some f -> <:expr< $lid:f$>> in let dyn = match typ with | None -> <:expr< None >> - | Some typ -> <:expr< Some (Geninterp.val_tag $make_topwit loc typ$) >> + | Some typ -> <:expr< Some (API.Geninterp.val_tag $make_topwit loc typ$) >> in let wit = <:expr< $lid:"wit_"^s$ >> in declare_str_items loc [ <:str_item< value ($lid:"wit_"^s$) = Genarg.make0 $se$ >>; - <:str_item< Genintern.register_intern0 $wit$ $glob$ >>; - <:str_item< Genintern.register_subst0 $wit$ $subst$ >>; - <:str_item< Geninterp.register_interp0 $wit$ $interp$ >>; - <:str_item< Geninterp.register_val0 $wit$ $dyn$ >>; + <:str_item< Grammar_API.Genintern.register_intern0 $wit$ $glob$ >>; + <:str_item< Grammar_API.Genintern.register_subst0 $wit$ $subst$ >>; + <:str_item< API.Geninterp.register_interp0 $wit$ $interp$ >>; + <:str_item< API.Geninterp.register_val0 $wit$ $dyn$ >>; make_extend loc s cl wit; <:str_item< do { Pptactic.declare_extra_genarg_pprule diff --git a/grammar/q_util.mlp b/grammar/q_util.mlp index 87262e1c8..1c2009ece 100644 --- a/grammar/q_util.mlp +++ b/grammar/q_util.mlp @@ -57,23 +57,23 @@ let mlexpr_of_option f = function | Some e -> <:expr< Some $f e$ >> let mlexpr_of_name f = function - | None -> <:expr< Anonymous >> - | Some e -> <:expr< Name $f e$ >> + | None -> <:expr< API.Names.Name.Anonymous >> + | Some e -> <:expr< API.Names.Name.Name $f e$ >> -let symbol_of_string s = <:expr< Extend.Atoken (CLexer.terminal $str:s$) >> +let symbol_of_string s = <:expr< Grammar_API.Extend.Atoken (Grammar_API.CLexer.terminal $str:s$) >> let rec mlexpr_of_prod_entry_key f = function - | Ulist1 s -> <:expr< Extend.Alist1 $mlexpr_of_prod_entry_key f s$ >> - | Ulist1sep (s,sep) -> <:expr< Extend.Alist1sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >> - | Ulist0 s -> <:expr< Extend.Alist0 $mlexpr_of_prod_entry_key f s$ >> - | Ulist0sep (s,sep) -> <:expr< Extend.Alist0sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >> - | Uopt s -> <:expr< Extend.Aopt $mlexpr_of_prod_entry_key f s$ >> - | Uentry e -> <:expr< Extend.Aentry $f e$ >> + | Ulist1 s -> <:expr< Grammar_API.Extend.Alist1 $mlexpr_of_prod_entry_key f s$ >> + | Ulist1sep (s,sep) -> <:expr< Grammar_API.Extend.Alist1sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >> + | Ulist0 s -> <:expr< Grammar_API.Extend.Alist0 $mlexpr_of_prod_entry_key f s$ >> + | Ulist0sep (s,sep) -> <:expr< Grammar_API.Extend.Alist0sep $mlexpr_of_prod_entry_key f s$ $symbol_of_string sep$ >> + | Uopt s -> <:expr< Grammar_API.Extend.Aopt $mlexpr_of_prod_entry_key f s$ >> + | Uentry e -> <:expr< Grammar_API.Extend.Aentry ($f e$) >> | Uentryl (e, l) -> (** Keep in sync with Pcoq! *) assert (e = "tactic"); - if l = 5 then <:expr< Extend.Aentry (Pltac.binder_tactic) >> - else <:expr< Extend.Aentryl (Pltac.tactic_expr) $mlexpr_of_int l$ >> + if l = 5 then <:expr< Grammar_API.Extend.Aentry Pltac.binder_tactic >> + else <:expr< Grammar_API.Extend.Aentryl (Pltac.tactic_expr) $mlexpr_of_int l$ >> let rec type_of_user_symbol = function | Ulist1 s | Ulist1sep (s, _) | Ulist0 s | Ulist0sep (s, _) -> diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp index 8e3dccf47..8f3f7a9de 100644 --- a/grammar/tacextend.mlp +++ b/grammar/tacextend.mlp @@ -25,7 +25,7 @@ let plugin_name = <:expr< __coq_plugin_name >> let mlexpr_of_ident id = (** Workaround for badly-designed generic arguments lacking a closure *) let id = "$" ^ id in - <:expr< Names.Id.of_string_soft $str:id$ >> + <:expr< API.Names.Id.of_string_soft $str:id$ >> let rec make_patt = function | [] -> <:patt< [] >> @@ -57,18 +57,18 @@ let make_fun_clauses loc s l = let get_argt e = <:expr< (fun e -> match e with [ Genarg.ExtraArg tag -> tag | _ -> assert False ]) $e$ >> let rec mlexpr_of_symbol = function -| Ulist1 s -> <:expr< Extend.Ulist1 $mlexpr_of_symbol s$ >> -| Ulist1sep (s,sep) -> <:expr< Extend.Ulist1sep $mlexpr_of_symbol s$ $str:sep$ >> -| Ulist0 s -> <:expr< Extend.Ulist0 $mlexpr_of_symbol s$ >> -| Ulist0sep (s,sep) -> <:expr< Extend.Ulist0sep $mlexpr_of_symbol s$ $str:sep$ >> -| Uopt s -> <:expr< Extend.Uopt $mlexpr_of_symbol s$ >> +| Ulist1 s -> <:expr< Grammar_API.Extend.Ulist1 $mlexpr_of_symbol s$ >> +| Ulist1sep (s,sep) -> <:expr< Grammar_API.Extend.Ulist1sep $mlexpr_of_symbol s$ $str:sep$ >> +| Ulist0 s -> <:expr< Grammar_API.Extend.Ulist0 $mlexpr_of_symbol s$ >> +| Ulist0sep (s,sep) -> <:expr< Grammar_API.Extend.Ulist0sep $mlexpr_of_symbol s$ $str:sep$ >> +| Uopt s -> <:expr< Grammar_API.Extend.Uopt $mlexpr_of_symbol s$ >> | Uentry e -> let arg = get_argt <:expr< $lid:"wit_"^e$ >> in - <:expr< Extend.Uentry (Genarg.ArgT.Any $arg$) >> + <:expr< Grammar_API.Extend.Uentry (Genarg.ArgT.Any $arg$) >> | Uentryl (e, l) -> assert (e = "tactic"); let arg = get_argt <:expr< Tacarg.wit_tactic >> in - <:expr< Extend.Uentryl (Genarg.ArgT.Any $arg$) $mlexpr_of_int l$>> + <:expr< Grammar_API.Extend.Uentryl (Genarg.ArgT.Any $arg$) $mlexpr_of_int l$>> let make_prod_item = function | ExtTerminal s -> <:expr< Tacentries.TacTerm $str:s$ >> @@ -113,12 +113,12 @@ let declare_tactic loc tacname ~level classification clause = match clause with the ML tactic retrieves its arguments in the [ist] environment instead. This is the rôle of the [lift_constr_tac_to_ml_tac] function. *) let body = <:expr< Tacexpr.TacFun ($vars$, Tacexpr.TacML (Loc.tag ( $ml$ , []))) >> in - let name = <:expr< Names.Id.of_string $name$ >> in + let name = <:expr< API.Names.Id.of_string $name$ >> in declare_str_items loc [ <:str_item< do { let obj () = Tacenv.register_ltac True False $name$ $body$ in let () = Tacenv.register_ml_tactic $se$ [|$tac$|] in - Mltop.declare_cache_obj obj $plugin_name$ } >> + API.Mltop.declare_cache_obj obj $plugin_name$ } >> ] | _ -> (** Otherwise we add parsing and printing rules to generate a call to a @@ -131,7 +131,7 @@ let declare_tactic loc tacname ~level classification clause = match clause with declare_str_items loc [ <:str_item< do { Tacenv.register_ml_tactic $se$ (Array.of_list $make_fun_clauses loc tacname clause$); - Mltop.declare_cache_obj $obj$ $plugin_name$; } >> + Grammar_API.Mltop.declare_cache_obj $obj$ $plugin_name$; } >> ] open Pcaml diff --git a/grammar/vernacextend.mlp b/grammar/vernacextend.mlp index 798b46523..6f0e9b7cf 100644 --- a/grammar/vernacextend.mlp +++ b/grammar/vernacextend.mlp @@ -100,12 +100,12 @@ let make_fun_classifiers loc s c l = mlexpr_of_list (fun x -> x) cl let make_prod_item = function - | ExtTerminal s -> <:expr< Egramml.GramTerminal $str:s$ >> + | ExtTerminal s -> <:expr< Grammar_API.Egramml.GramTerminal $str:s$ >> | ExtNonTerminal (g, ido) -> let nt = type_of_user_symbol g in - let base s = <:expr< Pcoq.genarg_grammar ($mk_extraarg loc s$) >> in + let base s = <:expr< Grammar_API.Pcoq.genarg_grammar ($mk_extraarg loc s$) >> in let typ = match ido with None -> None | Some _ -> Some nt in - <:expr< Egramml.GramNonTerminal ( Loc.tag ( $mlexpr_of_option (make_rawwit loc) typ$ , + <:expr< Grammar_API.Egramml.GramNonTerminal ( Loc.tag ( $mlexpr_of_option (make_rawwit loc) typ$ , $mlexpr_of_prod_entry_key base g$ ) ) >> let mlexpr_of_clause cl = @@ -122,9 +122,9 @@ let declare_command loc s c nt cl = let classl = make_fun_classifiers loc s c cl in declare_str_items loc [ <:str_item< do { - CList.iteri (fun i (depr, f) -> Vernacinterp.vinterp_add depr ($se$, i) f) $funcl$; - CList.iteri (fun i f -> Vernac_classifier.declare_vernac_classifier ($se$, i) f) $classl$; - CList.iteri (fun i r -> Egramml.extend_vernac_command_grammar ($se$, i) $nt$ r) $gl$; + CList.iteri (fun i (depr, f) -> Grammar_API.Vernacinterp.vinterp_add depr ($se$, i) f) $funcl$; + CList.iteri (fun i f -> API.Vernac_classifier.declare_vernac_classifier ($se$, i) f) $classl$; + CList.iteri (fun i r -> Grammar_API.Egramml.extend_vernac_command_grammar ($se$, i) $nt$ r) $gl$; } >> ] open Pcaml @@ -143,16 +143,16 @@ EXTEND | "DECLARE"; "PLUGIN"; name = STRING -> declare_str_items loc [ <:str_item< value __coq_plugin_name = $str:name$ >>; - <:str_item< value _ = Mltop.add_known_module __coq_plugin_name >>; + <:str_item< value _ = Grammar_API.Mltop.add_known_module __coq_plugin_name >>; ] ] ] ; classification: [ [ "CLASSIFIED"; "BY"; c = LIDENT -> <:expr< $lid:c$ >> | "CLASSIFIED"; "AS"; "SIDEFF" -> - <:expr< fun _ -> Vernac_classifier.classify_as_sideeff >> + <:expr< fun _ -> API.Vernac_classifier.classify_as_sideeff >> | "CLASSIFIED"; "AS"; "QUERY" -> - <:expr< fun _ -> Vernac_classifier.classify_as_query >> + <:expr< fun _ -> API.Vernac_classifier.classify_as_query >> ] ] ; deprecation: diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 79e0e6164..396dde046 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -45,8 +45,11 @@ let names_of_local_binders bl = (**********************************************************************) (* Functions on constr_expr *) +(* Note: redundant Numeral representations such as -0 and +0 (or different + numbers of leading zeros) are considered different here. *) + let prim_token_eq t1 t2 = match t1, t2 with -| Numeral i1, Numeral i2 -> Bigint.equal i1 i2 +| Numeral (n1,s1), Numeral (n2,s2) -> String.equal n1 n2 && s1 == s2 | String s1, String s2 -> String.equal s1 s2 | _ -> false diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 19ca8d50b..8a798bfb0 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -66,22 +66,138 @@ let print_universes = Detyping.print_universes (* This suppresses printing of primitive tokens (e.g. numeral) and notations *) let print_no_symbol = ref false -(* This tells which notations still not to used if print_no_symbol is true *) -let print_non_active_notations = ref ([] : interp_rule list) +(**********************************************************************) +(* Turning notations and scopes on and off for printing *) +module IRuleSet = Set.Make(struct + type t = interp_rule + let compare x y = Pervasives.compare x y + end) + +let inactive_notations_table = + Summary.ref ~name:"inactive_notations_table" (IRuleSet.empty) +let inactive_scopes_table = + Summary.ref ~name:"inactive_scopes_table" CString.Set.empty + +let show_scope scopt = + match scopt with + | None -> str "" + | Some sc -> spc () ++ str "in scope" ++ spc () ++ str sc + +let _show_inactive_notations () = + begin + if CString.Set.is_empty !inactive_scopes_table + then + Feedback.msg_notice (str "No inactive notation scopes.") + else + let _ = Feedback.msg_notice (str "Inactive notation scopes:") in + CString.Set.iter (fun sc -> Feedback.msg_notice (str " " ++ str sc)) + !inactive_scopes_table + end; + if IRuleSet.is_empty !inactive_notations_table + then + Feedback.msg_notice (str "No individual inactive notations.") + else + let _ = Feedback.msg_notice (str "Inactive notations:") in + IRuleSet.iter + (function + | NotationRule (scopt, ntn) -> + Feedback.msg_notice (str ntn ++ show_scope scopt) + | SynDefRule kn -> Feedback.msg_notice (str (Names.KerName.to_string kn))) + !inactive_notations_table + +let deactivate_notation nr = + match nr with + | SynDefRule kn -> + (* shouldn't we check wether it is well defined? *) + inactive_notations_table := IRuleSet.add nr !inactive_notations_table + | NotationRule (scopt, ntn) -> + match availability_of_notation (scopt, ntn) (scopt, []) with + | None -> user_err ~hdr:"Notation" + (str ntn ++ spc () ++ str "does not exist" + ++ (match scopt with + | None -> spc () ++ str "in the empty scope." + | Some _ -> show_scope scopt ++ str ".")) + | Some _ -> + if IRuleSet.mem nr !inactive_notations_table then + Feedback.msg_warning + (str "Notation" ++ spc () ++ str ntn ++ spc () + ++ str "is already inactive" ++ show_scope scopt ++ str ".") + else inactive_notations_table := IRuleSet.add nr !inactive_notations_table + +let reactivate_notation nr = + try + inactive_notations_table := + IRuleSet.remove nr !inactive_notations_table + with Not_found -> + match nr with + | NotationRule (scopt, ntn) -> + Feedback.msg_warning (str "Notation" ++ spc () ++ str ntn ++ spc () + ++ str "is already active" ++ show_scope scopt ++ + str ".") + | SynDefRule kn -> + Feedback.msg_warning + (str "Notation" ++ spc () ++ str (Names.KerName.to_string kn) + ++ spc () ++ str "is already active.") + + +let deactivate_scope sc = + ignore (find_scope sc); (* ensures that the scope exists *) + if CString.Set.mem sc !inactive_scopes_table + then + Feedback.msg_warning (str "Notation Scope" ++ spc () ++ str sc ++ spc () + ++ str "is already inactive.") + else + inactive_scopes_table := CString.Set.add sc !inactive_scopes_table + +let reactivate_scope sc = + try + inactive_scopes_table := CString.Set.remove sc !inactive_scopes_table + with Not_found -> + Feedback.msg_warning (str "Notation Scope" ++ spc () ++ str sc ++ spc () + ++ str "is already active.") + +let is_inactive_rule nr = + IRuleSet.mem nr !inactive_notations_table || + match nr with + | NotationRule (Some sc, ntn) -> CString.Set.mem sc !inactive_scopes_table + | NotationRule (None, ntn) -> false + | SynDefRule _ -> false + +(* args: notation, scope, activate/deactivate *) +let toggle_scope_printing ~scope ~activate = + if activate then + reactivate_scope scope + else + deactivate_scope scope + +let toggle_notation_printing ?scope ~notation ~activate = + if activate then + reactivate_notation (NotationRule (scope, notation)) + else + deactivate_notation (NotationRule (scope, notation)) (* This governs printing of projections using the dot notation symbols *) let print_projections = ref false let print_meta_as_hole = ref false -let with_arguments f = Flags.with_option print_arguments f -let with_implicits f = Flags.with_option print_implicits f -let with_coercions f = Flags.with_option print_coercions f let with_universes f = Flags.with_option print_universes f let with_meta_as_hole f = Flags.with_option print_meta_as_hole f let without_symbols f = Flags.with_option print_no_symbol f -let without_specific_symbols l f = - Flags.with_extra_values print_non_active_notations l f + +(* XXX: Where to put this in the library? Util maybe? *) +let protect_ref r nf f x = + let old_ref = !r in + r := nf !r; + try let res = f x in r := old_ref; res + with reraise -> + let reraise = Backtrace.add_backtrace reraise in + r := old_ref; + Exninfo.iraise reraise + +let without_specific_symbols l = + protect_ref inactive_notations_table + (fun tbl -> IRuleSet.(union (of_list l) tbl)) (**********************************************************************) (* Control printing of records *) @@ -239,23 +355,31 @@ let expand_curly_brackets loc mknot ntn l = let destPrim = function { CAst.v = CPrim t } -> Some t | _ -> None let destPatPrim = function { CAst.v = CPatPrim t } -> Some t | _ -> None +let is_number s = + let rec aux i = + Int.equal (String.length s) i || + match s.[i] with '0'..'9' -> aux (i+1) | _ -> false + in aux 0 + +let is_zero s = + let rec aux i = + Int.equal (String.length s) i || (s.[i] == '0' && aux (i+1)) + in aux 0 + let make_notation_gen loc ntn mknot mkprim destprim l = if has_curly_brackets ntn then expand_curly_brackets loc mknot ntn l else match ntn,List.map destprim l with (* Special case to avoid writing "- 3" for e.g. (Z.opp 3) *) - | "- _", [Some (Numeral p)] when Bigint.is_strictly_pos p -> + | "- _", [Some (Numeral (p,true))] when not (is_zero p) -> mknot (loc,ntn,([mknot (loc,"( _ )",l)])) | _ -> match decompose_notation_key ntn, l with - | [Terminal "-"; Terminal x], [] -> - (try mkprim (loc, Numeral (Bigint.neg (Bigint.of_string x))) - with Failure _ -> mknot (loc,ntn,[])) - | [Terminal x], [] -> - (try mkprim (loc, Numeral (Bigint.of_string x)) - with Failure _ -> mknot (loc,ntn,[])) - | _ -> - mknot (loc,ntn,l) + | [Terminal "-"; Terminal x], [] when is_number x -> + mkprim (loc, Numeral (x,false)) + | [Terminal x], [] when is_number x -> + mkprim (loc, Numeral (x,true)) + | _ -> mknot (loc,ntn,l) let make_notation loc ntn (terms,termlists,binders as subst) = if not (List.is_empty termlists) || not (List.is_empty binders) then @@ -288,17 +412,8 @@ let pattern_printable_in_both_syntax (ind,_ as c) = (* Better to use extern_glob_constr composed with injection/retraction ?? *) let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = - (* pboutill: There are letins in pat which is incompatible with notations and - not explicit application. *) - match pat with - | { loc; v = PatCstr(cstrsp,args,na) } - when !Flags.in_debugger||Inductiveops.constructor_has_local_defs cstrsp -> - let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in - let args = List.map (extern_cases_pattern_in_scope scopes vars) args in - CAst.make ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) - | _ -> try - if !Flags.raw_print || !print_no_symbol then raise No_match; + if !Flags.in_debugger || !Flags.raw_print || !print_no_symbol then raise No_match; let (na,sc,p) = uninterp_prim_token_cases_pattern pat in match availability_of_prim_token p sc scopes with | None -> raise No_match @@ -307,7 +422,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = insert_pat_alias ?loc (insert_pat_delimiters ?loc (CAst.make ?loc @@ CPatPrim p) key) na with No_match -> try - if !Flags.raw_print || !print_no_symbol then raise No_match; + if !Flags.in_debugger || !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation_pattern scopes vars pat (uninterp_cases_pattern_notations pat) with No_match -> @@ -321,21 +436,19 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = if !Flags.raw_print then raise Exit; let projs = Recordops.lookup_projections (fst cstrsp) in let rec ip projs args acc = - match projs with - | [] -> acc - | None :: q -> ip q args acc - | Some c :: q -> - match args with - | [] -> raise No_match - - - - - - | { CAst.v = CPatAtom None } :: tail -> ip q tail acc - (* we don't want to have 'x = _' in our patterns *) - | head :: tail -> ip q tail - ((extern_reference ?loc Id.Set.empty (ConstRef c), head) :: acc) + match projs, args with + | [], [] -> acc + | proj :: q, pat :: tail -> + let acc = + match proj, pat with + | _, { CAst.v = CPatAtom None } -> + (* we don't want to have 'x := _' in our patterns *) + acc + | Some c, _ -> + ((extern_reference ?loc Id.Set.empty (ConstRef c), pat) :: acc) + | _ -> raise No_match in + ip q tail acc + | _ -> assert false in CPatRecord(List.rev (ip projs args [])) with @@ -401,7 +514,7 @@ and extern_notation_pattern (tmp_scope,scopes as allscopes) vars t = function | [] -> raise No_match | (keyrule,pat,n as _rule)::rules -> try - if List.mem keyrule !print_non_active_notations then raise No_match; + if is_inactive_rule keyrule then raise No_match; let loc = t.loc in match t.v with | PatCstr (cstr,_,na) -> @@ -417,8 +530,8 @@ let rec extern_notation_ind_pattern allscopes vars ind args = function | [] -> raise No_match | (keyrule,pat,n as _rule)::rules -> try - if List.mem keyrule !print_non_active_notations then raise No_match; - apply_notation_to_pattern (IndRef ind) + if is_inactive_rule keyrule then raise No_match; + apply_notation_to_pattern (IndRef ind) (match_notation_constr_ind_pattern ind args pat) allscopes vars keyrule with No_match -> extern_notation_ind_pattern allscopes vars ind args rules @@ -888,7 +1001,7 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function | (keyrule,pat,n as _rule)::rules -> let loc = Glob_ops.loc_of_glob_constr t in try - if List.mem keyrule !print_non_active_notations then raise No_match; + if is_inactive_rule keyrule then raise No_match; (* Adjusts to the number of arguments expected by the notation *) let (t,args,argsscopes,argsimpls) = match t.v ,n with | GApp (f,args), Some n diff --git a/interp/constrextern.mli b/interp/constrextern.mli index ea627cff1..6c82168e4 100644 --- a/interp/constrextern.mli +++ b/interp/constrextern.mli @@ -59,16 +59,6 @@ val set_extern_reference : val get_extern_reference : unit -> (?loc:Loc.t -> Id.Set.t -> global_reference -> reference) -(** This governs printing of implicit arguments. If [with_implicits] is - on and not [with_arguments] then implicit args are printed prefixed - by "!"; if [with_implicits] and [with_arguments] are both on the - function and not the arguments is prefixed by "!" *) -val with_implicits : ('a -> 'b) -> 'a -> 'b -val with_arguments : ('a -> 'b) -> 'a -> 'b - -(** This forces printing of coercions *) -val with_coercions : ('a -> 'b) -> 'a -> 'b - (** This forces printing universe names of Type\{.\} *) val with_universes : ('a -> 'b) -> 'a -> 'b @@ -80,3 +70,13 @@ val without_specific_symbols : interp_rule list -> ('a -> 'b) -> 'a -> 'b (** This prints metas as anonymous holes *) val with_meta_as_hole : ('a -> 'b) -> 'a -> 'b + +(** Fine-grained activation and deactivation of notation printing. + *) +val toggle_scope_printing : + scope:Notation_term.scope_name -> activate:bool -> unit + +val toggle_notation_printing : + ?scope:Notation_term.scope_name -> notation:Constrexpr.notation -> activate:bool -> unit + + diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 6f17324a1..89827300c 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -786,7 +786,7 @@ let find_appl_head_data c = let scopes = find_arguments_scope ref in c, impls, scopes, [] | GApp ({ v = GRef (ref,_) },l) - when l != [] && Flags.version_strictly_greater Flags.V8_2 -> + when l != [] -> let n = List.length l in let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in @@ -963,6 +963,45 @@ let check_constructor_length env loc cstr len_pl pl0 = (error_wrong_numarg_constructor ?loc env cstr (Inductiveops.constructor_nrealargs cstr))) +open Term +open Declarations + +(* Similar to Cases.adjust_local_defs but on RCPat *) +let insert_local_defs_in_pattern (ind,j) l = + let (mib,mip) = Global.lookup_inductive ind in + if mip.mind_consnrealdecls.(j-1) = mip.mind_consnrealargs.(j-1) then + (* Optimisation *) l + else + let typi = mip.mind_nf_lc.(j-1) in + let (_,typi) = decompose_prod_n_assum (Context.Rel.length mib.mind_params_ctxt) typi in + let (decls,_) = decompose_prod_assum typi in + let rec aux decls args = + match decls, args with + | Context.Rel.Declaration.LocalDef _ :: decls, args -> (CAst.make @@ RCPatAtom None) :: aux decls args + | _, [] -> [] (* In particular, if there were trailing local defs, they have been inserted *) + | Context.Rel.Declaration.LocalAssum _ :: decls, a :: args -> a :: aux decls args + | _ -> assert false in + aux (List.rev decls) l + +let add_local_defs_and_check_length loc env g pl args = match g with + | ConstructRef cstr -> + (* We consider that no variables corresponding to local binders + have been given in the "explicit" arguments, which come from a + "@C args" notation or from a custom user notation *) + let pl' = insert_local_defs_in_pattern cstr pl in + let maxargs = Inductiveops.constructor_nalldecls cstr in + if List.length pl' + List.length args > maxargs then + error_wrong_numarg_constructor ?loc env cstr (Inductiveops.constructor_nrealargs cstr); + (* Two possibilities: either the args are given with explict + variables for local definitions, then we give the explicit args + extended with local defs, so that there is nothing more to be + added later on; or the args are not enough to have all arguments, + which a priori means local defs to add in the [args] part, so we + postpone the insertion of local defs in the explicit args *) + (* Note: further checks done later by check_constructor_length *) + if List.length pl' + List.length args = maxargs then pl' else pl + | _ -> pl + let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 = let impl_list = if Int.equal len_pl1 0 then select_impargs_size (List.length pl2) impls_st @@ -1180,6 +1219,11 @@ let alias_of als = match als.alias_ids with *) +let is_zero s = + let rec aux i = + Int.equal (String.length s) i || (s.[i] == '0' && aux (i+1)) + in aux 0 + let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2 let product_of_cases_patterns aliases idspl = @@ -1200,7 +1244,7 @@ let rec subst_pat_iterator y t = CAst.(map (function | RCPatAlias (p,a) -> RCPatAlias (subst_pat_iterator y t p,a) | RCPatOr pl -> RCPatOr (List.map (subst_pat_iterator y t) pl))) -let drop_notations_pattern looked_for = +let drop_notations_pattern looked_for genv = (* At toplevel, Constructors and Inductives are accepted, in recursive calls only constructor are allowed *) let ensure_kind top loc g = @@ -1292,9 +1336,9 @@ let drop_notations_pattern looked_for = (* but not scopes in expl_pl *) let (argscs1,_) = find_remaining_scopes expl_pl pl g in CAst.make ?loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) - | CPatNotation ("- _",([{ CAst.v = CPatPrim(Numeral p) }],[]),[]) - when Bigint.is_strictly_pos p -> - let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes in + | CPatNotation ("- _",([{ CAst.v = CPatPrim(Numeral (p,true)) }],[]),[]) + when not (is_zero p) -> + let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (p,false)) scopes in rcp_of_glob pat | CPatNotation ("( _ )",([a],[]),[]) -> in_pat top scopes a @@ -1355,9 +1399,9 @@ let drop_notations_pattern looked_for = | NApp (NRef g,pl) -> ensure_kind top loc g; let (argscs1,argscs2) = find_remaining_scopes pl args g in - CAst.make ?loc @@ RCPatCstr (g, - List.map2 (fun x -> in_not false loc (x,snd scopes) fullsubst []) argscs1 pl @ - List.map (in_pat false scopes) args, []) + let pl = List.map2 (fun x -> in_not false loc (x,snd scopes) fullsubst []) argscs1 pl in + let pl = add_local_defs_and_check_length loc genv g pl args in + CAst.make ?loc @@ RCPatCstr (g, pl @ List.map (in_pat false scopes) args, []) | NList (x,y,iter,terminator,lassoc) -> if not (List.is_empty args) then user_err ?loc (strbrk "Application of arguments to a recursive notation not supported in patterns."); @@ -1418,7 +1462,7 @@ let rec intern_pat genv aliases pat = let intern_cases_pattern genv scopes aliases pat = intern_pat genv aliases - (drop_notations_pattern (function ConstructRef _ -> () | _ -> raise Not_found) scopes pat) + (drop_notations_pattern (function ConstructRef _ -> () | _ -> raise Not_found) genv scopes pat) let _ = intern_cases_pattern_fwd := @@ -1427,7 +1471,7 @@ let _ = let intern_ind_pattern genv scopes pat = let no_not = try - drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) scopes pat + drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) genv scopes pat with InternalizationError(loc,NotAConstructor _) -> error_bad_inductive_type ?loc in let loc = no_not.CAst.loc in @@ -1600,9 +1644,9 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = CAst.make ?loc @@ GLetIn (snd na, inc1, int, intern (push_name_env ntnvars (impls_term_list inc1) env na) c2) - | CNotation ("- _",([{ CAst.v = CPrim (Numeral p) }],[],[])) - when Bigint.is_strictly_pos p -> - intern env (CAst.make ?loc @@ CPrim (Numeral (Bigint.neg p))) + | CNotation ("- _",([{ CAst.v = CPrim (Numeral (p,true)) }],[],[])) + when not (is_zero p) -> + intern env (CAst.make ?loc @@ CPrim (Numeral (p,false))) | CNotation ("( _ )",([a],[],[])) -> intern env a | CNotation (ntn,args) -> intern_notation intern env ntnvars loc ntn args diff --git a/interp/notation.ml b/interp/notation.ml index 23332f7c4..300f6b1dd 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -10,7 +10,6 @@ open CErrors open Util open Pp -open Bigint open Names open Term open Libnames @@ -319,16 +318,34 @@ let declare_prim_token_interpreter sc interp (patl,uninterp,b) = (glob_prim_constr_key pat) (sc,uninterp,b) !prim_token_key_table) patl -let mkNumeral n = Numeral n +let mkNumeral n = + if Bigint.is_pos_or_zero n then Numeral (Bigint.to_string n, true) + else Numeral (Bigint.to_string (Bigint.neg n), false) + +let ofNumeral n s = + if s then Bigint.of_string n else Bigint.neg (Bigint.of_string n) + let mkString = function | None -> None | Some s -> if Unicode.is_utf8 s then Some (String s) else None let delay dir int ?loc x = (dir, (fun () -> int ?loc x)) +type rawnum = Constrexpr.raw_natural_number * Constrexpr.sign + +let declare_rawnumeral_interpreter sc dir interp (patl,uninterp,inpat) = + declare_prim_token_interpreter sc + (fun cont ?loc -> function Numeral (n,s) -> delay dir interp ?loc (n,s) + | p -> cont ?loc p) + (patl, (fun r -> match uninterp r with + | None -> None + | Some (n,s) -> Some (Numeral (n,s))), inpat) + let declare_numeral_interpreter sc dir interp (patl,uninterp,inpat) = + let interp' ?loc (n,s) = interp ?loc (ofNumeral n s) in declare_prim_token_interpreter sc - (fun cont ?loc -> function Numeral n-> delay dir interp ?loc n | p -> cont ?loc p) + (fun cont ?loc -> function Numeral (n,s) -> delay dir interp' ?loc (n,s) + | p -> cont ?loc p) (patl, (fun r -> Option.map mkNumeral (uninterp r)), inpat) let declare_string_interpreter sc dir interp (patl,uninterp,inpat) = @@ -440,8 +457,8 @@ let find_notation ntn sc = (n.not_interp, n.not_location) let notation_of_prim_token = function - | Numeral n when is_pos_or_zero n -> to_string n - | Numeral n -> "- "^(to_string (neg n)) + | Numeral (n,true) -> n + | Numeral (n,false) -> "- "^n | String _ -> raise Not_found let find_prim_token check_allowed ?loc p sc = @@ -466,7 +483,8 @@ let interp_prim_token_gen ?loc g p local_scopes = with Not_found -> user_err ?loc ~hdr:"interp_prim_token" ((match p with - | Numeral n -> str "No interpretation for numeral " ++ str (to_string n) + | Numeral _ -> + str "No interpretation for numeral " ++ str (notation_of_prim_token p) | String s -> str "No interpretation for string " ++ qs s) ++ str ".") let interp_prim_token ?loc = diff --git a/interp/notation.mli b/interp/notation.mli index d271a88fe..c739ec12f 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -74,6 +74,11 @@ type 'a prim_token_interpreter = type 'a prim_token_uninterpreter = glob_constr list * (glob_constr -> 'a option) * cases_pattern_status +type rawnum = Constrexpr.raw_natural_number * Constrexpr.sign + +val declare_rawnumeral_interpreter : scope_name -> required_module -> + rawnum prim_token_interpreter -> rawnum prim_token_uninterpreter -> unit + val declare_numeral_interpreter : scope_name -> required_module -> bigint prim_token_interpreter -> bigint prim_token_uninterpreter -> unit diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 08b9fbe8e..33b93606e 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -1154,10 +1154,6 @@ let match_notation_constr u c (metas,pat) = metas ([],[],[]) (* Matching cases pattern *) -let add_patterns_for_params ind l = - let mib,_ = Global.lookup_inductive ind in - let nparams = mib.Declarations.mind_nparams in - Util.List.addn nparams (CAst.make @@ PatVar Anonymous) l let bind_env_cases_pattern (terms,x,termlists,y as sigma) var v = try @@ -1187,10 +1183,11 @@ let rec match_cases_pattern metas (terms,(),termlists,() as sigma) a1 a2 = | r1, NVar id2 when Id.List.mem_assoc id2 metas -> (bind_env_cases_pattern sigma id2 a1),(0,[]) | PatVar Anonymous, NHole _ -> sigma,(0,[]) | PatCstr ((ind,_ as r1),largs,_), NRef (ConstructRef r2) when eq_constructor r1 r2 -> - sigma,(0,add_patterns_for_params (fst r1) largs) + let l = try add_patterns_for_params_remove_local_defs r1 largs with Not_found -> raise No_match in + sigma,(0,l) | PatCstr ((ind,_ as r1),args1,_), NApp (NRef (ConstructRef r2),l2) when eq_constructor r1 r2 -> - let l1 = add_patterns_for_params (fst r1) args1 in + let l1 = try add_patterns_for_params_remove_local_defs r1 args1 with Not_found -> raise No_match in let le2 = List.length l2 in if Int.equal le2 0 (* Special case of a notation for a @Cstr *) || le2 > List.length l1 then diff --git a/intf/constrexpr.mli b/intf/constrexpr.ml index 614c097b5..593b190a6 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.ml @@ -31,8 +31,16 @@ type abstraction_kind = AbsLambda | AbsPi type proj_flag = int option (** [Some n] = proj of the n-th visible argument *) +(** Representation of integer literals that appear in Coq scripts. + We now use raw strings of digits in base 10 (big-endian), and a separate + sign flag. Note that this representation is not unique, due to possible + multiple leading zeros, and -0 = +0 *) + +type sign = bool +type raw_natural_number = string + type prim_token = - | Numeral of Bigint.bigint (** representation of integer literals that appear in Coq scripts. *) + | Numeral of raw_natural_number * sign | String of string type instance_expr = Misctypes.glob_level list diff --git a/intf/decl_kinds.mli b/intf/decl_kinds.ml index 8254b1b80..8254b1b80 100644 --- a/intf/decl_kinds.mli +++ b/intf/decl_kinds.ml diff --git a/intf/evar_kinds.mli b/intf/evar_kinds.ml index ac0d96e96..ac0d96e96 100644 --- a/intf/evar_kinds.mli +++ b/intf/evar_kinds.ml diff --git a/intf/extend.mli b/intf/extend.ml index 99401d06f..99401d06f 100644 --- a/intf/extend.mli +++ b/intf/extend.ml diff --git a/intf/genredexpr.mli b/intf/genredexpr.ml index 2a542e0ff..2a542e0ff 100644 --- a/intf/genredexpr.mli +++ b/intf/genredexpr.ml diff --git a/intf/glob_term.mli b/intf/glob_term.ml index 5da20c9d1..a35dae4aa 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.ml @@ -95,3 +95,19 @@ type closure = { and closed_glob_constr = { closure: closure; term: glob_constr } + +(** Ltac variable maps *) +type var_map = Pattern.constr_under_binders Id.Map.t +type uconstr_var_map = closed_glob_constr Id.Map.t +type unbound_ltac_var_map = Geninterp.Val.t Id.Map.t + +type ltac_var_map = { + ltac_constrs : var_map; + (** Ltac variables bound to constrs *) + ltac_uconstrs : uconstr_var_map; + (** Ltac variables bound to untyped constrs *) + ltac_idents: Id.t Id.Map.t; + (** Ltac variables bound to identifiers *) + ltac_genargs : unbound_ltac_var_map; + (** Ltac variables bound to other kinds of arguments *) +} diff --git a/intf/intf.mllib b/intf/intf.mllib new file mode 100644 index 000000000..523e4b265 --- /dev/null +++ b/intf/intf.mllib @@ -0,0 +1,12 @@ +Constrexpr +Evar_kinds +Genredexpr +Locus +Notation_term +Tactypes +Decl_kinds +Extend +Glob_term +Misctypes +Pattern +Vernacexpr diff --git a/intf/locus.mli b/intf/locus.ml index 57b398ab4..57b398ab4 100644 --- a/intf/locus.mli +++ b/intf/locus.ml diff --git a/intf/misctypes.mli b/intf/misctypes.ml index 2ab70a78e..2ab70a78e 100644 --- a/intf/misctypes.mli +++ b/intf/misctypes.ml diff --git a/intf/notation_term.mli b/intf/notation_term.ml index 753fa657a..753fa657a 100644 --- a/intf/notation_term.mli +++ b/intf/notation_term.ml diff --git a/intf/pattern.mli b/intf/pattern.ml index 48381cacd..48381cacd 100644 --- a/intf/pattern.mli +++ b/intf/pattern.ml diff --git a/intf/tactypes.mli b/intf/tactypes.ml index 5c1d31946..5c1d31946 100644 --- a/intf/tactypes.mli +++ b/intf/tactypes.ml diff --git a/intf/vernacexpr.mli b/intf/vernacexpr.ml index ab440c6b7..cabd06735 100644 --- a/intf/vernacexpr.mli +++ b/intf/vernacexpr.ml @@ -96,17 +96,13 @@ type locatable = type showable = | ShowGoal of goal_reference - | ShowGoalImplicitly of int option | ShowProof - | ShowNode | ShowScript | ShowExistentials | ShowUniverses - | ShowTree | ShowProofNames | ShowIntros of bool | ShowMatch of reference - | ShowThesis type comment = | CommentConstr of constr_expr diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli index 077756ac7..69a5e79b4 100644 --- a/kernel/cClosure.mli +++ b/kernel/cClosure.mli @@ -187,7 +187,7 @@ val create_clos_infos : ?evars:(existential->constr option) -> reds -> env -> clos_infos val oracle_of_infos : clos_infos -> Conv_oracle.oracle -val env_of_infos : clos_infos -> env +val env_of_infos : 'a infos -> env val infos_with_reds : clos_infos -> reds -> clos_infos diff --git a/kernel/declarations.mli b/kernel/declarations.ml index 71e228b19..71e228b19 100644 --- a/kernel/declarations.mli +++ b/kernel/declarations.ml diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib index 4c540a6d7..2f49982ce 100644 --- a/kernel/kernel.mllib +++ b/kernel/kernel.mllib @@ -42,3 +42,4 @@ Safe_typing Vm Csymtable Vconv +Declarations diff --git a/kernel/reduction.ml b/kernel/reduction.ml index 427ce04c5..b6786c045 100644 --- a/kernel/reduction.ml +++ b/kernel/reduction.ml @@ -356,17 +356,17 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = let (app1,app2) = if Conv_oracle.oracle_order Univ.out_punivs oracle l2r fl1 fl2 then match unfold_reference infos fl1 with - | Some def1 -> ((lft1, whd def1 v1), appr2) + | Some def1 -> ((lft1, (def1, v1)), appr2) | None -> (match unfold_reference infos fl2 with - | Some def2 -> (appr1, (lft2, whd def2 v2)) + | Some def2 -> (appr1, (lft2, (def2, v2))) | None -> raise NotConvertible) else match unfold_reference infos fl2 with - | Some def2 -> (appr1, (lft2, whd def2 v2)) + | Some def2 -> (appr1, (lft2, (def2, v2))) | None -> (match unfold_reference infos fl1 with - | Some def1 -> ((lft1, whd def1 v1), appr2) + | Some def1 -> ((lft1, (def1, v1)), appr2) | None -> raise NotConvertible) in eqappr cv_pb l2r infos app1 app2 cuniv) @@ -377,11 +377,11 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = form *) (match unfold_projection infos p1 c1 with | Some (def1,s1) -> - eqappr cv_pb l2r infos (lft1, whd def1 (s1 :: v1)) appr2 cuniv + eqappr cv_pb l2r infos (lft1, (def1, s1 :: v1)) appr2 cuniv | None -> match unfold_projection infos p2 c2 with | Some (def2,s2) -> - eqappr cv_pb l2r infos appr1 (lft2, whd def2 (s2 :: v2)) cuniv + eqappr cv_pb l2r infos appr1 (lft2, (def2, s2 :: v2)) cuniv | None -> if Constant.equal (Projection.constant p1) (Projection.constant p2) && compare_stack_shape v1 v2 then @@ -393,26 +393,26 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | (FProj (p1,c1), t2) -> (match unfold_projection infos p1 c1 with | Some (def1,s1) -> - eqappr cv_pb l2r infos (lft1, whd def1 (s1 :: v1)) appr2 cuniv + eqappr cv_pb l2r infos (lft1, (def1, s1 :: v1)) appr2 cuniv | None -> (match t2 with | FFlex fl2 -> (match unfold_reference infos fl2 with | Some def2 -> - eqappr cv_pb l2r infos appr1 (lft2, whd def2 v2) cuniv + eqappr cv_pb l2r infos appr1 (lft2, (def2, v2)) cuniv | None -> raise NotConvertible) | _ -> raise NotConvertible)) | (t1, FProj (p2,c2)) -> (match unfold_projection infos p2 c2 with | Some (def2,s2) -> - eqappr cv_pb l2r infos appr1 (lft2, whd def2 (s2 :: v2)) cuniv + eqappr cv_pb l2r infos appr1 (lft2, (def2, s2 :: v2)) cuniv | None -> (match t1 with | FFlex fl1 -> (match unfold_reference infos fl1 with | Some def1 -> - eqappr cv_pb l2r infos (lft1, whd def1 v1) appr2 cuniv + eqappr cv_pb l2r infos (lft1, (def1, v1)) appr2 cuniv | None -> raise NotConvertible) | _ -> raise NotConvertible)) @@ -458,7 +458,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | (FFlex fl1, c2) -> (match unfold_reference infos fl1 with | Some def1 -> - eqappr cv_pb l2r infos (lft1, whd def1 v1) appr2 cuniv + eqappr cv_pb l2r infos (lft1, (def1, v1)) appr2 cuniv | None -> match c2 with | FConstruct ((ind2,j2),u2) -> @@ -472,7 +472,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | (c1, FFlex fl2) -> (match unfold_reference infos fl2 with | Some def2 -> - eqappr cv_pb l2r infos appr1 (lft2, whd def2 v2) cuniv + eqappr cv_pb l2r infos appr1 (lft2, (def2, v2)) cuniv | None -> match c1 with | FConstruct ((ind1,j1),u1) -> diff --git a/kernel/retroknowledge.ml b/kernel/retroknowledge.ml index 970bc0fcc..ea53d00d7 100644 --- a/kernel/retroknowledge.ml +++ b/kernel/retroknowledge.ml @@ -83,7 +83,7 @@ type flags = {fastcomputation : bool} (* The [proactive] knowledge contains the mapping [field->entry]. *) module Proactive = - Map.Make (struct type t = field let compare = compare end) + Map.Make (struct type t = field let compare = Pervasives.compare end) type proactive = entry Proactive.t diff --git a/kernel/term.ml b/kernel/term.ml index 07a85329e..b90718358 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -143,7 +143,8 @@ let leq_constr_univs = Constr.leq_constr_univs let eq_constr_nounivs = Constr.eq_constr_nounivs let kind_of_term = Constr.kind -let constr_ord = Constr.compare +let compare = Constr.compare +let constr_ord = compare let fold_constr = Constr.fold let map_puniverses = Constr.map_puniverses let map_constr = Constr.map diff --git a/kernel/term.mli b/kernel/term.mli index 241ef322f..e729439f0 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -447,9 +447,12 @@ val eq_constr_nounivs : constr -> constr -> bool val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term (** Alias for [Constr.kind] *) -val constr_ord : constr -> constr -> int +val compare : constr -> constr -> int (** Alias for [Constr.compare] *) +val constr_ord : constr -> constr -> int +(** Alias for [Term.compare] *) + val fold_constr : ('a -> constr -> 'a) -> 'a -> constr -> 'a (** Alias for [Constr.fold] *) diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml index eeb9c421a..bdfd00a8d 100644 --- a/kernel/term_typing.ml +++ b/kernel/term_typing.ml @@ -344,11 +344,18 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx) if not (Id.Set.subset inferred_set declared_set) then let l = Id.Set.elements (Idset.diff inferred_set declared_set) in let n = List.length l in - user_err (Pp.(str "The following section " ++ - str (String.plural n "variable") ++ - str " " ++ str (String.conjugate_verb_to_be n) ++ - str " used but not declared:" ++ - fnl () ++ pr_sequence Id.print (List.rev l) ++ str ".")) in + let declared_vars = Pp.pr_sequence Id.print (Idset.elements declared_set) in + let inferred_vars = Pp.pr_sequence Id.print (Idset.elements inferred_set) in + let missing_vars = Pp.pr_sequence Id.print (List.rev l) in + user_err Pp.(prlist str + ["The following section "; (String.plural n "variable"); " "; + (String.conjugate_verb_to_be n); " used but not declared:"] ++ fnl () ++ + missing_vars ++ str "." ++ fnl () ++ fnl () ++ + str "You can either update your proof to not depend on " ++ missing_vars ++ + str ", or you can update your Proof line from" ++ fnl () ++ + str "Proof using " ++ declared_vars ++ fnl () ++ + str "to" ++ fnl () ++ + str "Proof using " ++ inferred_vars) in let sort evn l = List.filter (fun decl -> let id = NamedDecl.get_id decl in diff --git a/lib/coqProject_file.ml4 b/lib/coqProject_file.ml4 index 7a1660569..97aa90e07 100644 --- a/lib/coqProject_file.ml4 +++ b/lib/coqProject_file.ml4 @@ -11,6 +11,7 @@ type project = { makefile : string option; install_kind : install option; use_ocamlopt : bool; + bypass_API : bool; v_files : string list; mli_files : string list; @@ -42,11 +43,12 @@ and install = | UserInstall (* TODO generate with PPX *) -let mk_project project_file makefile install_kind use_ocamlopt = { +let mk_project project_file makefile install_kind use_ocamlopt bypass_API = { project_file; makefile; install_kind; use_ocamlopt; + bypass_API; v_files = []; mli_files = []; @@ -166,6 +168,8 @@ let process_cmd_line orig_dir proj args = aux { proj with defs = proj.defs @ [v,def] } r | "-arg" :: a :: r -> aux { proj with extra_args = proj.extra_args @ [a] } r + | "-bypass-API" :: r -> + aux { proj with bypass_API = true } r | f :: r -> let f = CUnix.correct_path f orig_dir in let proj = @@ -185,11 +189,11 @@ let process_cmd_line orig_dir proj args = (******************************* API ************************************) let cmdline_args_to_project ~curdir args = - process_cmd_line curdir (mk_project None None None true) args + process_cmd_line curdir (mk_project None None None true false) args let read_project_file f = process_cmd_line (Filename.dirname f) - (mk_project (Some f) None (Some NoInstall) true) (parse f) + (mk_project (Some f) None (Some NoInstall) true false) (parse f) let rec find_project_file ~from ~projfile_name = let fname = Filename.concat from projfile_name in diff --git a/lib/coqProject_file.mli b/lib/coqProject_file.mli index 8c8fc068a..19fc9227a 100644 --- a/lib/coqProject_file.mli +++ b/lib/coqProject_file.mli @@ -13,6 +13,7 @@ type project = { makefile : string option; install_kind : install option; use_ocamlopt : bool; + bypass_API : bool; v_files : string list; mli_files : string list; diff --git a/lib/envars.ml b/lib/envars.ml index bc8012297..47baf66a6 100644 --- a/lib/envars.ml +++ b/lib/envars.ml @@ -202,14 +202,7 @@ let xdg_dirs ~warn = (* Print the configuration information *) -let coq_src_subdirs = [ - "config" ; "dev" ; "lib" ; "kernel" ; "library" ; - "engine" ; "pretyping" ; "interp" ; "parsing" ; "proofs" ; - "tactics" ; "toplevel" ; "printing" ; "intf" ; - "grammar" ; "ide" ; "stm"; "vernac" ] @ - Coq_config.plugins_dirs - -let print_config ?(prefix_var_name="") f = +let print_config ?(prefix_var_name="") f coq_src_subdirs = let open Printf in fprintf f "%sLOCAL=%s\n" prefix_var_name (if Coq_config.local then "1" else "0"); fprintf f "%sCOQLIB=%s/\n" prefix_var_name (coqlib ()); diff --git a/lib/envars.mli b/lib/envars.mli index c8bbf17d9..edd13447f 100644 --- a/lib/envars.mli +++ b/lib/envars.mli @@ -76,7 +76,4 @@ val xdg_data_dirs : (string -> unit) -> string list val xdg_dirs : warn : (string -> unit) -> string list (** {6 Prints the configuration information } *) -val print_config : ?prefix_var_name:string -> out_channel -> unit - -(** Directories in which coq sources are found *) -val coq_src_subdirs : string list +val print_config : ?prefix_var_name:string -> out_channel -> string list -> unit diff --git a/lib/flags.ml b/lib/flags.ml index 6a3b7a426..13539bced 100644 --- a/lib/flags.ml +++ b/lib/flags.ml @@ -106,35 +106,27 @@ let we_are_parsing = ref false (* Current means no particular compatibility consideration. For correct comparisons, this constructor should remain the last one. *) -type compat_version = V8_2 | V8_3 | V8_4 | V8_5 | V8_6 | Current +type compat_version = VOld | V8_5 | V8_6 | Current let compat_version = ref Current let version_compare v1 v2 = match v1, v2 with -| V8_2, V8_2 -> 0 -| V8_2, (V8_3 | V8_4 | V8_5 | V8_6 | Current) -> -1 -| V8_3, V8_2 -> 1 -| V8_3, V8_3 -> 0 -| V8_3, (V8_4 | V8_5 | V8_6 | Current) -> -1 -| V8_4, (V8_2 | V8_3) -> 1 -| V8_4, V8_4 -> 0 -| V8_4, (V8_5 | V8_6 | Current) -> -1 -| V8_5, (V8_2 | V8_3 | V8_4) -> 1 -| V8_5, V8_5 -> 0 -| V8_5, (V8_6 | Current) -> -1 -| V8_6, (V8_2 | V8_3 | V8_4 | V8_5) -> 1 -| V8_6, V8_6 -> 0 -| V8_6, Current -> -1 -| Current, Current -> 0 -| Current, (V8_2 | V8_3 | V8_4 | V8_5 | V8_6) -> 1 + | VOld, VOld -> 0 + | VOld, _ -> -1 + | _, VOld -> 1 + | V8_5, V8_5 -> 0 + | V8_5, _ -> -1 + | _, V8_5 -> 1 + | V8_6, V8_6 -> 0 + | V8_6, _ -> -1 + | _, V8_6 -> 1 + | Current, Current -> 0 let version_strictly_greater v = version_compare !compat_version v > 0 let version_less_or_equal v = not (version_strictly_greater v) let pr_version = function - | V8_2 -> "8.2" - | V8_3 -> "8.3" - | V8_4 -> "8.4" + | VOld -> "old" | V8_5 -> "8.5" | V8_6 -> "8.6" | Current -> "current" @@ -157,7 +149,7 @@ let is_verbose () = not !quiet let auto_intros = ref true let make_auto_intros flag = auto_intros := flag -let is_auto_intros () = version_strictly_greater V8_2 && !auto_intros +let is_auto_intros () = !auto_intros let universe_polymorphism = ref false let make_universe_polymorphism b = universe_polymorphism := b diff --git a/lib/flags.mli b/lib/flags.mli index e2cf09474..0026aba2e 100644 --- a/lib/flags.mli +++ b/lib/flags.mli @@ -77,7 +77,7 @@ val raw_print : bool ref (* Univ print flag, never set anywere. Maybe should belong to Univ? *) val univ_print : bool ref -type compat_version = V8_2 | V8_3 | V8_4 | V8_5 | V8_6 | Current +type compat_version = VOld | V8_5 | V8_6 | Current val compat_version : compat_version ref val version_compare : compat_version -> compat_version -> int val version_strictly_greater : compat_version -> bool diff --git a/lib/pp.mli b/lib/pp.mli index 7a191b01a..45834dade 100644 --- a/lib/pp.mli +++ b/lib/pp.mli @@ -13,6 +13,7 @@ (* `Pp.t` or `Pp.std_ppcmds` is the main pretty printing document type *) (* in the Coq system. Documents are composed laying out boxes, and *) (* users can add arbitrary tag metadata that backends are free *) +(* to interpret. *) (* *) (* The datatype has a public view to allow serialization or advanced *) (* uses, however regular users are _strongly_ warned againt its use, *) diff --git a/library/declaremods.ml b/library/declaremods.ml index c98d4a7f3..187b749b8 100644 --- a/library/declaremods.ml +++ b/library/declaremods.ml @@ -589,7 +589,6 @@ let start_module interp_modast export id args res fs = openmod_info := { cur_typ = res_entry_o; cur_typs = subtyps }; let prefix = Lib.start_module export id mp fs in Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModule prefix); - Lib.add_frozen_state (); if_xml (Hook.get f_xml_start_module) mp; mp @@ -629,7 +628,6 @@ let end_module () = assert (eq_full_path (fst newoname) (fst oldoname)); assert (ModPath.equal (mp_of_kn (snd newoname)) mp); - Lib.add_frozen_state () (* to prevent recaching *); if_xml (Hook.get f_xml_end_module) mp; mp @@ -701,7 +699,6 @@ let start_modtype interp_modast id args mtys fs = openmodtype_info := sub_mty_l; let prefix = Lib.start_modtype id mp fs in Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModtype prefix); - Lib.add_frozen_state (); if_xml (Hook.get f_xml_start_module_type) mp; mp @@ -719,7 +716,6 @@ let end_modtype () = assert (eq_full_path (fst oname) (fst oldoname)); assert (ModPath.equal (mp_of_kn (snd oname)) mp); - Lib.add_frozen_state ()(* to prevent recaching *); if_xml (Hook.get f_xml_end_module_type) mp; mp @@ -894,8 +890,7 @@ let get_library_native_symbols dir = let start_library dir = let mp = Global.start_library dir in openmod_info := default_module_info; - Lib.start_compilation dir mp; - Lib.add_frozen_state () + Lib.start_compilation dir mp let end_library_hook = ref ignore let append_end_library_hook f = diff --git a/library/lib.ml b/library/lib.ml index 9d71a854f..f22f53ead 100644 --- a/library/lib.ml +++ b/library/lib.ml @@ -27,7 +27,6 @@ type node = | ClosedModule of library_segment | OpenedSection of object_prefix * Summary.frozen | ClosedSection of library_segment - | FrozenState of Summary.frozen and library_entry = object_name * node @@ -80,7 +79,6 @@ let classify_segment seg = | (_,OpenedModule (ty,_,_,_)) :: _ -> user_err ~hdr:"Lib.classify_segment" (str "there are still opened " ++ str (module_kind ty) ++ str "s") - | (_,FrozenState _) :: stk -> clean acc stk in clean ([],[],[]) (List.rev seg) @@ -254,10 +252,6 @@ let add_anonymous_leaf ?(cache_first = true) obj = cache_object (oname,obj) end -let add_frozen_state () = - add_anonymous_entry - (FrozenState (Summary.freeze_summaries ~marshallable:`No)) - (* Modules. *) let is_opening_node = function @@ -544,7 +538,6 @@ let discharge_item ((sp,_ as oname),e) = match e with | Leaf lobj -> Option.map (fun o -> (basename sp,o)) (discharge_object (oname,lobj)) - | FrozenState _ -> None | ClosedSection _ | ClosedModule _ -> None | OpenedSection _ | OpenedModule _ | CompilingLibrary _ -> anomaly (Pp.str "discharge_item.") @@ -585,8 +578,7 @@ let freeze ~marshallable = | n, ClosedModule _ -> Some (n,ClosedModule []) | n, OpenedSection (op, _) -> Some(n,OpenedSection(op,Summary.empty_frozen)) - | n, ClosedSection _ -> Some (n,ClosedSection []) - | _, FrozenState _ -> None) + | n, ClosedSection _ -> Some (n,ClosedSection [])) !lib_state.lib_stk in { !lib_state with lib_stk } | _ -> @@ -596,8 +588,7 @@ let unfreeze st = lib_state := st let init () = unfreeze initial_lib_state; - Summary.init_summaries (); - add_frozen_state () (* Stores e.g. the keywords declared in g_*.ml4 *) + Summary.init_summaries () (* Misc *) diff --git a/library/lib.mli b/library/lib.mli index 9f9d8c7e5..f47d6e1a5 100644 --- a/library/lib.mli +++ b/library/lib.mli @@ -23,7 +23,6 @@ type node = | ClosedModule of library_segment | OpenedSection of Libnames.object_prefix * Summary.frozen | ClosedSection of library_segment - | FrozenState of Summary.frozen and library_segment = (Libnames.object_name * node) list @@ -61,8 +60,6 @@ val pull_to_head : Libnames.object_name -> unit for each of them *) val add_leaves : Names.Id.t -> Libobject.obj list -> Libnames.object_name -val add_frozen_state : unit -> unit - (** {6 ... } *) (** The function [contents] gives access to the current entire segment *) @@ -123,8 +120,6 @@ val end_modtype : Libnames.object_name * Libnames.object_prefix * Summary.frozen * library_segment -(** [Lib.add_frozen_state] must be called after each of the above functions *) - (** {6 Compilation units } *) val start_compilation : Names.DirPath.t -> Names.module_path -> unit diff --git a/library/library.ml b/library/library.ml index 5a5f99cc5..db05ad2b7 100644 --- a/library/library.ml +++ b/library/library.ml @@ -575,7 +575,7 @@ let require_library_from_dirpath modrefl export = else add_anonymous_leaf (in_require (needed,modrefl,export)); if !Flags.xml_export then List.iter (Hook.get f_xml_require) modrefl; - add_frozen_state () + () (* the function called by Vernacentries.vernac_import *) diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index 890ce2dec..35ffa20d0 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -227,7 +227,7 @@ type prod_info = production_level * production_position type (_, _) entry = | TTName : ('self, Name.t Loc.located) entry | TTReference : ('self, reference) entry -| TTBigint : ('self, Bigint.bigint) entry +| TTBigint : ('self, Constrexpr.raw_natural_number) entry | TTBinder : ('self, local_binder_expr list) entry | TTConstr : prod_info * 'r target -> ('r, 'r) entry | TTConstrList : prod_info * Tok.t list * 'r target -> ('r, 'r list) entry @@ -337,8 +337,8 @@ match e with | TTBinderListF _ -> { subst with binders = (List.flatten v, false) :: subst.binders } | TTBigint -> begin match forpat with - | ForConstr -> push_constr subst (CAst.make @@ CPrim (Numeral v)) - | ForPattern -> push_constr subst (CAst.make @@ CPatPrim (Numeral v)) + | ForConstr -> push_constr subst (CAst.make @@ CPrim (Numeral (v,true))) + | ForPattern -> push_constr subst (CAst.make @@ CPatPrim (Numeral (v,true))) end | TTReference -> begin match forpat with diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 54bac253d..de7611802 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -203,7 +203,7 @@ GEXTEND Gram | c=match_constr -> c | "("; c = operconstr LEVEL "200"; ")" -> (match c.CAst.v with - CPrim (Numeral z) when Bigint.is_pos_or_zero z -> + | CPrim (Numeral (n,true)) -> CAst.make ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[])) | _ -> c) | "{|"; c = record_declaration; "|}" -> c @@ -280,7 +280,7 @@ GEXTEND Gram atomic_constr: [ [ g=global; i=instance -> CAst.make ~loc:!@loc @@ CRef (g,i) | s=sort -> CAst.make ~loc:!@loc @@ CSort s - | n=INT -> CAst.make ~loc:!@loc @@ CPrim (Numeral (Bigint.of_string n)) + | n=INT -> CAst.make ~loc:!@loc @@ CPrim (Numeral (n,true)) | s=string -> CAst.make ~loc:!@loc @@ CPrim (String s) | "_" -> CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None) | "?"; "["; id=ident; "]" -> CAst.make ~loc:!@loc @@ CHole (None, IntroIdentifier id, None) @@ -395,18 +395,18 @@ GEXTEND Gram | "_" -> CAst.make ~loc:!@loc @@ CPatAtom None | "("; p = pattern LEVEL "200"; ")" -> (match p.CAst.v with - | CPatPrim (Numeral z) when Bigint.is_pos_or_zero z -> + | CPatPrim (Numeral (n,true)) -> CAst.make ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) | _ -> p) | "("; p = pattern LEVEL "200"; ":"; ty = lconstr; ")" -> let p = match p with - | { CAst.v = CPatPrim (Numeral z) } when Bigint.is_pos_or_zero z -> + | { CAst.v = CPatPrim (Numeral (n,true)) } -> CAst.make ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) | _ -> p in CAst.make ~loc:!@loc @@ CPatCast (p, ty) - | n = INT -> CAst.make ~loc:!@loc @@ CPatPrim (Numeral (Bigint.of_string n)) + | n = INT -> CAst.make ~loc:!@loc @@ CPatPrim (Numeral (n,true)) | s = string -> CAst.make ~loc:!@loc @@ CPatPrim (String s) ] ] ; impl_ident_tail: diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index 78f75a73c..c77d6e204 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -114,7 +114,7 @@ GEXTEND Gram natural: [ [ i = INT -> my_int_of_string (!@loc) i ] ] ; - bigint: (* Negative numbers are dealt with specially *) - [ [ i = INT -> (Bigint.of_string i) ] ] + bigint: (* Negative numbers are dealt with elsewhere *) + [ [ i = INT -> i ] ] ; END diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index a3f9793bb..e96a68bc6 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -64,22 +64,14 @@ GEXTEND Gram | IDENT "Show" -> VernacShow (ShowGoal OpenSubgoals) | IDENT "Show"; n = natural -> VernacShow (ShowGoal (NthGoal n)) | IDENT "Show"; id = ident -> VernacShow (ShowGoal (GoalId id)) - | IDENT "Show"; IDENT "Goal" -> VernacShow (ShowGoal (GoalId (Names.Id.of_string "Goal"))) - | IDENT "Show"; IDENT "Goal"; n = string -> - VernacShow (ShowGoal (GoalUid n)) - | IDENT "Show"; IDENT "Implicit"; IDENT "Arguments"; n = OPT natural -> - VernacShow (ShowGoalImplicitly n) - | IDENT "Show"; IDENT "Node" -> VernacShow ShowNode | IDENT "Show"; IDENT "Script" -> VernacShow ShowScript | IDENT "Show"; IDENT "Existentials" -> VernacShow ShowExistentials | IDENT "Show"; IDENT "Universes" -> VernacShow ShowUniverses - | IDENT "Show"; IDENT "Tree" -> VernacShow ShowTree | IDENT "Show"; IDENT "Conjectures" -> VernacShow ShowProofNames | IDENT "Show"; IDENT "Proof" -> VernacShow ShowProof | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false) | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true) | IDENT "Show"; IDENT "Match"; id = reference -> VernacShow (ShowMatch id) - | IDENT "Show"; IDENT "Thesis" -> VernacShow ShowThesis | IDENT "Guarded" -> VernacCheckGuard (* Hints for Auto and EAuto *) | IDENT "Create"; IDENT "HintDb" ; diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 893605499..b605a44c8 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -51,6 +51,20 @@ let make_bullet s = | '*' -> Star n | _ -> assert false +let extraction_err ~loc = + if not (Mltop.module_is_known "extraction_plugin") then + CErrors.user_err ~loc (str "Please do first a Require Extraction.") + else + (* The right grammar entries should have been loaded. + We could only end here in case of syntax error. *) + raise (Stream.Error "unexpected end of command") + +let funind_err ~loc = + if not (Mltop.module_is_known "recdef_plugin") then + CErrors.user_err ~loc (str "Please do first a Require Import FunInd.") + else + raise (Stream.Error "unexpected end of command") (* Same as above... *) + GEXTEND Gram GLOBAL: vernac gallina_ext noedit_mode subprf; vernac: FIRST @@ -841,6 +855,22 @@ GEXTEND Gram | IDENT "DelPath"; dir = ne_string -> VernacRemoveLoadPath dir + (* Some plugins are not loaded initially anymore : extraction, + and funind. To ease this transition toward a mandatory Require, + we hack here the vernac grammar in order to get customized + error messages telling what to Require instead of the dreadful + "Illegal begin of vernac". Normally, these fake grammar entries + are overloaded later by the grammar extensions in these plugins. + This code is meant to be removed in a few releases, when this + transition is considered finished. *) + + | IDENT "Extraction" -> extraction_err ~loc:!@loc + | IDENT "Extract" -> extraction_err ~loc:!@loc + | IDENT "Recursive"; IDENT "Extraction" -> extraction_err ~loc:!@loc + | IDENT "Separate"; IDENT "Extraction" -> extraction_err ~loc:!@loc + | IDENT "Function" -> funind_err ~loc:!@loc + | IDENT "Functional" -> funind_err ~loc:!@loc + (* Type-Checking (pas dans le refman) *) | "Type"; c = lconstr -> VernacGlobalCheck c diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index 959e8ddf5..9fb3daaba 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -199,7 +199,7 @@ module Prim : val pattern_identref : Id.t located Gram.entry val base_ident : Id.t Gram.entry val natural : int Gram.entry - val bigint : Bigint.bigint Gram.entry + val bigint : Constrexpr.raw_natural_number Gram.entry val integer : int Gram.entry val string : string Gram.entry val lstring : string located Gram.entry diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 6281b2675..00e80d041 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -1,3 +1,5 @@ +open API + let contrib_name = "btauto" let init_constant dir s = diff --git a/plugins/btauto/vo.itarget b/plugins/btauto/vo.itarget deleted file mode 100644 index 1f72d3ef2..000000000 --- a/plugins/btauto/vo.itarget +++ /dev/null @@ -1,3 +0,0 @@ -Algebra.vo -Reflect.vo -Btauto.vo diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index ba398c385..5c7cad7ff 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -10,6 +10,7 @@ (* Downey,Sethi and Tarjan. *) (* Plus some e-matching and constructor handling by P. Corbineau *) +open API open CErrors open Util open Pp @@ -135,7 +136,7 @@ let family_eq f1 f2 = match f1, f2 with type term= Symb of constr - | Product of sorts * sorts + | Product of Sorts.t * Sorts.t | Eps of Id.t | Appli of term*term | Constructor of cinfo (* constructor arity + nhyps *) @@ -269,7 +270,7 @@ type state = mutable rew_depth:int; mutable changed:bool; by_type: Int.Set.t Typehash.t; - mutable gls:Proof_type.goal Tacmach.sigma} + mutable gls:Proof_type.goal Evd.sigma} let dummy_node = { @@ -456,13 +457,13 @@ let rec canonize_name sigma c = let func c = canonize_name sigma (EConstr.of_constr c) in match kind_of_term c with | Const (kn,u) -> - let canon_const = constant_of_kn (canonical_con kn) in + let canon_const = Constant.make1 (Constant.canonical kn) in (mkConstU (canon_const,u)) | Ind ((kn,i),u) -> - let canon_mind = mind_of_kn (canonical_mind kn) in + let canon_mind = MutInd.make1 (MutInd.canonical kn) in (mkIndU ((canon_mind,i),u)) | Construct (((kn,i),j),u) -> - let canon_mind = mind_of_kn (canonical_mind kn) in + let canon_mind = MutInd.make1 (MutInd.canonical kn) in mkConstructU (((canon_mind,i),j),u) | Prod (na,t,ct) -> mkProd (na,func t, func ct) @@ -474,7 +475,7 @@ let rec canonize_name sigma c = mkApp (func ct,Array.smartmap func l) | Proj(p,c) -> let p' = Projection.map (fun kn -> - constant_of_kn (canonical_con kn)) p in + Constant.make1 (Constant.canonical kn)) p in (mkProj (p', func c)) | _ -> c diff --git a/plugins/cc/ccalgo.mli b/plugins/cc/ccalgo.mli index c7fa2f56f..505029992 100644 --- a/plugins/cc/ccalgo.mli +++ b/plugins/cc/ccalgo.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Term open Names @@ -30,7 +31,7 @@ type cinfo = type term = Symb of constr - | Product of sorts * sorts + | Product of Sorts.t * Sorts.t | Eps of Id.t | Appli of term*term | Constructor of cinfo (* constructor arity + nhyps *) @@ -128,7 +129,7 @@ val axioms : forest -> (term * term) Constrhash.t val epsilons : forest -> pa_constructor list -val empty : int -> Proof_type.goal Tacmach.sigma -> state +val empty : int -> Proof_type.goal Evd.sigma -> state val add_term : state -> term -> int diff --git a/plugins/cc/ccproof.ml b/plugins/cc/ccproof.ml index 642ceba3d..eecb7bc98 100644 --- a/plugins/cc/ccproof.ml +++ b/plugins/cc/ccproof.ml @@ -9,6 +9,7 @@ (* This file uses the (non-compressed) union-find structure to generate *) (* proof-trees that will be transformed into proof-terms in cctac.ml4 *) +open API open CErrors open Term open Ccalgo diff --git a/plugins/cc/ccproof.mli b/plugins/cc/ccproof.mli index eacbfeac7..4e4d42f86 100644 --- a/plugins/cc/ccproof.mli +++ b/plugins/cc/ccproof.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Ccalgo open Term diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index b638f2360..0f5b80664 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -8,6 +8,7 @@ (* This file is the interface between the c-c algorithm and Coq *) +open API open Evd open Names open Inductiveops @@ -65,7 +66,7 @@ let rec decompose_term env sigma t= | Construct c -> let (((mind,i_ind),i_con),u)= c in let u = EInstance.kind sigma u in - let canon_mind = mind_of_kn (canonical_mind mind) in + let canon_mind = MutInd.make1 (MutInd.canonical mind) in let canon_ind = canon_mind,i_ind in let (oib,_)=Global.lookup_inductive (canon_ind) in let nargs=constructor_nallargs_env env (canon_ind,i_con) in @@ -75,16 +76,16 @@ let rec decompose_term env sigma t= | Ind c -> let (mind,i_ind),u = c in let u = EInstance.kind sigma u in - let canon_mind = mind_of_kn (canonical_mind mind) in - let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u))) + let canon_mind = MutInd.make1 (MutInd.canonical mind) in + let canon_ind = canon_mind,i_ind in (Symb (Term.mkIndU (canon_ind,u))) | Const (c,u) -> let u = EInstance.kind sigma u in - let canon_const = constant_of_kn (canonical_con c) in - (Symb (Constr.mkConstU (canon_const,u))) + let canon_const = Constant.make1 (Constant.canonical c) in + (Symb (Term.mkConstU (canon_const,u))) | Proj (p, c) -> - let canon_const kn = constant_of_kn (canonical_con kn) in + let canon_const kn = Constant.make1 (Constant.canonical kn) in let p' = Projection.map canon_const p in - (Appli (Symb (Constr.mkConst (Projection.constant p')), decompose_term env sigma c)) + (Appli (Symb (Term.mkConst (Projection.constant p')), decompose_term env sigma c)) | _ -> let t = Termops.strip_outer_cast sigma t in if closed0 sigma t then Symb (EConstr.to_constr sigma t) else raise Not_found @@ -197,7 +198,7 @@ let make_prb gls depth additionnal_terms = (fun decl -> let id = NamedDecl.get_id decl in begin - let cid=Constr.mkVar id in + let cid=Term.mkVar id in match litteral_of_constr env sigma (NamedDecl.get_type decl) with `Eq (t,a,b) -> add_equality state cid a b | `Neq (t,a,b) -> add_disequality state (Hyp cid) a b @@ -254,7 +255,7 @@ let app_global_with_holes f args n = Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let t = Tacmach.New.pf_get_type_of gl fc in let t = Termops.prod_applist sigma t (Array.to_list args) in let ans = mkApp (fc, args) in diff --git a/plugins/cc/cctac.mli b/plugins/cc/cctac.mli index b4bb62be8..ef32d2b83 100644 --- a/plugins/cc/cctac.mli +++ b/plugins/cc/cctac.mli @@ -7,6 +7,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open EConstr val proof_tac: Ccproof.proof -> unit Proofview.tactic diff --git a/plugins/cc/g_congruence.ml4 b/plugins/cc/g_congruence.ml4 index 7e76854b1..43b150c34 100644 --- a/plugins/cc/g_congruence.ml4 +++ b/plugins/cc/g_congruence.ml4 @@ -8,6 +8,7 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Ltac_plugin open Cctac open Stdarg diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml index b3ab29cce..31cbc8e25 100644 --- a/plugins/derive/derive.ml +++ b/plugins/derive/derive.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Context.Named.Declaration let map_const_entry_body (f:Term.constr->Term.constr) (x:Safe_typing.private_constants Entries.const_entry_body) diff --git a/plugins/derive/derive.mli b/plugins/derive/derive.mli index 9ea876f13..3a7e7b837 100644 --- a/plugins/derive/derive.mli +++ b/plugins/derive/derive.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + (** [start_deriving f suchthat lemma] starts a proof of [suchthat] (which can contain references to [f]) in the context extended by [f:=?x]. When the proof ends, [f] is defined as the value of [?x] diff --git a/plugins/derive/g_derive.ml4 b/plugins/derive/g_derive.ml4 index deadb3b4d..445923e01 100644 --- a/plugins/derive/g_derive.ml4 +++ b/plugins/derive/g_derive.ml4 @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Stdarg (*i camlp4deps: "grammar/grammar.cma" i*) diff --git a/plugins/derive/vo.itarget b/plugins/derive/vo.itarget deleted file mode 100644 index b48098219..000000000 --- a/plugins/derive/vo.itarget +++ /dev/null @@ -1 +0,0 @@ -Derive.vo
\ No newline at end of file diff --git a/plugins/extraction/ExtrHaskellBasic.v b/plugins/extraction/ExtrHaskellBasic.v index 294d61023..d08a81da6 100644 --- a/plugins/extraction/ExtrHaskellBasic.v +++ b/plugins/extraction/ExtrHaskellBasic.v @@ -1,5 +1,7 @@ (** Extraction to Haskell : use of basic Haskell types *) +Require Coq.extraction.Extraction. + Extract Inductive bool => "Prelude.Bool" [ "Prelude.True" "Prelude.False" ]. Extract Inductive option => "Prelude.Maybe" [ "Prelude.Just" "Prelude.Nothing" ]. Extract Inductive unit => "()" [ "()" ]. diff --git a/plugins/extraction/ExtrHaskellNatInt.v b/plugins/extraction/ExtrHaskellNatInt.v index e94e7d42b..267322d9e 100644 --- a/plugins/extraction/ExtrHaskellNatInt.v +++ b/plugins/extraction/ExtrHaskellNatInt.v @@ -1,5 +1,7 @@ (** Extraction of [nat] into Haskell's [Int] *) +Require Coq.extraction.Extraction. + Require Import Arith. Require Import ExtrHaskellNatNum. diff --git a/plugins/extraction/ExtrHaskellNatInteger.v b/plugins/extraction/ExtrHaskellNatInteger.v index 038f0ed81..4c5c71f58 100644 --- a/plugins/extraction/ExtrHaskellNatInteger.v +++ b/plugins/extraction/ExtrHaskellNatInteger.v @@ -1,5 +1,7 @@ (** Extraction of [nat] into Haskell's [Integer] *) +Require Coq.extraction.Extraction. + Require Import Arith. Require Import ExtrHaskellNatNum. diff --git a/plugins/extraction/ExtrHaskellNatNum.v b/plugins/extraction/ExtrHaskellNatNum.v index 244eb85fc..fabe9a4c6 100644 --- a/plugins/extraction/ExtrHaskellNatNum.v +++ b/plugins/extraction/ExtrHaskellNatNum.v @@ -6,6 +6,8 @@ * implements [Num]. *) +Require Coq.extraction.Extraction. + Require Import Arith. Require Import EqNat. diff --git a/plugins/extraction/ExtrHaskellString.v b/plugins/extraction/ExtrHaskellString.v index 3558f4f25..ac1f6f913 100644 --- a/plugins/extraction/ExtrHaskellString.v +++ b/plugins/extraction/ExtrHaskellString.v @@ -2,6 +2,8 @@ * Special handling of ascii and strings for extraction to Haskell. *) +Require Coq.extraction.Extraction. + Require Import Ascii. Require Import String. diff --git a/plugins/extraction/ExtrHaskellZInt.v b/plugins/extraction/ExtrHaskellZInt.v index 66690851a..0345ffc4e 100644 --- a/plugins/extraction/ExtrHaskellZInt.v +++ b/plugins/extraction/ExtrHaskellZInt.v @@ -1,5 +1,7 @@ (** Extraction of [Z] into Haskell's [Int] *) +Require Coq.extraction.Extraction. + Require Import ZArith. Require Import ExtrHaskellZNum. diff --git a/plugins/extraction/ExtrHaskellZInteger.v b/plugins/extraction/ExtrHaskellZInteger.v index f192f16ee..f7f9e2f80 100644 --- a/plugins/extraction/ExtrHaskellZInteger.v +++ b/plugins/extraction/ExtrHaskellZInteger.v @@ -1,5 +1,7 @@ (** Extraction of [Z] into Haskell's [Integer] *) +Require Coq.extraction.Extraction. + Require Import ZArith. Require Import ExtrHaskellZNum. diff --git a/plugins/extraction/ExtrHaskellZNum.v b/plugins/extraction/ExtrHaskellZNum.v index cbbfda75e..4141bd203 100644 --- a/plugins/extraction/ExtrHaskellZNum.v +++ b/plugins/extraction/ExtrHaskellZNum.v @@ -6,6 +6,8 @@ * implements [Num]. *) +Require Coq.extraction.Extraction. + Require Import ZArith. Require Import EqNat. diff --git a/plugins/extraction/ExtrOcamlBasic.v b/plugins/extraction/ExtrOcamlBasic.v index d9b000c2a..dfdc49863 100644 --- a/plugins/extraction/ExtrOcamlBasic.v +++ b/plugins/extraction/ExtrOcamlBasic.v @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +Require Coq.extraction.Extraction. + (** Extraction to Ocaml : use of basic Ocaml types *) Extract Inductive bool => bool [ true false ]. diff --git a/plugins/extraction/ExtrOcamlBigIntConv.v b/plugins/extraction/ExtrOcamlBigIntConv.v index c42938c8e..78ee46085 100644 --- a/plugins/extraction/ExtrOcamlBigIntConv.v +++ b/plugins/extraction/ExtrOcamlBigIntConv.v @@ -13,6 +13,8 @@ simplifies the use of [Big_int] (it can be found in the sources of Coq). *) +Require Coq.extraction.Extraction. + Require Import Arith ZArith. Parameter bigint : Type. diff --git a/plugins/extraction/ExtrOcamlIntConv.v b/plugins/extraction/ExtrOcamlIntConv.v index 515fa52df..fcfea352a 100644 --- a/plugins/extraction/ExtrOcamlIntConv.v +++ b/plugins/extraction/ExtrOcamlIntConv.v @@ -10,6 +10,8 @@ Nota: no check that [int] values aren't generating overflows *) +Require Coq.extraction.Extraction. + Require Import Arith ZArith. Parameter int : Type. diff --git a/plugins/extraction/ExtrOcamlNatBigInt.v b/plugins/extraction/ExtrOcamlNatBigInt.v index 3149e7029..e0837be62 100644 --- a/plugins/extraction/ExtrOcamlNatBigInt.v +++ b/plugins/extraction/ExtrOcamlNatBigInt.v @@ -8,6 +8,8 @@ (** Extraction of [nat] into Ocaml's [big_int] *) +Require Coq.extraction.Extraction. + Require Import Arith Even Div2 EqNat Euclid. Require Import ExtrOcamlBasic. diff --git a/plugins/extraction/ExtrOcamlNatInt.v b/plugins/extraction/ExtrOcamlNatInt.v index 7c607f7ae..80da72d43 100644 --- a/plugins/extraction/ExtrOcamlNatInt.v +++ b/plugins/extraction/ExtrOcamlNatInt.v @@ -8,6 +8,8 @@ (** Extraction of [nat] into Ocaml's [int] *) +Require Coq.extraction.Extraction. + Require Import Arith Even Div2 EqNat Euclid. Require Import ExtrOcamlBasic. diff --git a/plugins/extraction/ExtrOcamlString.v b/plugins/extraction/ExtrOcamlString.v index 6af591eed..64ca6c85d 100644 --- a/plugins/extraction/ExtrOcamlString.v +++ b/plugins/extraction/ExtrOcamlString.v @@ -8,6 +8,8 @@ (* Extraction to Ocaml : special handling of ascii and strings *) +Require Coq.extraction.Extraction. + Require Import Ascii String. Extract Inductive ascii => char diff --git a/plugins/extraction/ExtrOcamlZBigInt.v b/plugins/extraction/ExtrOcamlZBigInt.v index c9e8eac0c..66f188c84 100644 --- a/plugins/extraction/ExtrOcamlZBigInt.v +++ b/plugins/extraction/ExtrOcamlZBigInt.v @@ -8,6 +8,8 @@ (** Extraction of [positive], [N] and [Z] into Ocaml's [big_int] *) +Require Coq.extraction.Extraction. + Require Import ZArith NArith. Require Import ExtrOcamlBasic. diff --git a/plugins/extraction/ExtrOcamlZInt.v b/plugins/extraction/ExtrOcamlZInt.v index 4d33174b3..c93cfb9d4 100644 --- a/plugins/extraction/ExtrOcamlZInt.v +++ b/plugins/extraction/ExtrOcamlZInt.v @@ -8,6 +8,8 @@ (** Extraction of [positive], [N] and [Z] into Ocaml's [int] *) +Require Coq.extraction.Extraction. + Require Import ZArith NArith. Require Import ExtrOcamlBasic. diff --git a/plugins/extraction/Extraction.v b/plugins/extraction/Extraction.v new file mode 100644 index 000000000..ab1416b1d --- /dev/null +++ b/plugins/extraction/Extraction.v @@ -0,0 +1,9 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +Declare ML Module "extraction_plugin".
\ No newline at end of file diff --git a/plugins/extraction/common.ml b/plugins/extraction/common.ml index c498eb589..e66bf7e1b 100644 --- a/plugins/extraction/common.ml +++ b/plugins/extraction/common.ml @@ -6,9 +6,11 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Pp open Util open Names +open ModPath open Namegen open Nameops open Libnames @@ -44,7 +46,7 @@ let pp_apply2 st par args = let pr_binding = function | [] -> mt () - | l -> str " " ++ prlist_with_sep (fun () -> str " ") pr_id l + | l -> str " " ++ prlist_with_sep (fun () -> str " ") Id.print l let pp_tuple_light f = function | [] -> mt () @@ -273,8 +275,8 @@ let params_ren_add, params_ren_mem = seen at this level. *) -type visible_layer = { mp : module_path; - params : module_path list; +type visible_layer = { mp : ModPath.t; + params : ModPath.t list; mutable content : Label.t KMap.t; } let pop_visible, push_visible, get_visible = diff --git a/plugins/extraction/common.mli b/plugins/extraction/common.mli index b8e95afb3..004019e16 100644 --- a/plugins/extraction/common.mli +++ b/plugins/extraction/common.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Globnames open Miniml @@ -49,20 +50,20 @@ type phase = Pre | Impl | Intf val set_phase : phase -> unit val get_phase : unit -> phase -val opened_libraries : unit -> module_path list +val opened_libraries : unit -> ModPath.t list type kind = Term | Type | Cons | Mod val pp_global : kind -> global_reference -> string -val pp_module : module_path -> string +val pp_module : ModPath.t -> string -val top_visible_mp : unit -> module_path +val top_visible_mp : unit -> ModPath.t (* In [push_visible], the [module_path list] corresponds to module parameters, the innermost one coming first in the list *) -val push_visible : module_path -> module_path list -> unit +val push_visible : ModPath.t -> ModPath.t list -> unit val pop_visible : unit -> unit -val get_duplicate : module_path -> Label.t -> string option +val get_duplicate : ModPath.t -> Label.t -> string option type reset_kind = AllButExternal | Everything @@ -72,7 +73,7 @@ val set_keywords : Id.Set.t -> unit (** For instance: [mk_ind "Coq.Init.Datatypes" "nat"] *) -val mk_ind : string -> string -> mutual_inductive +val mk_ind : string -> string -> MutInd.t (** Special hack for constants of type Ascii.ascii : if an [Extract Inductive ascii => char] has been declared, then diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml index 2c85b185c..40ef6601d 100644 --- a/plugins/extraction/extract_env.ml +++ b/plugins/extraction/extract_env.ml @@ -6,10 +6,12 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Miniml open Term open Declarations open Names +open ModPath open Libnames open Globnames open Pp @@ -27,13 +29,13 @@ open Common let toplevel_env () = let get_reference = function | (_,kn), Lib.Leaf o -> - let mp,_,l = repr_kn kn in + let mp,_,l = KerName.repr kn in begin match Libobject.object_tag o with | "CONSTANT" -> - let constant = Global.lookup_constant (constant_of_kn kn) in + let constant = Global.lookup_constant (Constant.make1 kn) in Some (l, SFBconst constant) | "INDUCTIVE" -> - let inductive = Global.lookup_mind (mind_of_kn kn) in + let inductive = Global.lookup_mind (MutInd.make1 kn) in Some (l, SFBmind inductive) | "MODULE" -> let modl = Global.lookup_module (MPdot (mp, l)) in @@ -72,21 +74,21 @@ module type VISIT = sig (* Add the module_path and all its prefixes to the mp visit list. We'll keep all fields of these modules. *) - val add_mp_all : module_path -> unit + val add_mp_all : ModPath.t -> unit (* Add reference / ... in the visit lists. These functions silently add the mp of their arg in the mp list *) val add_ref : global_reference -> unit - val add_kn : kernel_name -> unit + val add_kn : KerName.t -> unit val add_decl_deps : ml_decl -> unit val add_spec_deps : ml_spec -> unit (* Test functions: is a particular object a needed dependency for the current extraction ? *) - val needed_ind : mutual_inductive -> bool - val needed_cst : constant -> bool - val needed_mp : module_path -> bool - val needed_mp_all : module_path -> bool + val needed_ind : MutInd.t -> bool + val needed_cst : Constant.t -> bool + val needed_mp : ModPath.t -> bool + val needed_mp_all : ModPath.t -> bool end module Visit : VISIT = struct @@ -101,8 +103,8 @@ module Visit : VISIT = struct v.kn <- KNset.empty; v.mp <- MPset.empty; v.mp_all <- MPset.empty - let needed_ind i = KNset.mem (user_mind i) v.kn - let needed_cst c = KNset.mem (user_con c) v.kn + let needed_ind i = KNset.mem (MutInd.user i) v.kn + let needed_cst c = KNset.mem (Constant.user c) v.kn let needed_mp mp = MPset.mem mp v.mp || MPset.mem mp v.mp_all let needed_mp_all mp = MPset.mem mp v.mp_all let add_mp mp = @@ -111,10 +113,10 @@ module Visit : VISIT = struct check_loaded_modfile mp; v.mp <- MPset.union (prefixes_mp mp) v.mp; v.mp_all <- MPset.add mp v.mp_all - let add_kn kn = v.kn <- KNset.add kn v.kn; add_mp (modpath kn) + let add_kn kn = v.kn <- KNset.add kn v.kn; add_mp (KerName.modpath kn) let add_ref = function - | ConstRef c -> add_kn (user_con c) - | IndRef (ind,_) | ConstructRef ((ind,_),_) -> add_kn (user_mind ind) + | ConstRef c -> add_kn (Constant.user c) + | IndRef (ind,_) | ConstructRef ((ind,_),_) -> add_kn (MutInd.user ind) | VarRef _ -> assert false let add_decl_deps = decl_iter_references add_ref add_ref add_ref let add_spec_deps = spec_iter_references add_ref add_ref add_ref diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli index 90f4f911b..4f0ed953c 100644 --- a/plugins/extraction/extract_env.mli +++ b/plugins/extraction/extract_env.mli @@ -8,6 +8,7 @@ (*s This module declares the extraction commands. *) +open API open Names open Libnames open Globnames @@ -20,12 +21,12 @@ val extraction_library : bool -> Id.t -> unit (* For debug / external output via coqtop.byte + Drop : *) val mono_environment : - global_reference list -> module_path list -> Miniml.ml_structure + global_reference list -> ModPath.t list -> Miniml.ml_structure (* Used by the Relation Extraction plugin *) val print_one_decl : - Miniml.ml_structure -> module_path -> Miniml.ml_decl -> Pp.std_ppcmds + Miniml.ml_structure -> ModPath.t -> Miniml.ml_decl -> Pp.std_ppcmds (* Used by Extraction Compute *) diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 92ece7ccf..2b7199a76 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -7,6 +7,7 @@ (************************************************************************) (*i*) +open API open Util open Names open Term @@ -31,7 +32,7 @@ open Context.Rel.Declaration exception I of inductive_kind (* A set of all fixpoint functions currently being extracted *) -let current_fixpoints = ref ([] : constant list) +let current_fixpoints = ref ([] : Constant.t list) let none = Evd.empty @@ -255,7 +256,7 @@ let rec extract_type env db j c args = let reason = if lvl == TypeScheme then Ktype else Kprop in Tarr (Tdummy reason, mld))) | Sort _ -> Tdummy Ktype (* The two logical cases. *) - | _ when sort_of env (applist (c, args)) == InProp -> Tdummy Kprop + | _ when sort_of env (applistc c args) == InProp -> Tdummy Kprop | Rel n -> (match lookup_rel n env with | LocalDef (_,t,_) -> extract_type env db j (lift n t) args @@ -276,7 +277,7 @@ let rec extract_type env db j c args = | Undef _ | OpaqueDef _ -> mlt | Def _ when is_custom r -> mlt | Def lbody -> - let newc = applist (Mod_subst.force_constr lbody, args) in + let newc = applistc (Mod_subst.force_constr lbody) args in let mlt' = extract_type env db j newc [] in (* ML type abbreviations interact badly with Coq *) (* reduction, so [mlt] and [mlt'] might be different: *) @@ -290,7 +291,7 @@ let rec extract_type env db j c args = | Undef _ | OpaqueDef _ -> Tunknown (* Brutal approx ... *) | Def lbody -> (* We try to reduce. *) - let newc = applist (Mod_subst.force_constr lbody, args) in + let newc = applistc (Mod_subst.force_constr lbody) args in extract_type env db j newc [])) | Ind ((kn,i),u) -> let s = (extract_ind env kn).ind_packets.(i).ip_sign in @@ -361,14 +362,14 @@ and extract_really_ind env kn mib = (cf Vector and bug #2570) *) let equiv = if lang () != Ocaml || - (not (modular ()) && at_toplevel (mind_modpath kn)) || - KerName.equal (canonical_mind kn) (user_mind kn) + (not (modular ()) && at_toplevel (MutInd.modpath kn)) || + KerName.equal (MutInd.canonical kn) (MutInd.user kn) then NoEquiv else begin - ignore (extract_ind env (mind_of_kn (canonical_mind kn))); - Equiv (canonical_mind kn) + ignore (extract_ind env (MutInd.make1 (MutInd.canonical kn))); + Equiv (MutInd.canonical kn) end in (* Everything concerning parameters. *) @@ -864,7 +865,7 @@ let decomp_lams_eta_n n m env c t = (* we'd better keep rels' as long as possible. *) let rels = (List.firstn d rels) @ rels' in let eta_args = List.rev_map mkRel (List.interval 1 d) in - rels, applist (lift d c,eta_args) + rels, applistc (lift d c) eta_args (* Let's try to identify some situation where extracted code will allow generalisation of type variables *) diff --git a/plugins/extraction/extraction.mli b/plugins/extraction/extraction.mli index cdda777a6..26268fb17 100644 --- a/plugins/extraction/extraction.mli +++ b/plugins/extraction/extraction.mli @@ -8,24 +8,25 @@ (*s Extraction from Coq terms to Miniml. *) +open API open Names open Term open Declarations open Environ open Miniml -val extract_constant : env -> constant -> constant_body -> ml_decl +val extract_constant : env -> Constant.t -> constant_body -> ml_decl -val extract_constant_spec : env -> constant -> constant_body -> ml_spec +val extract_constant_spec : env -> Constant.t -> constant_body -> ml_spec (** For extracting "module ... with ..." declaration *) val extract_with_type : env -> constr -> ( Id.t list * ml_type ) option val extract_fixpoint : - env -> constant array -> (constr, types) prec_declaration -> ml_decl + env -> Constant.t array -> (constr, types) prec_declaration -> ml_decl -val extract_inductive : env -> mutual_inductive -> ml_ind +val extract_inductive : env -> MutInd.t -> ml_ind (** For extraction compute *) diff --git a/plugins/extraction/g_extraction.ml4 b/plugins/extraction/g_extraction.ml4 index 3ed959cf2..76b435410 100644 --- a/plugins/extraction/g_extraction.ml4 +++ b/plugins/extraction/g_extraction.ml4 @@ -8,6 +8,9 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API.Pcoq.Prim + DECLARE PLUGIN "extraction_plugin" (* ML names *) @@ -15,10 +18,8 @@ DECLARE PLUGIN "extraction_plugin" open Ltac_plugin open Genarg open Stdarg -open Pcoq.Prim open Pp open Names -open Nameops open Table open Extract_env @@ -33,7 +34,7 @@ END let pr_int_or_id _ _ _ = function | ArgInt i -> int i - | ArgId id -> pr_id id + | ArgId id -> Id.print id ARGUMENT EXTEND int_or_id PRINTED BY pr_int_or_id diff --git a/plugins/extraction/haskell.ml b/plugins/extraction/haskell.ml index eb13fd675..4bd207a98 100644 --- a/plugins/extraction/haskell.ml +++ b/plugins/extraction/haskell.ml @@ -8,11 +8,11 @@ (*s Production of Haskell syntax. *) +open API open Pp open CErrors open Util open Names -open Nameops open Globnames open Table open Miniml @@ -93,7 +93,7 @@ let preamble mod_name comment used_modules usf = let pp_abst = function | [] -> (mt ()) | l -> (str "\\" ++ - prlist_with_sep (fun () -> (str " ")) pr_id l ++ + prlist_with_sep (fun () -> (str " ")) Id.print l ++ str " ->" ++ spc ()) (*s The pretty-printer for haskell syntax *) @@ -109,7 +109,7 @@ let rec pp_type par vl t = let rec pp_rec par = function | Tmeta _ | Tvar' _ -> assert false | Tvar i -> - (try pr_id (List.nth vl (pred i)) + (try Id.print (List.nth vl (pred i)) with Failure _ -> (str "a" ++ int i)) | Tglob (r,[]) -> pp_global Type r | Tglob (IndRef(kn,0),l) @@ -148,7 +148,7 @@ let rec pp_expr par env args = (* Try to survive to the occurrence of a Dummy rel. TODO: we should get rid of this hack (cf. #592) *) let id = if Id.equal id dummy_name then Id.of_string "__" else id in - apply (pr_id id) + apply (Id.print id) | MLapp (f,args') -> let stl = List.map (pp_expr true env []) args' in pp_expr par env (stl @ args) f @@ -159,7 +159,7 @@ let rec pp_expr par env args = apply2 st | MLletin (id,a1,a2) -> let i,env' = push_vars [id_of_mlid id] env in - let pp_id = pr_id (List.hd i) + let pp_id = Id.print (List.hd i) and pp_a1 = pp_expr false env [] a1 and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in let pp_def = @@ -223,10 +223,10 @@ and pp_cons_pat par r ppl = and pp_gen_pat par ids env = function | Pcons (r,l) -> pp_cons_pat par r (List.map (pp_gen_pat true ids env) l) - | Pusual r -> pp_cons_pat par r (List.map pr_id ids) + | Pusual r -> pp_cons_pat par r (List.map Id.print ids) | Ptuple l -> pp_boxed_tuple (pp_gen_pat false ids env) l | Pwild -> str "_" - | Prel n -> pr_id (get_db_name n env) + | Prel n -> Id.print (get_db_name n env) and pp_one_pat env (ids,p,t) = let ids',env' = push_vars (List.rev_map id_of_mlid ids) env in @@ -251,10 +251,10 @@ and pp_fix par env i (ids,bl) args = (v 0 (v 1 (str "let {" ++ fnl () ++ prvect_with_sep (fun () -> str ";" ++ fnl ()) - (fun (fi,ti) -> pp_function env (pr_id fi) ti) + (fun (fi,ti) -> pp_function env (Id.print fi) ti) (Array.map2 (fun a b -> a,b) ids bl) ++ str "}") ++ - fnl () ++ str "in " ++ pp_apply (pr_id ids.(i)) false args)) + fnl () ++ str "in " ++ pp_apply (Id.print ids.(i)) false args)) and pp_function env f t = let bl,t' = collect_lams t in @@ -266,19 +266,19 @@ and pp_function env f t = (*s Pretty-printing of inductive types declaration. *) let pp_logical_ind packet = - pp_comment (pr_id packet.ip_typename ++ str " : logical inductive") ++ + pp_comment (Id.print packet.ip_typename ++ str " : logical inductive") ++ pp_comment (str "with constructors : " ++ - prvect_with_sep spc pr_id packet.ip_consnames) + prvect_with_sep spc Id.print packet.ip_consnames) let pp_singleton kn packet = let name = pp_global Type (IndRef (kn,0)) in let l = rename_tvars keywords packet.ip_vars in hov 2 (str "type " ++ name ++ spc () ++ - prlist_with_sep spc pr_id l ++ + prlist_with_sep spc Id.print l ++ (if not (List.is_empty l) then str " " else mt ()) ++ str "=" ++ spc () ++ pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ pp_comment (str "singleton inductive, whose constructor was " ++ - pr_id packet.ip_consnames.(0))) + Id.print packet.ip_consnames.(0))) let pp_one_ind ip pl cv = let pl = rename_tvars keywords pl in @@ -330,7 +330,7 @@ let pp_decl = function let ids,s = find_type_custom r in prlist (fun id -> str (id^" ")) ids ++ str "=" ++ spc () ++ str s with Not_found -> - prlist (fun id -> pr_id id ++ str " ") l ++ + prlist (fun id -> Id.print id ++ str " ") l ++ if t == Taxiom then str "= () -- AXIOM TO BE REALIZED" ++ fnl () else str "=" ++ spc () ++ pp_type false l t in diff --git a/plugins/extraction/json.ml b/plugins/extraction/json.ml index e43c47d05..1bf19f186 100644 --- a/plugins/extraction/json.ml +++ b/plugins/extraction/json.ml @@ -1,3 +1,4 @@ +open API open Pp open Util open Names diff --git a/plugins/extraction/miniml.mli b/plugins/extraction/miniml.mli index db3361522..ec28f4996 100644 --- a/plugins/extraction/miniml.mli +++ b/plugins/extraction/miniml.mli @@ -8,6 +8,7 @@ (*s Target language for extraction: a core ML called MiniML. *) +open API open Pp open Names open Globnames @@ -82,7 +83,7 @@ type ml_ind_packet = { type equiv = | NoEquiv - | Equiv of kernel_name + | Equiv of KerName.t | RenEquiv of string type ml_ind = { @@ -137,13 +138,13 @@ and ml_pattern = (*s ML declarations. *) type ml_decl = - | Dind of mutual_inductive * ml_ind + | Dind of MutInd.t * ml_ind | Dtype of global_reference * Id.t list * ml_type | Dterm of global_reference * ml_ast * ml_type | Dfix of global_reference array * ml_ast array * ml_type array type ml_spec = - | Sind of mutual_inductive * ml_ind + | Sind of MutInd.t * ml_ind | Stype of global_reference * Id.t list * ml_type option | Sval of global_reference * ml_type @@ -153,14 +154,14 @@ type ml_specif = | Smodtype of ml_module_type and ml_module_type = - | MTident of module_path + | MTident of ModPath.t | MTfunsig of MBId.t * ml_module_type * ml_module_type - | MTsig of module_path * ml_module_sig + | MTsig of ModPath.t * ml_module_sig | MTwith of ml_module_type * ml_with_declaration and ml_with_declaration = | ML_With_type of Id.t list * Id.t list * ml_type - | ML_With_module of Id.t list * module_path + | ML_With_module of Id.t list * ModPath.t and ml_module_sig = (Label.t * ml_specif) list @@ -170,9 +171,9 @@ type ml_structure_elem = | SEmodtype of ml_module_type and ml_module_expr = - | MEident of module_path + | MEident of ModPath.t | MEfunctor of MBId.t * ml_module_type * ml_module_expr - | MEstruct of module_path * ml_module_structure + | MEstruct of ModPath.t * ml_module_structure | MEapply of ml_module_expr * ml_module_expr and ml_module_structure = (Label.t * ml_structure_elem) list @@ -184,9 +185,9 @@ and ml_module = (* NB: we do not translate the [mod_equiv] field, since [mod_equiv = mp] implies that [mod_expr = MEBident mp]. Same with [msb_equiv]. *) -type ml_structure = (module_path * ml_module_structure) list +type ml_structure = (ModPath.t * ml_module_structure) list -type ml_signature = (module_path * ml_module_sig) list +type ml_signature = (ModPath.t * ml_module_sig) list type ml_flat_structure = ml_structure_elem list @@ -202,10 +203,10 @@ type language_descr = { (* Concerning the source file *) file_suffix : string; - file_naming : module_path -> string; + file_naming : ModPath.t -> string; (* the second argument is a comment to add to the preamble *) preamble : - Id.t -> std_ppcmds option -> module_path list -> unsafe_needs -> + Id.t -> std_ppcmds option -> ModPath.t list -> unsafe_needs -> std_ppcmds; pp_struct : ml_structure -> std_ppcmds; @@ -213,7 +214,7 @@ type language_descr = { sig_suffix : string option; (* the second argument is a comment to add to the preamble *) sig_preamble : - Id.t -> std_ppcmds option -> module_path list -> unsafe_needs -> + Id.t -> std_ppcmds option -> ModPath.t list -> unsafe_needs -> std_ppcmds; pp_sig : ml_signature -> std_ppcmds; diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml index 402fe4ffe..3a70a5020 100644 --- a/plugins/extraction/mlutil.ml +++ b/plugins/extraction/mlutil.ml @@ -7,6 +7,7 @@ (************************************************************************) (*i*) +open API open Util open Names open Libnames @@ -28,9 +29,9 @@ let dummy_name = Id.of_string "_" let anonymous = Id anonymous_name let id_of_name = function - | Anonymous -> anonymous_name - | Name id when Id.equal id dummy_name -> anonymous_name - | Name id -> id + | Name.Anonymous -> anonymous_name + | Name.Name id when Id.equal id dummy_name -> anonymous_name + | Name.Name id -> id let id_of_mlid = function | Dummy -> dummy_name @@ -1487,7 +1488,7 @@ let inline_test r t = let con_of_string s = let d, id = Libnames.split_dirpath (dirpath_of_string s) in - Constant.make2 (MPfile d) (Label.of_id id) + Constant.make2 (ModPath.MPfile d) (Label.of_id id) let manual_inline_set = List.fold_right (fun x -> Cset_env.add (con_of_string x)) diff --git a/plugins/extraction/mlutil.mli b/plugins/extraction/mlutil.mli index c66755249..6924dc9ff 100644 --- a/plugins/extraction/mlutil.mli +++ b/plugins/extraction/mlutil.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Globnames open Miniml @@ -48,7 +49,7 @@ end (*s Utility functions over ML types without meta *) -val type_mem_kn : mutual_inductive -> ml_type -> bool +val type_mem_kn : MutInd.t -> ml_type -> bool val type_maxvar : ml_type -> int diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml index b67b9931e..6c38813e4 100644 --- a/plugins/extraction/modutil.ml +++ b/plugins/extraction/modutil.ml @@ -6,7 +6,9 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names +open ModPath open Globnames open CErrors open Util @@ -110,7 +112,7 @@ let ind_iter_references do_term do_cons do_type kn ind = do_type (IndRef ip); if lang () == Ocaml then (match ind.ind_equiv with - | Miniml.Equiv kne -> do_type (IndRef (mind_of_kn kne, snd ip)); + | Miniml.Equiv kne -> do_type (IndRef (MutInd.make1 kne, snd ip)); | _ -> ()); Array.iteri (fun j -> cons_iter (ip,j+1)) p.ip_types in diff --git a/plugins/extraction/modutil.mli b/plugins/extraction/modutil.mli index dc8708249..9a67baa96 100644 --- a/plugins/extraction/modutil.mli +++ b/plugins/extraction/modutil.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Globnames open Miniml @@ -25,7 +26,7 @@ val signature_of_structure : ml_structure -> ml_signature val mtyp_of_mexpr : ml_module_expr -> ml_module_type -val msid_of_mt : ml_module_type -> module_path +val msid_of_mt : ml_module_type -> ModPath.t val get_decl_in_structure : global_reference -> ml_structure -> ml_decl @@ -36,5 +37,5 @@ val get_decl_in_structure : global_reference -> ml_structure -> ml_decl optimizations. The first argument is the list of objects we want to appear. *) -val optimize_struct : global_reference list * module_path list -> +val optimize_struct : global_reference list * ModPath.t list -> ml_structure -> ml_structure diff --git a/plugins/extraction/ocaml.ml b/plugins/extraction/ocaml.ml index 4399fc561..16feaf4d6 100644 --- a/plugins/extraction/ocaml.ml +++ b/plugins/extraction/ocaml.ml @@ -8,11 +8,12 @@ (*s Production of Ocaml syntax. *) +open API open Pp open CErrors open Util open Names -open Nameops +open ModPath open Globnames open Table open Miniml @@ -28,7 +29,7 @@ let pp_tvar id = str ("'" ^ Id.to_string id) let pp_abst = function | [] -> mt () | l -> - str "fun " ++ prlist_with_sep (fun () -> str " ") pr_id l ++ + str "fun " ++ prlist_with_sep (fun () -> str " ") Id.print l ++ str " ->" ++ spc () let pp_parameters l = @@ -182,7 +183,7 @@ let rec pp_expr par env args = (* Try to survive to the occurrence of a Dummy rel. TODO: we should get rid of this hack (cf. #592) *) let id = if Id.equal id dummy_name then Id.of_string "__" else id in - apply (pr_id id) + apply (Id.print id) | MLapp (f,args') -> let stl = List.map (pp_expr true env []) args' in pp_expr par env (stl @ args) f @@ -194,7 +195,7 @@ let rec pp_expr par env args = apply2 st | MLletin (id,a1,a2) -> let i,env' = push_vars [id_of_mlid id] env in - let pp_id = pr_id (List.hd i) + let pp_id = Id.print (List.hd i) and pp_a1 = pp_expr false env [] a1 and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in hv 0 (apply2 (pp_letin pp_id pp_a1 pp_a2)) @@ -330,10 +331,10 @@ and pp_cons_pat r ppl = and pp_gen_pat ids env = function | Pcons (r, l) -> pp_cons_pat r (List.map (pp_gen_pat ids env) l) - | Pusual r -> pp_cons_pat r (List.map pr_id ids) + | Pusual r -> pp_cons_pat r (List.map Id.print ids) | Ptuple l -> pp_boxed_tuple (pp_gen_pat ids env) l | Pwild -> str "_" - | Prel n -> pr_id (get_db_name n env) + | Prel n -> Id.print (get_db_name n env) and pp_ifthenelse env expr pv = match pv with | [|([],tru,the);([],fal,els)|] when @@ -372,7 +373,7 @@ and pp_function env t = v 0 (pp_pat env' pv) else pr_binding (List.rev bl) ++ - str " = match " ++ pr_id (List.hd bl) ++ str " with" ++ fnl () ++ + str " = match " ++ Id.print (List.hd bl) ++ str " with" ++ fnl () ++ v 0 (pp_pat env' pv) | _ -> pr_binding (List.rev bl) ++ @@ -387,10 +388,10 @@ and pp_fix par env i (ids,bl) args = (v 0 (str "let rec " ++ prvect_with_sep (fun () -> fnl () ++ str "and ") - (fun (fi,ti) -> pr_id fi ++ pp_function env ti) + (fun (fi,ti) -> Id.print fi ++ pp_function env ti) (Array.map2 (fun id b -> (id,b)) ids bl) ++ fnl () ++ - hov 2 (str "in " ++ pp_apply (pr_id ids.(i)) false args))) + hov 2 (str "in " ++ pp_apply (Id.print ids.(i)) false args))) (* Ad-hoc double-newline in v boxes, with enough negative whitespace to avoid indenting the intermediate blank line *) @@ -431,7 +432,7 @@ let pp_Dfix (rv,c,t) = let pp_equiv param_list name = function | NoEquiv, _ -> mt () | Equiv kn, i -> - str " = " ++ pp_parameters param_list ++ pp_global Type (IndRef (mind_of_kn kn,i)) + str " = " ++ pp_parameters param_list ++ pp_global Type (IndRef (MutInd.make1 kn,i)) | RenEquiv ren, _ -> str " = " ++ pp_parameters param_list ++ str (ren^".") ++ name @@ -451,10 +452,10 @@ let pp_one_ind prefix ip_equiv pl name cnames ctyps = else fnl () ++ v 0 (prvecti pp_constructor ctyps) let pp_logical_ind packet = - pp_comment (pr_id packet.ip_typename ++ str " : logical inductive") ++ + pp_comment (Id.print packet.ip_typename ++ str " : logical inductive") ++ fnl () ++ pp_comment (str "with constructors : " ++ - prvect_with_sep spc pr_id packet.ip_consnames) ++ + prvect_with_sep spc Id.print packet.ip_consnames) ++ fnl () let pp_singleton kn packet = @@ -463,7 +464,7 @@ let pp_singleton kn packet = hov 2 (str "type " ++ pp_parameters l ++ name ++ str " =" ++ spc () ++ pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ pp_comment (str "singleton inductive, whose constructor was " ++ - pr_id packet.ip_consnames.(0))) + Id.print packet.ip_consnames.(0))) let pp_record kn fields ip_equiv packet = let ind = IndRef (kn,0) in diff --git a/plugins/extraction/scheme.ml b/plugins/extraction/scheme.ml index 3c81564e3..55168cc29 100644 --- a/plugins/extraction/scheme.ml +++ b/plugins/extraction/scheme.ml @@ -8,6 +8,7 @@ (*s Production of Scheme syntax. *) +open API open Pp open CErrors open Util diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml index 29dd8ff4f..b82c5257e 100644 --- a/plugins/extraction/table.ml +++ b/plugins/extraction/table.ml @@ -6,10 +6,11 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names +open ModPath open Term open Declarations -open Nameops open Namegen open Libobject open Goptions @@ -35,14 +36,14 @@ module Refset' = Refset_env let occur_kn_in_ref kn = function | IndRef (kn',_) - | ConstructRef ((kn',_),_) -> Names.eq_mind kn kn' + | ConstructRef ((kn',_),_) -> MutInd.equal kn kn' | ConstRef _ -> false | VarRef _ -> assert false let repr_of_r = function - | ConstRef kn -> repr_con kn + | ConstRef kn -> Constant.repr3 kn | IndRef (kn,_) - | ConstructRef ((kn,_),_) -> repr_mind kn + | ConstructRef ((kn,_),_) -> MutInd.repr3 kn | VarRef _ -> assert false let modpath_of_r r = @@ -64,7 +65,7 @@ let raw_string_of_modfile = function | _ -> assert false let is_toplevel mp = - ModPath.equal mp initial_path || ModPath.equal mp (Lib.current_mp ()) + ModPath.equal mp ModPath.initial || ModPath.equal mp (Lib.current_mp ()) let at_toplevel mp = is_modfile mp || is_toplevel mp @@ -264,8 +265,8 @@ let safe_basename_of_global r = anomaly (Pp.str "Inductive object unknown to extraction and not globally visible.") in match r with - | ConstRef kn -> Label.to_id (con_label kn) - | IndRef (kn,0) -> Label.to_id (mind_label kn) + | ConstRef kn -> Label.to_id (Constant.label kn) + | IndRef (kn,0) -> Label.to_id (MutInd.label kn) | IndRef (kn,i) -> (try (unsafe_lookup_ind kn).ind_packets.(i).ip_typename with Not_found -> last_chance r) @@ -286,8 +287,8 @@ let safe_pr_long_global r = try Printer.pr_global r with Not_found -> match r with | ConstRef kn -> - let mp,_,l = repr_con kn in - str ((string_of_mp mp)^"."^(Label.to_string l)) + let mp,_,l = Constant.repr3 kn in + str ((ModPath.to_string mp)^"."^(Label.to_string l)) | _ -> assert false let pr_long_mp mp = @@ -416,7 +417,7 @@ let error_singleton_become_prop id og = str " (or in its mutual block)" | None -> mt () in - err (str "The informative inductive type " ++ pr_id id ++ + err (str "The informative inductive type " ++ Id.print id ++ str " has a Prop instance" ++ loc ++ str "." ++ fnl () ++ str "This happens when a sort-polymorphic singleton inductive type\n" ++ str "has logical parameters, such as (I,I) : (True * True) : Prop.\n" ++ @@ -721,7 +722,7 @@ let add_implicits r l = let i = List.index Name.equal (Name id) names in Int.Set.add i s with Not_found -> - err (str "No argument " ++ pr_id id ++ str " for " ++ + err (str "No argument " ++ Id.print id ++ str " for " ++ safe_pr_global r) in let ints = List.fold_left add_arg Int.Set.empty l in @@ -799,7 +800,7 @@ let extraction_blacklist l = (* Printing part *) let print_extraction_blacklist () = - prlist_with_sep fnl pr_id (Id.Set.elements !blacklist_table) + prlist_with_sep fnl Id.print (Id.Set.elements !blacklist_table) (* Reset part *) diff --git a/plugins/extraction/table.mli b/plugins/extraction/table.mli index 15a08756c..cfe75bf4e 100644 --- a/plugins/extraction/table.mli +++ b/plugins/extraction/table.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Libnames open Globnames @@ -21,22 +22,22 @@ val safe_basename_of_global : global_reference -> Id.t val warning_axioms : unit -> unit val warning_opaques : bool -> unit -val warning_ambiguous_name : ?loc:Loc.t -> qualid * module_path * global_reference -> unit +val warning_ambiguous_name : ?loc:Loc.t -> qualid * ModPath.t * global_reference -> unit val warning_id : string -> unit val error_axiom_scheme : global_reference -> int -> 'a val error_constant : global_reference -> 'a val error_inductive : global_reference -> 'a val error_nb_cons : unit -> 'a -val error_module_clash : module_path -> module_path -> 'a -val error_no_module_expr : module_path -> 'a +val error_module_clash : ModPath.t -> ModPath.t -> 'a +val error_no_module_expr : ModPath.t -> 'a val error_singleton_become_prop : Id.t -> global_reference option -> 'a val error_unknown_module : qualid -> 'a val error_scheme : unit -> 'a val error_not_visible : global_reference -> 'a -val error_MPfile_as_mod : module_path -> bool -> 'a +val error_MPfile_as_mod : ModPath.t -> bool -> 'a val check_inside_module : unit -> unit val check_inside_section : unit -> unit -val check_loaded_modfile : module_path -> unit +val check_loaded_modfile : ModPath.t -> unit val msg_of_implicit : kill_reason -> string val err_or_warn_remaining_implicit : kill_reason -> unit @@ -44,22 +45,22 @@ val info_file : string -> unit (*s utilities about [module_path] and [kernel_names] and [global_reference] *) -val occur_kn_in_ref : mutual_inductive -> global_reference -> bool -val repr_of_r : global_reference -> module_path * DirPath.t * Label.t -val modpath_of_r : global_reference -> module_path +val occur_kn_in_ref : MutInd.t -> global_reference -> bool +val repr_of_r : global_reference -> ModPath.t * DirPath.t * Label.t +val modpath_of_r : global_reference -> ModPath.t val label_of_r : global_reference -> Label.t -val base_mp : module_path -> module_path -val is_modfile : module_path -> bool -val string_of_modfile : module_path -> string -val file_of_modfile : module_path -> string -val is_toplevel : module_path -> bool -val at_toplevel : module_path -> bool -val mp_length : module_path -> int -val prefixes_mp : module_path -> MPset.t +val base_mp : ModPath.t -> ModPath.t +val is_modfile : ModPath.t -> bool +val string_of_modfile : ModPath.t -> string +val file_of_modfile : ModPath.t -> string +val is_toplevel : ModPath.t -> bool +val at_toplevel : ModPath.t -> bool +val mp_length : ModPath.t -> int +val prefixes_mp : ModPath.t -> MPset.t val common_prefix_from_list : - module_path -> module_path list -> module_path option -val get_nth_label_mp : int -> module_path -> Label.t -val labels_of_ref : global_reference -> module_path * Label.t list + ModPath.t -> ModPath.t list -> ModPath.t option +val get_nth_label_mp : int -> ModPath.t -> Label.t +val labels_of_ref : global_reference -> ModPath.t * Label.t list (*s Some table-related operations *) @@ -71,16 +72,16 @@ val labels_of_ref : global_reference -> module_path * Label.t list [mutual_inductive_body] as checksum. In both case, we should ideally also check the env *) -val add_typedef : constant -> constant_body -> ml_type -> unit -val lookup_typedef : constant -> constant_body -> ml_type option +val add_typedef : Constant.t -> constant_body -> ml_type -> unit +val lookup_typedef : Constant.t -> constant_body -> ml_type option -val add_cst_type : constant -> constant_body -> ml_schema -> unit -val lookup_cst_type : constant -> constant_body -> ml_schema option +val add_cst_type : Constant.t -> constant_body -> ml_schema -> unit +val lookup_cst_type : Constant.t -> constant_body -> ml_schema option -val add_ind : mutual_inductive -> mutual_inductive_body -> ml_ind -> unit -val lookup_ind : mutual_inductive -> mutual_inductive_body -> ml_ind option +val add_ind : MutInd.t -> mutual_inductive_body -> ml_ind -> unit +val lookup_ind : MutInd.t -> mutual_inductive_body -> ml_ind option -val add_inductive_kind : mutual_inductive -> inductive_kind -> unit +val add_inductive_kind : MutInd.t -> inductive_kind -> unit val is_coinductive : global_reference -> bool val is_coinductive_type : ml_type -> bool (* What are the fields of a record (empty for a non-record) *) @@ -88,10 +89,10 @@ val get_record_fields : global_reference -> global_reference option list val record_fields_of_type : ml_type -> global_reference option list -val add_recursors : Environ.env -> mutual_inductive -> unit +val add_recursors : Environ.env -> MutInd.t -> unit val is_recursor : global_reference -> bool -val add_projection : int -> constant -> inductive -> unit +val add_projection : int -> Constant.t -> inductive -> unit val is_projection : global_reference -> bool val projection_arity : global_reference -> int val projection_info : global_reference -> inductive * int (* arity *) diff --git a/plugins/extraction/vo.itarget b/plugins/extraction/vo.itarget deleted file mode 100644 index 9c30c5eb3..000000000 --- a/plugins/extraction/vo.itarget +++ /dev/null @@ -1,16 +0,0 @@ -ExtrHaskellBasic.vo -ExtrHaskellNatNum.vo -ExtrHaskellNatInt.vo -ExtrHaskellNatInteger.vo -ExtrHaskellZNum.vo -ExtrHaskellZInt.vo -ExtrHaskellZInteger.vo -ExtrHaskellString.vo -ExtrOcamlBasic.vo -ExtrOcamlIntConv.vo -ExtrOcamlBigIntConv.vo -ExtrOcamlNatInt.vo -ExtrOcamlNatBigInt.vo -ExtrOcamlZInt.vo -ExtrOcamlZBigInt.vo -ExtrOcamlString.vo diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index 9900792ca..314a2b2f9 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Hipattern open Names open Term diff --git a/plugins/firstorder/formula.mli b/plugins/firstorder/formula.mli index 3f438c04a..a31de5e61 100644 --- a/plugins/firstorder/formula.mli +++ b/plugins/firstorder/formula.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Term open EConstr open Globnames diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4 index e3fab6d01..139baaeb3 100644 --- a/plugins/firstorder/g_ground.ml4 +++ b/plugins/firstorder/g_ground.ml4 @@ -8,6 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API open Ltac_plugin open Formula open Sequent diff --git a/plugins/firstorder/ground.ml b/plugins/firstorder/ground.ml index 0fa3089e7..a5a81bb16 100644 --- a/plugins/firstorder/ground.ml +++ b/plugins/firstorder/ground.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Ltac_plugin open Formula open Sequent diff --git a/plugins/firstorder/ground.mli b/plugins/firstorder/ground.mli index 4fd1e38a2..aaf79ae88 100644 --- a/plugins/firstorder/ground.mli +++ b/plugins/firstorder/ground.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + val ground_tac: unit Proofview.tactic -> ((Sequent.t -> unit Proofview.tactic) -> unit Proofview.tactic) -> unit Proofview.tactic diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index e1d765a42..92372fe29 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Unify open Rules open CErrors diff --git a/plugins/firstorder/instances.mli b/plugins/firstorder/instances.mli index 47550f314..b0e4b2690 100644 --- a/plugins/firstorder/instances.mli +++ b/plugins/firstorder/instances.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Globnames open Rules diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index b7fe25a32..72ede1f7d 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open CErrors open Util open Names diff --git a/plugins/firstorder/rules.mli b/plugins/firstorder/rules.mli index fb2173083..682047075 100644 --- a/plugins/firstorder/rules.mli +++ b/plugins/firstorder/rules.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Term open EConstr open Names diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index 826afc35b..435ca1986 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term +open API open EConstr open CErrors open Util @@ -57,11 +57,11 @@ end module OrderedConstr= struct - type t=Constr.t - let compare=constr_ord + type t=Term.constr + let compare=Term.compare end -type h_item = global_reference * (int*Constr.t) option +type h_item = global_reference * (int*Term.constr) option module Hitem= struct diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli index 6ed251f34..e24eca7cb 100644 --- a/plugins/firstorder/sequent.mli +++ b/plugins/firstorder/sequent.mli @@ -6,15 +6,16 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open EConstr open Formula open Globnames -module OrderedConstr: Set.OrderedType with type t=Constr.t +module OrderedConstr: Set.OrderedType with type t=Term.constr -module CM: CSig.MapS with type key=Constr.t +module CM: CSig.MapS with type key=Term.constr -type h_item = global_reference * (int*Constr.t) option +type h_item = global_reference * (int*Term.constr) option module History: Set.S with type elt = h_item diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index 49bf07155..e1adebe8d 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Term open EConstr @@ -54,12 +55,12 @@ let unif evd t1 t2= | Meta i,_ -> let t=subst_meta !sigma nt2 in if Int.Set.is_empty (free_rels evd t) && - not (occur_term evd (EConstr.mkMeta i) t) then + not (dependent evd (EConstr.mkMeta i) t) then bind i t else raise (UFAIL(nt1,nt2)) | _,Meta i -> let t=subst_meta !sigma nt1 in if Int.Set.is_empty (free_rels evd t) && - not (occur_term evd (EConstr.mkMeta i) t) then + not (dependent evd (EConstr.mkMeta i) t) then bind i t else raise (UFAIL(nt1,nt2)) | Cast(_,_,_),_->Queue.add (strip_outer_cast evd nt1,nt2) bige | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast evd nt2) bige diff --git a/plugins/firstorder/unify.mli b/plugins/firstorder/unify.mli index c9cca9bd8..7f1fb9bd0 100644 --- a/plugins/firstorder/unify.mli +++ b/plugins/firstorder/unify.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Term open EConstr diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index 317444cf1..b44307590 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -12,6 +12,7 @@ des inéquations et équations sont entiers. En attendant la tactique Field. *) +open API open Term open Tactics open Names @@ -76,8 +77,8 @@ let flin_emult a f = type ineq = Rlt | Rle | Rgt | Rge let string_of_R_constant kn = - match Names.repr_con kn with - | MPfile dir, sec_dir, id when + match Constant.repr3 kn with + | ModPath.MPfile dir, sec_dir, id when sec_dir = DirPath.empty && DirPath.to_string dir = "Coq.Reals.Rdefinitions" -> Label.to_string id diff --git a/plugins/fourier/vo.itarget b/plugins/fourier/vo.itarget deleted file mode 100644 index 87d82dacc..000000000 --- a/plugins/fourier/vo.itarget +++ /dev/null @@ -1,2 +0,0 @@ -Fourier_util.vo -Fourier.vo diff --git a/plugins/funind/FunInd.v b/plugins/funind/FunInd.v new file mode 100644 index 000000000..e40aea776 --- /dev/null +++ b/plugins/funind/FunInd.v @@ -0,0 +1,10 @@ +(************************************************************************) +(* 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 Coq.extraction.Extraction. +Declare ML Module "recdef_plugin". diff --git a/plugins/funind/Recdef.v b/plugins/funind/Recdef.v index e4433247b..64f43b833 100644 --- a/plugins/funind/Recdef.v +++ b/plugins/funind/Recdef.v @@ -6,8 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +Require Export Coq.funind.FunInd. Require Import PeanoNat. - Require Compare_dec. Require Wf_nat. diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index fd4962398..ef894b239 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -1,3 +1,4 @@ +open API open Printer open CErrors open Util @@ -105,7 +106,7 @@ let make_refl_eq constructor type_of_t t = type pte_info = { - proving_tac : (Id.t list -> Tacmach.tactic); + proving_tac : (Id.t list -> Proof_type.tactic); is_valid : constr -> bool } @@ -687,7 +688,7 @@ let instanciate_hyps_with_args (do_prove:Id.t list -> tactic) hyps args_id = let build_proof (interactive_proof:bool) - (fnames:constant list) + (fnames:Constant.t list) ptes_infos dyn_infos : tactic = @@ -707,13 +708,13 @@ let build_proof let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t in - tclTHENSEQ + tclTHENLIST [ Proofview.V82.of_tactic (generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps))); thin dyn_infos.rec_hyps; Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],t] None); (fun g -> observe_tac "toto" ( - tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t); + tclTHENLIST [Proofview.V82.of_tactic (Simple.case t); (fun g' -> let g'_nb_prod = nb_prod (project g') (pf_concl g') in let nb_instanciate_partial = g'_nb_prod - g_nb_prod in @@ -981,14 +982,14 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let eqn = mkApp(Lazy.force eq,[|type_of_f;eq_lhs;eq_rhs|]) in let lemma_type = it_mkProd_or_LetIn eqn type_ctxt in (* Pp.msgnl (str "lemma type " ++ Printer.pr_lconstr lemma_type ++ fnl () ++ str "f_body " ++ Printer.pr_lconstr f_body); *) - let f_id = Label.to_id (con_label (fst (destConst evd f))) in + let f_id = Label.to_id (Constant.label (fst (destConst evd f))) in let prove_replacement = - tclTHENSEQ + tclTHENLIST [ tclDO (nb_params + rec_args_num + 1) (Proofview.V82.of_tactic intro); observe_tac "" (fun g -> let rec_id = pf_nth_hyp_id g 1 in - tclTHENSEQ + tclTHENLIST [observe_tac "generalize_non_dep in generate_equation_lemma" (generalize_non_dep rec_id); observe_tac "h_case" (Proofview.V82.of_tactic (simplest_case (mkVar rec_id))); (Proofview.V82.of_tactic intros_reflexivity)] g @@ -1018,7 +1019,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let finfos = find_Function_infos (fst (destConst !evd f)) (*FIXME*) in mkConst (Option.get finfos.equation_lemma) with (Not_found | Option.IsNone as e) -> - let f_id = Label.to_id (con_label (fst (destConst !evd f))) in + let f_id = Label.to_id (Constant.label (fst (destConst !evd f))) in (*i The next call to mk_equation_id is valid since we will construct the lemma Ensures by: obvious i*) @@ -1241,7 +1242,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam other_fix_infos 0) in let first_tac : tactic = (* every operations until fix creations *) - tclTHENSEQ + tclTHENLIST [ observe_tac "introducing params" (Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.params))); observe_tac "introducing predictes" (Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.predicates))); observe_tac "introducing branches" (Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.branches))); @@ -1259,7 +1260,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam in let fix_info = Id.Map.find pte ptes_to_fix in let nb_args = fix_info.nb_realargs in - tclTHENSEQ + tclTHENLIST [ (* observe_tac ("introducing args") *) (tclDO nb_args (Proofview.V82.of_tactic intro)); (fun g -> (* replacement of the function by its body *) @@ -1278,7 +1279,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam eq_hyps = [] } in - tclTHENSEQ + tclTHENLIST [ observe_tac "do_replace" (do_replace evd @@ -1321,7 +1322,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam ] gl with Not_found -> let nb_args = min (princ_info.nargs) (List.length ctxt) in - tclTHENSEQ + tclTHENLIST [ tclDO nb_args (Proofview.V82.of_tactic intro); (fun g -> (* replacement of the function by its body *) @@ -1342,7 +1343,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam } in let fname = destConst (project g) (fst (decompose_app (project g) (List.hd (List.rev pte_args)))) in - tclTHENSEQ + tclTHENLIST [Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))]); let do_prove = build_proof @@ -1401,7 +1402,7 @@ let prove_with_tcc tcc_lemma_constr eqs : tactic = fun gls -> (* let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in *) (* let ids = hid::pf_ids_of_hyps gls in *) - tclTHENSEQ + tclTHENLIST [ (* generalize [lemma]; *) (* h_intro hid; *) @@ -1456,13 +1457,13 @@ let rec rewrite_eqs_in_eqs eqs = let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = fun gls -> - (tclTHENSEQ + (tclTHENLIST [ backtrack_eqs_until_hrec hrec eqs; (* observe_tac ("new_prove_with_tcc ( applying "^(Id.to_string hrec)^" )" ) *) (tclTHENS (* We must have exactly ONE subgoal !*) (Proofview.V82.of_tactic (apply (mkVar hrec))) - [ tclTHENSEQ + [ tclTHENLIST [ (Proofview.V82.of_tactic (keep (tcc_hyps@eqs))); (Proofview.V82.of_tactic (apply (Lazy.force acc_inv))); @@ -1616,7 +1617,7 @@ let prove_principle_for_gen (Id.of_string "prov") hyps in - tclTHENSEQ + tclTHENLIST [ Proofview.V82.of_tactic (generalize [lemma]); Proofview.V82.of_tactic (Simple.intro hid); @@ -1635,7 +1636,7 @@ let prove_principle_for_gen ] gls in - tclTHENSEQ + tclTHENLIST [ observe_tac "start_tac" start_tac; h_intros diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli index 61752aa33..5bb288678 100644 --- a/plugins/funind/functional_principles_proofs.mli +++ b/plugins/funind/functional_principles_proofs.mli @@ -1,19 +1,20 @@ +open API open Names val prove_princ_for_struct : Evd.evar_map ref -> bool -> - int -> constant array -> EConstr.constr array -> int -> Tacmach.tactic + int -> Constant.t array -> EConstr.constr array -> int -> Proof_type.tactic val prove_principle_for_gen : - constant*constant*constant -> (* name of the function, the functional and the fixpoint equation *) + Constant.t * Constant.t * Constant.t -> (* name of the function, the functional and the fixpoint equation *) Indfun_common.tcc_lemma_value ref -> (* a pointer to the obligation proofs lemma *) bool -> (* is that function uses measure *) int -> (* the number of recursive argument *) EConstr.types -> (* the type of the recursive argument *) EConstr.constr -> (* the wf relation used to prove the function *) - Tacmach.tactic + Proof_type.tactic (* val is_pte : rel_declaration -> bool *) diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index b8070ff88..8ffd15f9f 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -1,3 +1,4 @@ +open API open Printer open CErrors open Util @@ -149,7 +150,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = ([],[]) in let new_f,binders_to_remove_from_f = compute_new_princ_type remove env f in - applist(new_f, new_args), + applistc new_f new_args, list_union_eq eq_constr binders_to_remove_from_f binders_to_remove | LetIn(x,v,t,b) -> compute_new_princ_type_for_letin remove env x v t b @@ -329,7 +330,7 @@ let generate_functional_principle (evd: Evd.evar_map ref) match new_princ_name with | Some (id) -> id,id | None -> - let id_of_f = Label.to_id (con_label (fst f)) in + let id_of_f = Label.to_id (Constant.label (fst f)) in id_of_f,Indrec.make_elimination_ident id_of_f (family_of_sort type_sort) in let names = ref [new_princ_name] in @@ -370,12 +371,12 @@ let generate_functional_principle (evd: Evd.evar_map ref) begin begin try - let id = Pfedit.get_current_proof_name () in + let id = Proof_global.get_current_proof_name () in let s = Id.to_string id in let n = String.length "___________princ_________" in if String.length s >= n then if String.equal (String.sub s 0 n) "___________princ_________" - then Pfedit.delete_current_proof () + then Proof_global.discard_current () else () else () with e when CErrors.noncritical e -> () @@ -388,14 +389,14 @@ let generate_functional_principle (evd: Evd.evar_map ref) exception Not_Rec let get_funs_constant mp dp = - let get_funs_constant const e : (Names.constant*int) array = + let get_funs_constant const e : (Names.Constant.t*int) array = match kind_of_term ((strip_lam e)) with | Fix((_,(na,_,_))) -> Array.mapi (fun i na -> match na with | Name id -> - let const = make_con mp dp (Label.of_id id) in + let const = Constant.make3 mp dp (Label.of_id id) in const,i | Anonymous -> anomaly (Pp.str "Anonymous fix.") @@ -523,12 +524,12 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Safe_typing.private_con begin begin try - let id = Pfedit.get_current_proof_name () in + let id = Proof_global.get_current_proof_name () in let s = Id.to_string id in let n = String.length "___________princ_________" in if String.length s >= n then if String.equal (String.sub s 0 n) "___________princ_________" - then Pfedit.delete_current_proof () + then Proof_global.discard_current () else () else () with e when CErrors.noncritical e -> () @@ -655,7 +656,7 @@ let build_case_scheme fa = user_err ~hdr:"FunInd.build_case_scheme" (str "Cannot find " ++ Libnames.pr_reference f) in let first_fun,u = destConst funs in - let funs_mp,funs_dp,_ = Names.repr_con first_fun in + let funs_mp,funs_dp,_ = Constant.repr3 first_fun in let first_fun_kn = try fst (find_Function_infos first_fun).graph_ind with Not_found -> raise No_graph_found in let this_block_funs_indexes = get_funs_constant funs_mp funs_dp first_fun in let this_block_funs = Array.map (fun (c,_) -> (c,u)) this_block_funs_indexes in diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli index 45ad332fc..bb2b2d918 100644 --- a/plugins/funind/functional_principles_types.mli +++ b/plugins/funind/functional_principles_types.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Term open Misctypes @@ -17,7 +18,7 @@ val generate_functional_principle : (* induction principle on rel *) types -> (* *) - sorts array option -> + Sorts.t array option -> (* Name of the new principle *) (Id.t) option -> (* the compute functions to use *) @@ -27,10 +28,10 @@ val generate_functional_principle : (* The tactic to use to make the proof w.r the number of params *) - (EConstr.constr array -> int -> Tacmach.tactic) -> + (EConstr.constr array -> int -> Proof_type.tactic) -> unit -val compute_new_princ_type_from_rel : constr array -> sorts array -> +val compute_new_princ_type_from_rel : constr array -> Sorts.t array -> types -> types diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index d28e0aba0..1258c9286 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API open Ltac_plugin open Util open Pp @@ -164,7 +166,7 @@ VERNAC COMMAND EXTEND Function END let pr_fun_scheme_arg (princ_name,fun_name,s) = - Nameops.pr_id princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++ + Names.Id.print princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++ Libnames.pr_reference fun_name ++ spc() ++ str "Sort " ++ Ppconstr.pr_glob_sort s diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 785633e25..0e2ca4900 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -1,3 +1,4 @@ +open API open Printer open Pp open Names diff --git a/plugins/funind/glob_term_to_relation.mli b/plugins/funind/glob_term_to_relation.mli index 0cab5a6d3..7ad7de079 100644 --- a/plugins/funind/glob_term_to_relation.mli +++ b/plugins/funind/glob_term_to_relation.mli @@ -1,3 +1,4 @@ +open API open Names (* diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 6fd496f50..726a8203d 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -1,3 +1,4 @@ +open API open Pp open Glob_term open CErrors @@ -578,8 +579,8 @@ let ids_of_pat = ids_of_pat Id.Set.empty let id_of_name = function - | Names.Anonymous -> Id.of_string "x" - | Names.Name x -> x + | Anonymous -> Id.of_string "x" + | Name x -> x (* TODO: finish Rec caes *) let ids_of_glob_constr c = @@ -721,7 +722,7 @@ let resolve_and_replace_implicits ?(flags=Pretyping.all_and_fail_flags) ?(expect (* we first (pseudo) understand [rt] and get back the computed evar_map *) (* FIXME : JF (30/03/2017) I'm not completely sure to have split understand as needed. If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *) - let ctx,_ = Pretyping.ise_pretype_gen flags env sigma Pretyping.empty_lvar expected_type rt in + let ctx,_ = Pretyping.ise_pretype_gen flags env sigma Glob_ops.empty_lvar expected_type rt in let ctx, f = Evarutil.nf_evars_and_universes ctx in (* then we map [rt] to replace the implicit holes by their values *) diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli index 99a258de9..b6d2c4543 100644 --- a/plugins/funind/glob_termops.mli +++ b/plugins/funind/glob_termops.mli @@ -1,3 +1,4 @@ +open API open Names open Glob_term open Misctypes diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index f1a9758e8..d12aa7f42 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -1,3 +1,4 @@ +open API open CErrors open Util open Names @@ -64,7 +65,7 @@ let functional_induction with_clean c princl pat = (or f_rec, f_rect) i*) let princ_name = Indrec.make_elimination_ident - (Label.to_id (con_label c')) + (Label.to_id (Constant.label c')) (Tacticals.elimination_sort_of_goal g) in try @@ -341,8 +342,8 @@ let error_error names e = let generate_principle (evd:Evd.evar_map ref) pconstants on_error is_general do_built (fix_rec_l:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) recdefs interactive_proof - (continue_proof : int -> Names.constant array -> EConstr.constr array -> int -> - Tacmach.tactic) : unit = + (continue_proof : int -> Names.Constant.t array -> EConstr.constr array -> int -> + Proof_type.tactic) : unit = let names = List.map (function (((_, name),_),_,_,_,_),_ -> name) fix_rec_l in let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in let funs_args = List.map fst fun_bodies in @@ -445,7 +446,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let generate_correction_proof_wf f_ref tcc_lemma_ref is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation - (_: int) (_:Names.constant array) (_:EConstr.constr array) (_:int) : Tacmach.tactic = + (_: int) (_:Names.Constant.t array) (_:EConstr.constr array) (_:int) : Proof_type.tactic = Functional_principles_proofs.prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation @@ -898,14 +899,14 @@ let make_graph (f_ref:global_reference) = in l | _ -> - let id = Label.to_id (con_label c) in + let id = Label.to_id (Constant.label c) in [(((Loc.tag id),None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]] in - let mp,dp,_ = repr_con c in + let mp,dp,_ = Constant.repr3 c in do_generate_principle [c,Univ.Instance.empty] error_error false false expr_list; (* We register the infos *) List.iter - (fun ((((_,id),_),_,_,_,_),_) -> add_Function false (make_con mp dp (Label.of_id id))) + (fun ((((_,id),_),_,_,_,_),_) -> add_Function false (Constant.make3 mp dp (Label.of_id id))) expr_list) let do_generate_principle = do_generate_principle [] warning_error true diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli index ba89fe4a7..33420d813 100644 --- a/plugins/funind/indfun.mli +++ b/plugins/funind/indfun.mli @@ -1,3 +1,4 @@ +open API open Misctypes val warn_cannot_define_graph : ?loc:Loc.t -> Pp.std_ppcmds * Pp.std_ppcmds -> unit @@ -15,7 +16,7 @@ val functional_induction : EConstr.constr -> (EConstr.constr * EConstr.constr bindings) option -> Tacexpr.or_and_intro_pattern option -> - Proof_type.goal Tacmach.sigma -> Proof_type.goal list Evd.sigma + Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma val make_graph : Globnames.global_reference -> unit diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index a73425543..6fe6888f3 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -1,3 +1,4 @@ +open API open Names open Pp open Libnames @@ -108,7 +109,7 @@ let const_of_id id = try Constrintern.locate_reference princ_ref with Not_found -> CErrors.user_err ~hdr:"IndFun.const_of_id" - (str "cannot find " ++ Nameops.pr_id id) + (str "cannot find " ++ Id.print id) let def_of_const t = match (Term.kind_of_term t) with @@ -160,7 +161,7 @@ let save with_clean id const (locality,_,kind) hook = let kn = declare_constant id ~local (DefinitionEntry const, k) in (locality, ConstRef kn) in - if with_clean then Pfedit.delete_current_proof (); + if with_clean then Proof_global.discard_current (); CEphemeron.iter_opt hook (fun f -> Lemmas.call_hook fix_exn f l r); definition_message id @@ -172,7 +173,7 @@ let cook_proof _ = let get_proof_clean do_reduce = let result = cook_proof do_reduce in - Pfedit.delete_current_proof (); + Proof_global.discard_current (); result let with_full_print f a = @@ -216,14 +217,14 @@ let with_full_print f a = type function_info = { - function_constant : constant; + function_constant : Constant.t; graph_ind : inductive; - equation_lemma : constant option; - correctness_lemma : constant option; - completeness_lemma : constant option; - rect_lemma : constant option; - rec_lemma : constant option; - prop_lemma : constant option; + equation_lemma : Constant.t option; + correctness_lemma : Constant.t option; + completeness_lemma : Constant.t option; + rect_lemma : Constant.t option; + rec_lemma : Constant.t option; + prop_lemma : Constant.t option; is_general : bool; (* Has this function been defined using general recursive definition *) } @@ -388,7 +389,7 @@ let update_Function finfo = let add_Function is_general f = - let f_id = Label.to_id (con_label f) in + let f_id = Label.to_id (Constant.label f) in let equation_lemma = find_or_none (mk_equation_id f_id) and correctness_lemma = find_or_none (mk_correct_id f_id) and completeness_lemma = find_or_none (mk_complete_id f_id) @@ -547,5 +548,5 @@ let compose_prod l b = prodn (List.length l) l b type tcc_lemma_value = | Undefined - | Value of Constr.constr + | Value of Term.constr | Not_needed diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli index 5ef8f05bb..6b40c9171 100644 --- a/plugins/funind/indfun_common.mli +++ b/plugins/funind/indfun_common.mli @@ -1,3 +1,4 @@ +open API open Names open Pp @@ -22,7 +23,7 @@ val array_get_start : 'a array -> 'a array val id_of_name : Name.t -> Id.t val locate_ind : Libnames.reference -> inductive -val locate_constant : Libnames.reference -> constant +val locate_constant : Libnames.reference -> Constant.t val locate_with_msg : Pp.std_ppcmds -> (Libnames.reference -> 'a) -> Libnames.reference -> 'a @@ -69,21 +70,21 @@ val with_full_print : ('a -> 'b) -> 'a -> 'b type function_info = { - function_constant : constant; + function_constant : Constant.t; graph_ind : inductive; - equation_lemma : constant option; - correctness_lemma : constant option; - completeness_lemma : constant option; - rect_lemma : constant option; - rec_lemma : constant option; - prop_lemma : constant option; + equation_lemma : Constant.t option; + correctness_lemma : Constant.t option; + completeness_lemma : Constant.t option; + rect_lemma : Constant.t option; + rec_lemma : Constant.t option; + prop_lemma : Constant.t option; is_general : bool; } -val find_Function_infos : constant -> function_info +val find_Function_infos : Constant.t -> function_info val find_Function_of_graph : inductive -> function_info (* WARNING: To be used just after the graph definition !!! *) -val add_Function : bool -> constant -> unit +val add_Function : bool -> Constant.t -> unit val update_Function : function_info -> unit @@ -122,5 +123,5 @@ val compose_prod : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t type tcc_lemma_value = | Undefined - | Value of Constr.constr + | Value of Term.constr | Not_needed diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index bcfa6b931..ebdb490e3 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Ltac_plugin open Declarations open CErrors @@ -217,7 +218,7 @@ let rec generate_fresh_id x avoid i = \end{enumerate} *) -let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes lemmas_types_infos i : tactic = +let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes lemmas_types_infos i : Proof_type.tactic = fun g -> (* first of all we recreate the lemmas types to be used as predicates of the induction principle that is~: @@ -341,7 +342,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes in (* observe (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor); *) ( - tclTHENSEQ + tclTHENLIST [ observe_tac("h_intro_patterns ") (let l = (List.nth intro_pats (pred i)) in match l with @@ -414,7 +415,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes in (params_bindings@lemmas_bindings) in - tclTHENSEQ + tclTHENLIST [ observe_tac "principle" (Proofview.V82.of_tactic (assert_by (Name principle_id) @@ -467,7 +468,7 @@ let tauto = let rec intros_with_rewrite g = observe_tac "intros_with_rewrite" intros_with_rewrite_aux g -and intros_with_rewrite_aux : tactic = +and intros_with_rewrite_aux : Proof_type.tactic = fun g -> let eq_ind = make_eq () in let sigma = project g in @@ -479,16 +480,16 @@ and intros_with_rewrite_aux : tactic = if Reductionops.is_conv (pf_env g) (project g) args.(1) args.(2) then let id = pf_get_new_id (Id.of_string "y") g in - tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g + tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g else if isVar sigma args.(1) && (Environ.evaluable_named (destVar sigma args.(1)) (pf_env g)) - then tclTHENSEQ[ + then tclTHENLIST[ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))]); tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))] ((destVar sigma args.(1)),Locus.InHyp) ))) (pf_ids_of_hyps g); intros_with_rewrite ] g else if isVar sigma args.(2) && (Environ.evaluable_named (destVar sigma args.(2)) (pf_env g)) - then tclTHENSEQ[ + then tclTHENLIST[ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))]); tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))] ((destVar sigma args.(2)),Locus.InHyp) ))) (pf_ids_of_hyps g); @@ -497,7 +498,7 @@ and intros_with_rewrite_aux : tactic = else if isVar sigma args.(1) then let id = pf_get_new_id (Id.of_string "y") g in - tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); + tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); generalize_dependent_of (destVar sigma args.(1)) id; tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); intros_with_rewrite @@ -506,7 +507,7 @@ and intros_with_rewrite_aux : tactic = else if isVar sigma args.(2) then let id = pf_get_new_id (Id.of_string "y") g in - tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); + tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); generalize_dependent_of (destVar sigma args.(2)) id; tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id))); intros_with_rewrite @@ -515,7 +516,7 @@ and intros_with_rewrite_aux : tactic = else begin let id = pf_get_new_id (Id.of_string "y") g in - tclTHENSEQ[ + tclTHENLIST[ Proofview.V82.of_tactic (Simple.intro id); tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); intros_with_rewrite @@ -524,12 +525,12 @@ and intros_with_rewrite_aux : tactic = | Ind _ when EConstr.eq_constr sigma t (EConstr.of_constr (Universes.constr_of_global @@ Coqlib.build_coq_False ())) -> Proofview.V82.of_tactic tauto g | Case(_,_,v,_) -> - tclTHENSEQ[ + tclTHENLIST[ Proofview.V82.of_tactic (simplest_case v); intros_with_rewrite ] g | LetIn _ -> - tclTHENSEQ[ + tclTHENLIST[ Proofview.V82.of_tactic (reduce (Genredexpr.Cbv {Redops.all_flags @@ -541,10 +542,10 @@ and intros_with_rewrite_aux : tactic = ] g | _ -> let id = pf_get_new_id (Id.of_string "y") g in - tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);intros_with_rewrite] g + tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id);intros_with_rewrite] g end | LetIn _ -> - tclTHENSEQ[ + tclTHENLIST[ Proofview.V82.of_tactic (reduce (Genredexpr.Cbv {Redops.all_flags @@ -561,7 +562,7 @@ let rec reflexivity_with_destruct_cases g = try match EConstr.kind (project g) (snd (destApp (project g) (pf_concl g))).(2) with | Case(_,_,v,_) -> - tclTHENSEQ[ + tclTHENLIST[ Proofview.V82.of_tactic (simplest_case v); Proofview.V82.of_tactic intros; observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases @@ -581,7 +582,7 @@ let rec reflexivity_with_destruct_cases g = if Equality.discriminable (pf_env g) (project g) t1 t2 then Proofview.V82.of_tactic (Equality.discrHyp id) g else if Equality.injectable (pf_env g) (project g) t1 t2 - then tclTHENSEQ [Proofview.V82.of_tactic (Equality.injHyp None id);thin [id];intros_with_rewrite] g + then tclTHENLIST [Proofview.V82.of_tactic (Equality.injHyp None id);thin [id];intros_with_rewrite] g else tclIDTAC g | _ -> tclIDTAC g ) @@ -628,7 +629,7 @@ let rec reflexivity_with_destruct_cases g = *) -let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = +let prove_fun_complete funcs graphs schemes lemmas_types_infos i : Proof_type.tactic = fun g -> (* We compute the types of the different mutually recursive lemmas in $\zeta$ normal form @@ -672,7 +673,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = using [f_equation] if it is recursive (that is the graph is infinite or unfold if the graph is finite *) - let rewrite_tac j ids : tactic = + let rewrite_tac j ids : Proof_type.tactic = let graph_def = graphs.(j) in let infos = try find_Function_infos (fst (destConst (project g) funcs.(j))) @@ -685,7 +686,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = try Option.get (infos).equation_lemma with Option.IsNone -> anomaly (Pp.str "Cannot find equation lemma.") in - tclTHENSEQ[ + tclTHENLIST[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) ids; Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_lemma)); (* Don't forget to $\zeta$ normlize the term since the principles @@ -721,7 +722,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = end in let this_branche_ids = List.nth intro_pats (pred i) in - tclTHENSEQ[ + tclTHENLIST[ (* we expand the definition of the function *) observe_tac "rewrite_tac" (rewrite_tac this_ind_number this_branche_ids); (* introduce hypothesis with some rewrite *) @@ -734,7 +735,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = let params_names = fst (List.chop princ_infos.nparams args_names) in let open EConstr in let params = List.map mkVar params_names in - tclTHENSEQ + tclTHENLIST [ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]); observe_tac "h_generalize" (Proofview.V82.of_tactic (generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)])); @@ -806,7 +807,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( in Array.iteri (fun i f_as_constant -> - let f_id = Label.to_id (con_label (fst f_as_constant)) in + let f_id = Label.to_id (Constant.label (fst f_as_constant)) in (*i The next call to mk_correct_id is valid since we are constructing the lemma Ensures by: obvious i*) @@ -871,7 +872,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( in Array.iteri (fun i f_as_constant -> - let f_id = Label.to_id (con_label (fst f_as_constant)) in + let f_id = Label.to_id (Constant.label (fst f_as_constant)) in (*i The next call to mk_complete_id is valid since we are constructing the lemma Ensures by: obvious i*) @@ -922,7 +923,7 @@ let revert_graph kn post_tac hid g = | None -> tclIDTAC g | Some f_complete -> let f_args,res = Array.chop (Array.length args - 1) args in - tclTHENSEQ + tclTHENLIST [ Proofview.V82.of_tactic (generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])]); thin [hid]; @@ -952,7 +953,7 @@ let revert_graph kn post_tac hid g = \end{enumerate} *) -let functional_inversion kn hid fconst f_correct : tactic = +let functional_inversion kn hid fconst f_correct : Proof_type.tactic = fun g -> let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in let sigma = project g in @@ -967,7 +968,7 @@ let functional_inversion kn hid fconst f_correct : tactic = ((fun hid -> tclIDTAC),f_args,args.(1)) | _ -> (fun hid -> tclFAIL 1 (mt ())),[||],args.(2) in - tclTHENSEQ[ + tclTHENLIST [ pre_tac hid; Proofview.V82.of_tactic (generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])]); thin [hid]; diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 763443717..c75f7f868 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -8,6 +8,7 @@ (* Merging of induction principles. *) +open API open Globnames open Tactics open Indfun_common @@ -892,7 +893,7 @@ let find_Function_infos_safe (id:Id.t): Indfun_common.function_info = locate_constant f_ref in try find_Function_infos (kn_of_id id) with Not_found -> - user_err ~hdr:"indfun" (Nameops.pr_id id ++ str " has no functional scheme") + user_err ~hdr:"indfun" (Id.print id ++ str " has no functional scheme") (** [merge id1 id2 args1 args2 id] builds and declares a new inductive type called [id], representing the merged graphs of both graphs diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index ff397d2e9..3cd20a950 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + module CVars = Vars open Term @@ -75,7 +77,7 @@ let def_of_const t = | _ -> raise Not_found) with Not_found -> anomaly (str "Cannot find definition of constant " ++ - (Id.print (Label.to_id (con_label (fst sp)))) ++ str ".") + (Id.print (Label.to_id (Constant.label (fst sp)))) ++ str ".") ) |_ -> assert false @@ -170,7 +172,7 @@ let simpl_iter clause = clause (* Others ugly things ... *) -let (value_f:Constr.constr list -> global_reference -> Constr.constr) = +let (value_f:Term.constr list -> global_reference -> Term.constr) = let open Term in fun al fterm -> let rev_x_id_l = @@ -202,7 +204,7 @@ let (value_f:Constr.constr list -> global_reference -> Constr.constr) = let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in it_mkLambda_or_LetIn body context -let (declare_f : Id.t -> logical_kind -> Constr.constr list -> global_reference -> global_reference) = +let (declare_f : Id.t -> logical_kind -> Term.constr list -> global_reference -> global_reference) = fun f_id kind input_type fterm_ref -> declare_fun f_id kind (value_f input_type fterm_ref);; @@ -311,7 +313,7 @@ let check_not_nested sigma forbidden e = | Var x -> if Id.List.mem x forbidden then user_err ~hdr:"Recdef.check_not_nested" - (str "check_not_nested: failure " ++ pr_id x) + (str "check_not_nested: failure " ++ Id.print x) | Meta _ | Evar _ | Sort _ -> () | Cast(e,_,t) -> check_not_nested e;check_not_nested t | Prod(_,t,b) -> check_not_nested t;check_not_nested b @@ -448,7 +450,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g = check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; jinfo.otherS () expr_info continuation_tac expr_info g with e when CErrors.noncritical e -> - user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id) + user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id) end | Lambda(n,t,b) -> begin @@ -456,7 +458,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g = check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; jinfo.otherS () expr_info continuation_tac expr_info g with e when CErrors.noncritical e -> - user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id) + user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id) end | Case(ci,t,a,l) -> begin @@ -681,7 +683,7 @@ let pf_typel l tac = introduced back later; the result is the pair of the tactic and the list of hypotheses that have been generalized and cleared. *) let mkDestructEq : - Id.t list -> constr -> goal sigma -> tactic * Id.t list = + Id.t list -> constr -> goal Evd.sigma -> tactic * Id.t list = fun not_on_hyp expr g -> let hyps = pf_hyps g in let to_revert = @@ -689,7 +691,7 @@ let mkDestructEq : (fun decl -> let open Context.Named.Declaration in let id = get_id decl in - if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) expr (get_type decl)) + if Id.List.mem id not_on_hyp || not (Termops.dependent (project g) expr (get_type decl)) then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in let type_of_expr = pf_unsafe_type_of g expr in @@ -848,7 +850,7 @@ let rec prove_le g = try let matching_fun = pf_is_matching g - (Pattern.PApp(Pattern.PRef (reference_of_constr (EConstr.Unsafe.to_constr (le ()))),[|Pattern.PVar (destVar sigma x);Pattern.PMeta None|])) in + (Pattern.PApp(Pattern.PRef (Globnames.global_of_constr (EConstr.Unsafe.to_constr (le ()))),[|Pattern.PVar (destVar sigma x);Pattern.PMeta None|])) in let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g) in let y = @@ -868,7 +870,7 @@ let rec make_rewrite_list expr_info max = function | [] -> tclIDTAC | (_,p,hp)::l -> observe_tac (str "make_rewrite_list") (tclTHENS - (observe_tac (str "rewrite heq on " ++ pr_id p ) ( + (observe_tac (str "rewrite heq on " ++ Id.print p ) ( (fun g -> let sigma = project g in let t_eq = compute_renamed_type g (mkVar hp) in @@ -963,7 +965,7 @@ let rec destruct_hex expr_info acc l = onNthHypId 1 (fun hp -> onNthHypId 2 (fun p -> observe_tac - (str "destruct_hex after " ++ pr_id hp ++ spc () ++ pr_id p) + (str "destruct_hex after " ++ Id.print hp ++ spc () ++ Id.print p) (destruct_hex expr_info ((v,p,hp)::acc) l) ) ) @@ -1293,7 +1295,7 @@ let is_opaque_constant c = let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) = (* Pp.msgnl (str "gls_type := " ++ Printer.pr_lconstr gls_type); *) - let current_proof_name = get_current_proof_name () in + let current_proof_name = Proof_global.get_current_proof_name () in let name = match goal_name with | Some s -> s | None -> @@ -1455,7 +1457,7 @@ let start_equation (f:global_reference) (term_f:global_reference) let (com_eqn : int -> Id.t -> global_reference -> global_reference -> global_reference - -> Constr.constr -> unit) = + -> Term.constr -> unit) = fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type -> let open CVars in let opacity = diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli index 80f02e01c..e1a072799 100644 --- a/plugins/funind/recdef.mli +++ b/plugins/funind/recdef.mli @@ -1,4 +1,4 @@ - +open API (* val evaluable_of_global_reference : Libnames.global_reference -> Names.evaluable_global_reference *) val tclUSER_if_not_mes : diff --git a/plugins/funind/vo.itarget b/plugins/funind/vo.itarget deleted file mode 100644 index 33c968302..000000000 --- a/plugins/funind/vo.itarget +++ /dev/null @@ -1 +0,0 @@ -Recdef.vo diff --git a/plugins/ltac/coretactics.ml4 b/plugins/ltac/coretactics.ml4 index ea1660d90..07b8746fb 100644 --- a/plugins/ltac/coretactics.ml4 +++ b/plugins/ltac/coretactics.ml4 @@ -8,13 +8,14 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Util -open Names open Locus open Misctypes open Genredexpr open Stdarg open Extraargs +open Names DECLARE PLUGIN "coretactics" @@ -306,7 +307,7 @@ let initial_atomic () = let nocl = {onhyps=Some[];concl_occs=AllOccurrences} in let iter (s, t) = let body = TacAtom (Loc.tag t) in - Tacenv.register_ltac false false (Id.of_string s) body + Tacenv.register_ltac false false (Names.Id.of_string s) body in let () = List.iter iter [ "red", TacReduce(Red false,nocl); @@ -316,7 +317,7 @@ let initial_atomic () = "intros", TacIntroPattern (false,[]); ] in - let iter (s, t) = Tacenv.register_ltac false false (Id.of_string s) t in + let iter (s, t) = Tacenv.register_ltac false false (Names.Id.of_string s) t in List.iter iter [ "idtac",TacId []; "fail", TacFail(TacLocal,ArgArg 0,[]); diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 7db484d82..7ecfa57f6 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Names open Term @@ -27,7 +28,7 @@ let instantiate_evar evk (ist,rawc) sigma = let filtered = Evd.evar_filtered_env evi in let constrvars = Tacinterp.extract_ltac_constr_values ist filtered in let lvar = { - Pretyping.ltac_constrs = constrvars; + Glob_term.ltac_constrs = constrvars; ltac_uconstrs = Names.Id.Map.empty; ltac_idents = Names.Id.Map.empty; ltac_genargs = ist.Geninterp.lfun; @@ -86,16 +87,16 @@ let let_evar name typ = let _ = Typing.e_sort_of env sigma typ in let sigma = !sigma in let id = match name with - | Names.Anonymous -> + | Name.Anonymous -> let id = Namegen.id_of_name_using_hdchar env sigma typ name in Namegen.next_ident_away_in_goal id (Termops.ids_of_named_context (Environ.named_context env)) - | Names.Name id -> id + | Name.Name id -> id in let (sigma, evar) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) typ in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) - (Tactics.letin_tac None (Names.Name id) evar None Locusops.nowhere) + (Tactics.letin_tac None (Name.Name id) evar None Locusops.nowhere) end - + let hget_evar n = let open EConstr in Proofview.Goal.nf_enter begin fun gl -> @@ -107,6 +108,5 @@ let hget_evar n = if n <= 0 then user_err Pp.(str "Incorrect existential variable index."); let ev = List.nth evl (n-1) in let ev_type = EConstr.existential_type sigma ev in - Tactics.change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl)) + Tactics.change_concl (mkLetIn (Name.Anonymous,mkEvar ev,ev_type,concl)) end - diff --git a/plugins/ltac/evar_tactics.mli b/plugins/ltac/evar_tactics.mli index cfe747665..7c734cd9a 100644 --- a/plugins/ltac/evar_tactics.mli +++ b/plugins/ltac/evar_tactics.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Tacexpr open Locus diff --git a/plugins/ltac/extraargs.ml4 b/plugins/ltac/extraargs.ml4 index fdb8d3461..44f33ab80 100644 --- a/plugins/ltac/extraargs.ml4 +++ b/plugins/ltac/extraargs.ml4 @@ -8,6 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API open Pp open Genarg open Stdarg @@ -83,7 +85,7 @@ let pr_int_list_full _prc _prlc _prt l = pr_int_list l let pr_occurrences _prc _prlc _prt l = match l with | ArgArg x -> pr_int_list x - | ArgVar (loc, id) -> Nameops.pr_id id + | ArgVar (loc, id) -> Id.print id let occurrences_of = function | [] -> NoOccurrences @@ -199,8 +201,8 @@ let pr_gen_place pr_id = function | HypLocation (id,InHypValueOnly) -> str "in (Value of " ++ pr_id id ++ str ")" -let pr_loc_place _ _ _ = pr_gen_place (fun (_,id) -> Nameops.pr_id id) -let pr_place _ _ _ = pr_gen_place Nameops.pr_id +let pr_loc_place _ _ _ = pr_gen_place (fun (_,id) -> Id.print id) +let pr_place _ _ _ = pr_gen_place Id.print let pr_hloc = pr_loc_place () () () let intern_place ist = function @@ -236,7 +238,7 @@ ARGUMENT EXTEND hloc END -let pr_rename _ _ _ (n, m) = Nameops.pr_id n ++ str " into " ++ Nameops.pr_id m +let pr_rename _ _ _ (n, m) = Id.print n ++ str " into " ++ Id.print m ARGUMENT EXTEND rename TYPED AS ident * ident diff --git a/plugins/ltac/extraargs.mli b/plugins/ltac/extraargs.mli index 9b4167512..b2b3f8b6b 100644 --- a/plugins/ltac/extraargs.mli +++ b/plugins/ltac/extraargs.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Tacexpr open Names open Constrexpr diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 8afe3053d..7259faecd 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -8,6 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API open Pp open Genarg open Stdarg @@ -363,7 +365,7 @@ let refine_tac ist simple with_classes c = let update = begin fun sigma -> c env sigma end in - let refine = Refine.refine ~unsafe:true update in + let refine = Refine.refine ~typecheck:false update in if simple then refine else refine <*> Tactics.New.reduce_after_refine <*> @@ -462,8 +464,8 @@ open Evar_tactics (* TODO: add support for some test similar to g_constr.name_colon so that expressions like "evar (list A)" do not raise a syntax error *) TACTIC EXTEND evar - [ "evar" test_lpar_id_colon "(" ident(id) ":" lconstr(typ) ")" ] -> [ let_evar (Name id) typ ] -| [ "evar" constr(typ) ] -> [ let_evar Anonymous typ ] + [ "evar" test_lpar_id_colon "(" ident(id) ":" lconstr(typ) ")" ] -> [ let_evar (Name.Name id) typ ] +| [ "evar" constr(typ) ] -> [ let_evar Name.Anonymous typ ] END TACTIC EXTEND instantiate @@ -514,7 +516,7 @@ let cache_transitivity_lemma (_,(left,lem)) = let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref) -let inTransitivity : bool * Constr.constr -> obj = +let inTransitivity : bool * Term.constr -> obj = declare_object {(default_object "TRANSITIVITY-STEPS") with cache_function = cache_transitivity_lemma; open_function = (fun i o -> if Int.equal i 1 then cache_transitivity_lemma o); @@ -682,7 +684,7 @@ let hResolve id c occ t = let sigma = Evd.merge_universe_context sigma ctx in let t_constr_type = Retyping.get_type_of env sigma t_constr in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) - (change_concl (mkLetIn (Anonymous,t_constr,t_constr_type,concl))) + (change_concl (mkLetIn (Name.Anonymous,t_constr,t_constr_type,concl))) end let hResolve_auto id c t = diff --git a/plugins/ltac/extratactics.mli b/plugins/ltac/extratactics.mli index 18334dafe..c7ec26967 100644 --- a/plugins/ltac/extratactics.mli +++ b/plugins/ltac/extratactics.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + val discrHyp : Names.Id.t -> unit Proofview.tactic val injHyp : Names.Id.t -> unit Proofview.tactic diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4 index 2c2a4b850..dfd8e88a9 100644 --- a/plugins/ltac/g_auto.ml4 +++ b/plugins/ltac/g_auto.ml4 @@ -8,6 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API open Pp open Genarg open Stdarg @@ -15,7 +17,6 @@ open Pcoq.Prim open Pcoq.Constr open Pltac open Hints -open Names DECLARE PLUGIN "g_auto" diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4 index dd5307638..905cfd02a 100644 --- a/plugins/ltac/g_class.ml4 +++ b/plugins/ltac/g_class.ml4 @@ -8,10 +8,10 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Class_tactics open Stdarg open Tacarg -open Names DECLARE PLUGIN "g_class" diff --git a/plugins/ltac/g_eqdecide.ml4 b/plugins/ltac/g_eqdecide.ml4 index 679aa1127..570cd4e69 100644 --- a/plugins/ltac/g_eqdecide.ml4 +++ b/plugins/ltac/g_eqdecide.ml4 @@ -14,8 +14,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Eqdecide -open Names DECLARE PLUGIN "g_eqdecide" diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 index 36ac10bfe..4bab31b85 100644 --- a/plugins/ltac/g_ltac.ml4 +++ b/plugins/ltac/g_ltac.ml4 @@ -8,6 +8,9 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API + DECLARE PLUGIN "ltac_plugin" open Util @@ -228,8 +231,8 @@ GEXTEND Gram | "multimatch" -> General ] ] ; input_fun: - [ [ "_" -> Anonymous - | l = ident -> Name l ] ] + [ [ "_" -> Name.Anonymous + | l = ident -> Name.Name l ] ] ; let_clause: [ [ id = identref; ":="; te = tactic_expr -> @@ -396,7 +399,7 @@ let pr_ltac_selector = function | SelectNth i -> int i ++ str ":" | SelectList l -> str "[" ++ prlist_with_sep (fun () -> str ", ") pr_range_selector l ++ str "]" ++ str ":" -| SelectId id -> str "[" ++ Nameops.pr_id id ++ str "]" ++ str ":" +| SelectId id -> str "[" ++ Id.print id ++ str "]" ++ str ":" | SelectAll -> str "all" ++ str ":" VERNAC ARGUMENT EXTEND ltac_selector PRINTED BY pr_ltac_selector @@ -466,14 +469,14 @@ let pr_ltac_production_item = function | None -> mt () | Some sep -> str "," ++ spc () ++ quote (str sep) in - str arg ++ str "(" ++ Nameops.pr_id id ++ sep ++ str ")" + str arg ++ str "(" ++ Id.print id ++ sep ++ str ")" VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY pr_ltac_production_item | [ string(s) ] -> [ Tacentries.TacTerm s ] | [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> - [ Tacentries.TacNonTerm (Loc.tag ~loc ((Names.Id.to_string nt, sep), Some p)) ] + [ Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, sep), Some p)) ] | [ ident(nt) ] -> - [ Tacentries.TacNonTerm (Loc.tag ~loc ((Names.Id.to_string nt, None), None)) ] + [ Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, None), None)) ] END VERNAC COMMAND EXTEND VernacTacticNotation @@ -496,7 +499,7 @@ let pr_ltac_ref = Libnames.pr_reference let pr_tacdef_body tacdef_body = let id, redef, body = match tacdef_body with - | TacticDefinition ((_,id), body) -> Nameops.pr_id id, false, body + | TacticDefinition ((_,id), body) -> Id.print id, false, body | TacticRedefinition (id, body) -> pr_ltac_ref id, true, body in let idl, body = @@ -504,8 +507,8 @@ let pr_tacdef_body tacdef_body = | Tacexpr.TacFun (idl,b) -> idl,b | _ -> [], body in id ++ - prlist (function Anonymous -> str " _" - | Name id -> spc () ++ Nameops.pr_id id) idl + prlist (function Name.Anonymous -> str " _" + | Name.Name id -> spc () ++ Id.print id) idl ++ (if redef then str" ::=" else str" :=") ++ brk(1,1) ++ Pptactic.pr_raw_tactic body diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4 index 4dceb0331..18e62a211 100644 --- a/plugins/ltac/g_obligations.ml4 +++ b/plugins/ltac/g_obligations.ml4 @@ -12,7 +12,8 @@ Syntax for the subtac terms and types. Elaborated from correctness/psyntax.ml4 by Jean-Christophe Filliâtre *) - +open API +open Grammar_API open Libnames open Constrexpr open Constrexpr_ops diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4 index 25258ffa9..e6ddc5cc1 100644 --- a/plugins/ltac/g_rewrite.ml4 +++ b/plugins/ltac/g_rewrite.ml4 @@ -10,6 +10,8 @@ (* Syntax for rewriting with strategies *) +open API +open Grammar_API open Names open Misctypes open Locus diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4 index 83bfd0233..804f73450 100644 --- a/plugins/ltac/g_tactic.ml4 +++ b/plugins/ltac/g_tactic.ml4 @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Pp open CErrors open Util @@ -137,14 +139,16 @@ let destruction_arg_of_constr (c,lbind as clbind) = match lbind with end | _ -> ElimOnConstr clbind +let mkNumeral n = Numeral (string_of_int (abs n), 0<=n) + let mkTacCase with_evar = function | [(clear,ElimOnConstr cl),(None,None),None],None -> TacCase (with_evar,(clear,cl)) (* Reinterpret numbers as a notation for terms *) | [(clear,ElimOnAnonHyp n),(None,None),None],None -> TacCase (with_evar, - (clear,(CAst.make @@ CPrim (Numeral (Bigint.of_int n)), - NoBindings))) + (clear,(CAst.make @@ CPrim (mkNumeral n), + NoBindings))) (* Reinterpret ident as notations for variables in the context *) (* because we don't know if they are quantified or not *) | [(clear,ElimOnIdent id),(None,None),None],None -> @@ -475,7 +479,7 @@ GEXTEND Gram | -> None ] ] ; as_name: - [ [ "as"; id = ident -> Names.Name id | -> Names.Anonymous ] ] + [ [ "as"; id = ident ->Names.Name.Name id | -> Names.Name.Anonymous ] ] ; by_tactic: [ [ "by"; tac = tactic_expr LEVEL "3" -> Some tac @@ -538,7 +542,7 @@ GEXTEND Gram TacAtom (Loc.tag ~loc:!@loc @@ TacMutualCofix (id,List.map mk_cofix_tac fd)) | IDENT "pose"; (id,b) = bindings_with_parameters -> - TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (false,Names.Name id,b,Locusops.nowhere,true,None)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (false,Names.Name.Name id,b,Locusops.nowhere,true,None)) | IDENT "pose"; b = constr; na = as_name -> TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (false,na,b,Locusops.nowhere,true,None)) | IDENT "epose"; (id,b) = bindings_with_parameters -> @@ -546,7 +550,7 @@ GEXTEND Gram | IDENT "epose"; b = constr; na = as_name -> TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (true,na,b,Locusops.nowhere,true,None)) | IDENT "set"; (id,c) = bindings_with_parameters; p = clause_dft_concl -> - TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (false,Names.Name id,c,p,true,None)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (false,Names.Name.Name id,c,p,true,None)) | IDENT "set"; c = constr; na = as_name; p = clause_dft_concl -> TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (false,na,c,p,true,None)) | IDENT "eset"; (id,c) = bindings_with_parameters; p = clause_dft_concl -> @@ -598,9 +602,9 @@ GEXTEND Gram TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (true,false,Some tac,ipat,c)) | IDENT "generalize"; c = constr -> - TacAtom (Loc.tag ~loc:!@loc @@ TacGeneralize [((AllOccurrences,c),Names.Anonymous)]) + TacAtom (Loc.tag ~loc:!@loc @@ TacGeneralize [((AllOccurrences,c),Names.Name.Anonymous)]) | IDENT "generalize"; c = constr; l = LIST1 constr -> - let gen_everywhere c = ((AllOccurrences,c),Names.Anonymous) in + let gen_everywhere c = ((AllOccurrences,c),Names.Name.Anonymous) in TacAtom (Loc.tag ~loc:!@loc @@ TacGeneralize (List.map gen_everywhere (c::l))) | IDENT "generalize"; c = constr; lookup_at_as_comma; nl = occs; na = as_name; diff --git a/plugins/ltac/pltac.ml b/plugins/ltac/pltac.ml index 7e979d269..84c5d3a44 100644 --- a/plugins/ltac/pltac.ml +++ b/plugins/ltac/pltac.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Pcoq (* Main entry for extensions *) diff --git a/plugins/ltac/pltac.mli b/plugins/ltac/pltac.mli index 810e1ec39..9261a11c7 100644 --- a/plugins/ltac/pltac.mli +++ b/plugins/ltac/pltac.mli @@ -8,6 +8,8 @@ (** Ltac parsing entries *) +open API +open Grammar_API open Loc open Names open Pcoq diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index 9446f9df4..8300a55e3 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Pp open Names open Namegen @@ -334,11 +335,11 @@ type 'a extra_genarg_printer = | ArgVar (loc,id) -> pr_with_comments ?loc (pr_id id) let pr_ltac_constant kn = - if !Flags.in_debugger then pr_kn kn + if !Flags.in_debugger then KerName.print kn else try pr_qualid (Nametab.shortest_qualid_of_tactic kn) with Not_found -> (* local tactic not accessible anymore *) - str "<" ++ pr_kn kn ++ str ">" + str "<" ++ KerName.print kn ++ str ">" let pr_evaluable_reference_env env = function | EvalVarRef id -> pr_id id @@ -481,7 +482,7 @@ type 'a extra_genarg_printer = | SelectNth i -> int i ++ str ":" | SelectList l -> str "[" ++ prlist_with_sep (fun () -> str ", ") pr_range_selector l ++ str "]" ++ str ":" - | SelectId id -> str "[" ++ Nameops.pr_id id ++ str "]" ++ str ":" + | SelectId id -> str "[" ++ Id.print id ++ str "]" ++ str ":" | SelectAll -> str "all" ++ str ":" let pr_lazy = function diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli index 4265c416b..519283759 100644 --- a/plugins/ltac/pptactic.mli +++ b/plugins/ltac/pptactic.mli @@ -9,6 +9,7 @@ (** This module implements pretty-printers for tactic_expr syntactic objects and their subcomponents. *) +open API open Pp open Genarg open Geninterp diff --git a/plugins/ltac/profile_ltac.ml b/plugins/ltac/profile_ltac.ml index b237e917d..020b3048f 100644 --- a/plugins/ltac/profile_ltac.ml +++ b/plugins/ltac/profile_ltac.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Unicode open Pp open Printer @@ -246,7 +247,7 @@ let string_of_call ck = (match ck with | Tacexpr.LtacNotationCall s -> Pptactic.pr_alias_key s | Tacexpr.LtacNameCall cst -> Pptactic.pr_ltac_constant cst - | Tacexpr.LtacVarCall (id, t) -> Nameops.pr_id id + | Tacexpr.LtacVarCall (id, t) -> Names.Id.print id | Tacexpr.LtacAtomCall te -> (Pptactic.pr_glob_tactic (Global.env ()) (Tacexpr.TacAtom (Loc.tag te))) diff --git a/plugins/ltac/profile_ltac.mli b/plugins/ltac/profile_ltac.mli index e5e2e4197..09fc549c6 100644 --- a/plugins/ltac/profile_ltac.mli +++ b/plugins/ltac/profile_ltac.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + (** Ltac profiling primitives *) val do_profile : diff --git a/plugins/ltac/profile_ltac_tactics.ml4 b/plugins/ltac/profile_ltac_tactics.ml4 index 8cb76d81c..83fb6963b 100644 --- a/plugins/ltac/profile_ltac_tactics.ml4 +++ b/plugins/ltac/profile_ltac_tactics.ml4 @@ -10,6 +10,7 @@ (** Ltac profiling entrypoints *) +open API open Profile_ltac open Stdarg diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 68dc1fd37..fad181c89 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Pp open CErrors @@ -426,7 +427,7 @@ let split_head = function | [] -> assert(false) let eq_pb (ty, env, x, y as pb) (ty', env', x', y' as pb') = - pb == pb' || (ty == ty' && Constr.equal x x' && Constr.equal y y') + pb == pb' || (ty == ty' && Term.eq_constr x x' && Term.eq_constr y y') let problem_inclusion x y = List.for_all (fun pb -> List.exists (fun pb' -> eq_pb pb pb') y) x @@ -956,7 +957,7 @@ let fold_match ?(force=false) env sigma c = let unfold_match env sigma sk app = match EConstr.kind sigma app with - | App (f', args) when eq_constant (fst (destConst sigma f')) sk -> + | App (f', args) when Constant.equal (fst (destConst sigma f')) sk -> let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in let v = EConstr.of_constr v in Reductionops.whd_beta sigma (mkApp (v, args)) @@ -1370,7 +1371,7 @@ module Strategies = fail cs let inj_open hint = (); fun sigma -> - let ctx = Evd.evar_universe_context_of hint.Autorewrite.rew_ctx in + let ctx = UState.of_context_set hint.Autorewrite.rew_ctx in let sigma = Evd.merge_universe_context sigma ctx in (sigma, (EConstr.of_constr hint.Autorewrite.rew_lemma, NoBindings)) @@ -1471,7 +1472,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul let evars = (!evdref, Evar.Set.empty) in let evars, cstr = let prop, (evars, arrow) = - if is_prop_sort sort then true, app_poly_sort true env evars impl [||] + if Sorts.is_prop sort then true, app_poly_sort true env evars impl [||] else false, app_poly_sort false env evars TypeGlobal.arrow [||] in match is_hyp with @@ -1538,7 +1539,7 @@ let assert_replacing id newt tac = | d :: rem -> insert_dependent env sigma (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in - Refine.refine ~unsafe:false begin fun sigma -> + Refine.refine ~typecheck:true begin fun sigma -> let (sigma, ev) = Evarutil.new_evar env' sigma concl in let (sigma, ev') = Evarutil.new_evar env sigma newt in let map d = @@ -1572,7 +1573,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = match clause, prf with | Some id, Some p -> let tac = tclTHENLIST [ - Refine.refine ~unsafe:false (fun h -> (h,p)); + Refine.refine ~typecheck:true (fun h -> (h,p)); Proofview.Unsafe.tclNEWGOALS gls; ] in Proofview.Unsafe.tclEVARS undef <*> @@ -1589,7 +1590,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = let (sigma, ev) = Evarutil.new_evar env sigma newt in (sigma, mkApp (p, [| ev |])) end in - Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls + Refine.refine ~typecheck:true make <*> Proofview.Unsafe.tclNEWGOALS gls end | None, None -> Proofview.Unsafe.tclEVARS undef <*> @@ -1964,7 +1965,7 @@ let add_morphism_infer glob m n = if Lib.is_modtype () then let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest instance_id (Entries.ParameterEntry - (None,poly,(instance,Evd.evar_context_universe_context uctx),None), + (None,poly,(instance,UState.context uctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical) in add_instance (Typeclasses.new_instance diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli index 6683d753b..d7f92fd6e 100644 --- a/plugins/ltac/rewrite.mli +++ b/plugins/ltac/rewrite.mli @@ -6,8 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names -open Constr open Environ open EConstr open Constrexpr @@ -38,7 +38,7 @@ type ('constr,'redexpr) strategy_ast = type rewrite_proof = | RewPrf of constr * constr - | RewCast of cast_kind + | RewCast of Term.cast_kind type evars = evar_map * Evar.Set.t (* goal evars, constraint evars *) diff --git a/plugins/ltac/tacarg.ml b/plugins/ltac/tacarg.ml index 42552c484..2c9bf14be 100644 --- a/plugins/ltac/tacarg.ml +++ b/plugins/ltac/tacarg.ml @@ -8,6 +8,7 @@ (** Generic arguments based on Ltac. *) +open API open Genarg open Geninterp open Tacexpr diff --git a/plugins/ltac/tacarg.mli b/plugins/ltac/tacarg.mli index bfa423db2..e82cb516c 100644 --- a/plugins/ltac/tacarg.mli +++ b/plugins/ltac/tacarg.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Genarg open Tacexpr open Constrexpr diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml index e037bb4b2..117a16b0a 100644 --- a/plugins/ltac/taccoerce.ml +++ b/plugins/ltac/taccoerce.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Names open Term @@ -131,8 +132,8 @@ let coerce_var_to_ident fresh env sigma v = let coerce_to_ident_not_fresh env sigma v = let g = sigma in let id_of_name = function - | Names.Anonymous -> Id.of_string "x" - | Names.Name x -> x in + | Name.Anonymous -> Id.of_string "x" + | Name.Name x -> x in let v = Value.normalize v in let fail () = raise (CannotCoerceTo "an identifier") in if has_type v (topwit wit_intro_pattern) then diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli index 9883c03c4..2c02171d0 100644 --- a/plugins/ltac/taccoerce.mli +++ b/plugins/ltac/taccoerce.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Names open EConstr diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index f44ccbd3b..270225e23 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Pp open CErrors open Util @@ -417,7 +419,7 @@ let is_defined_tac kn = let warn_unusable_identifier = CWarnings.create ~name:"unusable-identifier" ~category:"parsing" - (fun id -> strbrk "The Ltac name" ++ spc () ++ pr_id id ++ spc () ++ + (fun id -> strbrk "The Ltac name" ++ spc () ++ Id.print id ++ spc () ++ strbrk "may be unusable because of a conflict with a notation.") let register_ltac local tacl = @@ -425,7 +427,7 @@ let register_ltac local tacl = match tactic_body with | Tacexpr.TacticDefinition ((loc,id), body) -> let kn = Lib.make_kn id in - let id_pp = pr_id id in + let id_pp = Id.print id in let () = if is_defined_tac kn then CErrors.user_err ?loc (str "There is already an Ltac named " ++ id_pp ++ str".") @@ -473,7 +475,7 @@ let register_ltac local tacl = let iter (def, tac) = match def with | NewTac id -> Tacenv.register_ltac false local id tac; - Flags.if_verbose Feedback.msg_info (Nameops.pr_id id ++ str " is defined") + Flags.if_verbose Feedback.msg_info (Id.print id ++ str " is defined") | UpdateTac kn -> Tacenv.redefine_ltac local kn tac; let name = Nametab.shortest_qualid_of_tactic kn in diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli index 07aa7ad82..c5223052c 100644 --- a/plugins/ltac/tacentries.mli +++ b/plugins/ltac/tacentries.mli @@ -8,6 +8,8 @@ (** Ltac toplevel command entries. *) +open API +open Grammar_API open Vernacexpr open Tacexpr diff --git a/plugins/ltac/tacenv.ml b/plugins/ltac/tacenv.ml index efb7e780d..14b5e00c7 100644 --- a/plugins/ltac/tacenv.ml +++ b/plugins/ltac/tacenv.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Pp open Names diff --git a/plugins/ltac/tacenv.mli b/plugins/ltac/tacenv.mli index d1e2a7bbe..2295852ce 100644 --- a/plugins/ltac/tacenv.mli +++ b/plugins/ltac/tacenv.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Tacexpr open Geninterp diff --git a/plugins/ltac/tacexpr.mli b/plugins/ltac/tacexpr.mli index cfb698cd8..67893bd11 100644 --- a/plugins/ltac/tacexpr.mli +++ b/plugins/ltac/tacexpr.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Loc open Names open Constrexpr @@ -385,7 +386,7 @@ type ltac_call_kind = | LtacNameCall of ltac_constant | LtacAtomCall of glob_atomic_tactic_expr | LtacVarCall of Id.t * glob_tactic_expr - | LtacConstrInterp of Glob_term.glob_constr * Pretyping.ltac_var_map + | LtacConstrInterp of Glob_term.glob_constr * Glob_term.ltac_var_map type ltac_trace = ltac_call_kind Loc.located list diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index d201cf949..bc1dd26d9 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Pattern open Pp open Genredexpr @@ -14,7 +16,6 @@ open Tacred open CErrors open Util open Names -open Nameops open Libnames open Globnames open Nametab diff --git a/plugins/ltac/tacintern.mli b/plugins/ltac/tacintern.mli index 8ad52ca02..1841ab42b 100644 --- a/plugins/ltac/tacintern.mli +++ b/plugins/ltac/tacintern.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Pp open Names open Tacexpr diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index ff76d06cf..0cd3ee2f9 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Constrintern open Patternops open Pp @@ -20,7 +22,6 @@ open Nameops open Libnames open Globnames open Nametab -open Pfedit open Refiner open Tacmach.New open Tactic_debug @@ -90,7 +91,7 @@ type value = Val.t (** Abstract application, to print ltac functions *) type appl = | UnnamedAppl (** For generic applications: nothing is printed *) - | GlbAppl of (Names.kernel_name * Val.t list) list + | GlbAppl of (Names.KerName.t * Val.t list) list (** For calls to global constants, some may alias other. *) let push_appl appl args = match appl with @@ -255,7 +256,7 @@ let pr_closure env ist body = let pr_sep () = fnl () in let pr_iarg (id, arg) = let arg = pr_argument_type arg in - hov 0 (pr_id id ++ spc () ++ str ":" ++ spc () ++ arg) + hov 0 (Id.print id ++ spc () ++ str ":" ++ spc () ++ arg) in let pp_iargs = v 0 (prlist_with_sep pr_sep pr_iarg (Id.Map.bindings ist)) in pp_body ++ fnl() ++ str "in environment " ++ fnl() ++ pp_iargs @@ -312,7 +313,7 @@ let append_trace trace v = let coerce_to_tactic loc id v = let v = Value.normalize v in let fail () = user_err ?loc - (str "Variable " ++ pr_id id ++ str " should be bound to a tactic.") + (str "Variable " ++ Id.print id ++ str " should be bound to a tactic.") in let v = Value.normalize v in if has_type v (topwit wit_tacvalue) then @@ -367,7 +368,7 @@ let debugging_exception_step ist signal_anomaly e pp = pp() ++ spc() ++ str "raised the exception" ++ fnl() ++ explain_exc e) let error_ltac_variable ?loc id env v s = - user_err ?loc (str "Ltac variable " ++ pr_id id ++ + user_err ?loc (str "Ltac variable " ++ Id.print id ++ strbrk " is bound to" ++ spc () ++ pr_value env v ++ spc () ++ strbrk "which cannot be coerced to " ++ str s ++ str".") @@ -401,7 +402,7 @@ let interp_int ist locid = try try_interp_ltac_var coerce_to_int ist None locid with Not_found -> user_err ?loc:(fst locid) ~hdr:"interp_int" - (str "Unbound variable " ++ pr_id (snd locid) ++ str".") + (str "Unbound variable " ++ Id.print (snd locid) ++ str".") let interp_int_or_var ist = function | ArgVar locid -> interp_int ist locid @@ -603,10 +604,10 @@ let interp_gen kind ist pattern_mode flags env sigma c = let { closure = constrvars ; term } = interp_glob_closure ist env sigma ~kind:kind_for_intern ~pattern_mode c in let vars = { - Pretyping.ltac_constrs = constrvars.typed; - Pretyping.ltac_uconstrs = constrvars.untyped; - Pretyping.ltac_idents = constrvars.idents; - Pretyping.ltac_genargs = ist.lfun; + Glob_term.ltac_constrs = constrvars.typed; + Glob_term.ltac_uconstrs = constrvars.untyped; + Glob_term.ltac_idents = constrvars.idents; + Glob_term.ltac_genargs = ist.lfun; } in (* Jason Gross: To avoid unnecessary modifications to tacinterp, as suggested by Arnaud Spiwack, we run push_trace immediately. We do @@ -627,7 +628,7 @@ let interp_gen kind ist pattern_mode flags env sigma c = let constr_flags () = { use_typeclasses = true; solve_unification_constraints = true; - use_hook = solve_by_implicit_tactic (); + use_hook = Pfedit.solve_by_implicit_tactic (); fail_evar = true; expand_evars = true } @@ -642,14 +643,14 @@ let interp_type = interp_constr_gen IsType let open_constr_use_classes_flags () = { use_typeclasses = true; solve_unification_constraints = true; - use_hook = solve_by_implicit_tactic (); + use_hook = Pfedit.solve_by_implicit_tactic (); fail_evar = false; expand_evars = true } let open_constr_no_classes_flags () = { use_typeclasses = false; solve_unification_constraints = true; - use_hook = solve_by_implicit_tactic (); + use_hook = Pfedit.solve_by_implicit_tactic (); fail_evar = false; expand_evars = true } @@ -780,7 +781,7 @@ let interp_may_eval f ist env sigma = function with | Not_found -> user_err ?loc ~hdr:"interp_may_eval" - (str "Unbound context identifier" ++ pr_id s ++ str".")) + (str "Unbound context identifier" ++ Id.print s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in let (sigma, t) = Typing.type_of ~refresh:true env sigma c_interp in @@ -856,7 +857,7 @@ let rec message_of_value v = end else if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in - Ftactic.enter begin fun gl -> Ftactic.return (pr_id id) end + Ftactic.enter begin fun gl -> Ftactic.return (Id.print id) end else match Value.to_list v with | Some l -> Ftactic.List.map message_of_value l >>= fun l -> @@ -871,7 +872,7 @@ let interp_message_token ist = function | MsgIdent (loc,id) -> let v = try Some (Id.Map.find id ist.lfun) with Not_found -> None in match v with - | None -> Ftactic.lift (Tacticals.New.tclZEROMSG (pr_id id ++ str" not found.")) + | None -> Ftactic.lift (Tacticals.New.tclZEROMSG (Id.print id ++ str" not found.")) | Some v -> message_of_value v let interp_message ist l = @@ -1008,7 +1009,7 @@ let interp_destruction_arg ist gl arg = | keep,ElimOnAnonHyp n as x -> x | keep,ElimOnIdent (loc,id) -> let error () = user_err ?loc - (strbrk "Cannot coerce " ++ pr_id id ++ + (strbrk "Cannot coerce " ++ Id.print id ++ strbrk " neither to a quantified hypothesis nor to a term.") in let try_cast_id id' = @@ -1019,7 +1020,7 @@ let interp_destruction_arg ist gl arg = try (sigma, (constr_of_id env id', NoBindings)) with Not_found -> user_err ?loc ~hdr:"interp_destruction_arg" ( - pr_id id ++ strbrk " binds to " ++ pr_id id' ++ strbrk " which is neither a declared nor a quantified hypothesis.") + Id.print id ++ strbrk " binds to " ++ Id.print id' ++ strbrk " which is neither a declared nor a quantified hypothesis.") end) in try @@ -1086,7 +1087,7 @@ let read_pattern lfun ist env sigma = function let cons_and_check_name id l = if Id.List.mem id l then user_err ~hdr:"read_match_goal_hyps" ( - str "Hypothesis pattern-matching variable " ++ pr_id id ++ + str "Hypothesis pattern-matching variable " ++ Id.print id ++ str " used twice in the same pattern.") else id::l diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli index fb50a6434..a1841afe3 100644 --- a/plugins/ltac/tacinterp.mli +++ b/plugins/ltac/tacinterp.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open Tactic_debug open EConstr diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml index 2858df313..6d33724f1 100644 --- a/plugins/ltac/tacsubst.ml +++ b/plugins/ltac/tacsubst.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API +open Grammar_API open Util open Tacexpr open Mod_subst diff --git a/plugins/ltac/tacsubst.mli b/plugins/ltac/tacsubst.mli index c1bf27257..2cfe8fac9 100644 --- a/plugins/ltac/tacsubst.mli +++ b/plugins/ltac/tacsubst.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Tacexpr open Mod_subst open Genarg diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml index e6d0370f3..53dc80023 100644 --- a/plugins/ltac/tactic_debug.ml +++ b/plugins/ltac/tactic_debug.ml @@ -6,13 +6,12 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Names open Pp open Tacexpr open Termops -open Nameops - let (ltac_trace_info : ltac_trace Exninfo.t) = Exninfo.make () @@ -258,14 +257,14 @@ let db_pattern_rule debug num r = (* Prints the hypothesis pattern identifier if it exists *) let hyp_bound = function | Anonymous -> str " (unbound)" - | Name id -> str " (bound to " ++ pr_id id ++ str ")" + | Name id -> str " (bound to " ++ Id.print id ++ str ")" (* Prints a matched hypothesis *) let db_matched_hyp debug env sigma (id,_,c) ido = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then - msg_tac_debug (str "Hypothesis " ++ pr_id id ++ hyp_bound ido ++ + msg_tac_debug (str "Hypothesis " ++ Id.print id ++ hyp_bound ido ++ str " has been matched: " ++ print_constr_env env sigma c) else return () @@ -360,18 +359,18 @@ let explain_ltac_call_trace last trace loc = | Tacexpr.LtacMLCall t -> quote (Pptactic.pr_glob_tactic (Global.env()) t) | Tacexpr.LtacVarCall (id,t) -> - quote (Nameops.pr_id id) ++ strbrk " (bound to " ++ + quote (Id.print id) ++ strbrk " (bound to " ++ Pptactic.pr_glob_tactic (Global.env()) t ++ str ")" | Tacexpr.LtacAtomCall te -> quote (Pptactic.pr_glob_tactic (Global.env()) (Tacexpr.TacAtom (Loc.tag te))) - | Tacexpr.LtacConstrInterp (c, { Pretyping.ltac_constrs = vars }) -> + | Tacexpr.LtacConstrInterp (c, { Glob_term.ltac_constrs = vars }) -> quote (Printer.pr_glob_constr_env (Global.env()) c) ++ (if not (Id.Map.is_empty vars) then strbrk " (with " ++ prlist_with_sep pr_comma (fun (id,c) -> - pr_id id ++ str ":=" ++ Printer.pr_lconstr_under_binders c) + Id.print id ++ str ":=" ++ Printer.pr_lconstr_under_binders c) (List.rev (Id.Map.bindings vars)) ++ str ")" else mt()) in diff --git a/plugins/ltac/tactic_debug.mli b/plugins/ltac/tactic_debug.mli index ac35464c4..6cfaed305 100644 --- a/plugins/ltac/tactic_debug.mli +++ b/plugins/ltac/tactic_debug.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Environ open Pattern open Names diff --git a/plugins/ltac/tactic_matching.ml b/plugins/ltac/tactic_matching.ml index 5b5cd06cc..6dcef414c 100644 --- a/plugins/ltac/tactic_matching.ml +++ b/plugins/ltac/tactic_matching.ml @@ -9,6 +9,7 @@ (** This file extends Matching with the main logic for Ltac's (lazy)match and (lazy)match goal. *) +open API open Names open Tacexpr open Context.Named.Declaration diff --git a/plugins/ltac/tactic_matching.mli b/plugins/ltac/tactic_matching.mli index 300b546f1..304eec463 100644 --- a/plugins/ltac/tactic_matching.mli +++ b/plugins/ltac/tactic_matching.mli @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + (** This file extends Matching with the main logic for Ltac's (lazy)match and (lazy)match goal. *) diff --git a/plugins/ltac/tactic_option.ml b/plugins/ltac/tactic_option.ml index a5ba3b837..53dfe22a9 100644 --- a/plugins/ltac/tactic_option.ml +++ b/plugins/ltac/tactic_option.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Libobject open Pp diff --git a/plugins/ltac/tactic_option.mli b/plugins/ltac/tactic_option.mli index ed759a76d..2817b54a1 100644 --- a/plugins/ltac/tactic_option.mli +++ b/plugins/ltac/tactic_option.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Tacexpr open Vernacexpr diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml index d8e21d81d..2a8ed7238 100644 --- a/plugins/ltac/tauto.ml +++ b/plugins/ltac/tauto.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Term open EConstr open Hipattern @@ -65,7 +66,7 @@ let negation_unfolding = ref true (* Whether inner iff are unfolded *) let iff_unfolding = ref false -let unfold_iff () = !iff_unfolding || Flags.version_less_or_equal Flags.V8_2 +let unfold_iff () = !iff_unfolding open Goptions let _ = @@ -78,7 +79,7 @@ let _ = let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "unfolding of iff in intuition"; optkey = ["Intuition";"Iff";"Unfolding"]; optread = (fun () -> !iff_unfolding); @@ -196,7 +197,7 @@ let flatten_contravariant_disj _ ist = let make_unfold name = let dir = DirPath.make (List.map Id.of_string ["Logic"; "Init"; "Coq"]) in - let const = Constant.make2 (MPfile dir) (Label.make name) in + let const = Constant.make2 (ModPath.MPfile dir) (Label.make name) in (Locus.AllOccurrences, ArgArg (EvalConstRef const, None)) let u_iff = make_unfold "iff" diff --git a/plugins/ltac/vo.itarget b/plugins/ltac/vo.itarget deleted file mode 100644 index a28fb770b..000000000 --- a/plugins/ltac/vo.itarget +++ /dev/null @@ -1 +0,0 @@ -Ltac.vo diff --git a/plugins/micromega/MExtraction.v b/plugins/micromega/MExtraction.v index 4d5c3b1d5..95f135c8f 100644 --- a/plugins/micromega/MExtraction.v +++ b/plugins/micromega/MExtraction.v @@ -14,6 +14,7 @@ (* Used to generate micromega.ml *) +Require Extraction. Require Import ZMicromega. Require Import QMicromega. Require Import RMicromega. @@ -48,7 +49,10 @@ Extract Constant Rmult => "( * )". Extract Constant Ropp => "fun x -> - x". Extract Constant Rinv => "fun x -> 1 / x". -Extraction "plugins/micromega/micromega.ml" +(** We now extract to stdout, see comment in Makefile.build *) + +(*Extraction "plugins/micromega/micromega.ml" *) +Recursive Extraction List.map simpl_cone (*map_cone indexes*) denorm Qpower vm_add n_of_Z N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find. diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 83f374346..fba1966df 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -16,11 +16,11 @@ (* *) (************************************************************************) +open API open Pp open Mutils open Goptions - -module Term = EConstr +open Names (** * Debug flag @@ -109,8 +109,8 @@ type 'cst atom = 'cst Micromega.formula type 'cst formula = | TT | FF - | X of Term.constr - | A of 'cst atom * tag * Term.constr + | X of EConstr.constr + | A of 'cst atom * tag * EConstr.constr | C of 'cst formula * 'cst formula | D of 'cst formula * 'cst formula | N of 'cst formula @@ -328,9 +328,6 @@ let selecti s m = module M = struct - open Constr - open EConstr - (** * Location of the Coq libraries. *) @@ -602,10 +599,10 @@ struct let get_left_construct sigma term = match EConstr.kind sigma term with - | Constr.Construct((_,i),_) -> (i,[| |]) - | Constr.App(l,rst) -> + | Term.Construct((_,i),_) -> (i,[| |]) + | Term.App(l,rst) -> (match EConstr.kind sigma l with - | Constr.Construct((_,i),_) -> (i,rst) + | Term.Construct((_,i),_) -> (i,rst) | _ -> raise ParseError ) | _ -> raise ParseError @@ -626,7 +623,7 @@ struct let rec dump_nat x = match x with | Mc.O -> Lazy.force coq_O - | Mc.S p -> Term.mkApp(Lazy.force coq_S,[| dump_nat p |]) + | Mc.S p -> EConstr.mkApp(Lazy.force coq_S,[| dump_nat p |]) let rec parse_positive sigma term = let (i,c) = get_left_construct sigma term in @@ -639,28 +636,28 @@ struct let rec dump_positive x = match x with | Mc.XH -> Lazy.force coq_xH - | Mc.XO p -> Term.mkApp(Lazy.force coq_xO,[| dump_positive p |]) - | Mc.XI p -> Term.mkApp(Lazy.force coq_xI,[| dump_positive p |]) + | Mc.XO p -> EConstr.mkApp(Lazy.force coq_xO,[| dump_positive p |]) + | Mc.XI p -> EConstr.mkApp(Lazy.force coq_xI,[| dump_positive p |]) let pp_positive o x = Printf.fprintf o "%i" (CoqToCaml.positive x) let dump_n x = match x with | Mc.N0 -> Lazy.force coq_N0 - | Mc.Npos p -> Term.mkApp(Lazy.force coq_Npos,[| dump_positive p|]) + | Mc.Npos p -> EConstr.mkApp(Lazy.force coq_Npos,[| dump_positive p|]) let rec dump_index x = match x with | Mc.XH -> Lazy.force coq_xH - | Mc.XO p -> Term.mkApp(Lazy.force coq_xO,[| dump_index p |]) - | Mc.XI p -> Term.mkApp(Lazy.force coq_xI,[| dump_index p |]) + | Mc.XO p -> EConstr.mkApp(Lazy.force coq_xO,[| dump_index p |]) + | Mc.XI p -> EConstr.mkApp(Lazy.force coq_xI,[| dump_index p |]) let pp_index o x = Printf.fprintf o "%i" (CoqToCaml.index x) let pp_n o x = output_string o (string_of_int (CoqToCaml.n x)) let dump_pair t1 t2 dump_t1 dump_t2 (x,y) = - Term.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|]) + EConstr.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|]) let parse_z sigma term = let (i,c) = get_left_construct sigma term in @@ -673,23 +670,23 @@ struct let dump_z x = match x with | Mc.Z0 ->Lazy.force coq_ZERO - | Mc.Zpos p -> Term.mkApp(Lazy.force coq_POS,[| dump_positive p|]) - | Mc.Zneg p -> Term.mkApp(Lazy.force coq_NEG,[| dump_positive p|]) + | Mc.Zpos p -> EConstr.mkApp(Lazy.force coq_POS,[| dump_positive p|]) + | Mc.Zneg p -> EConstr.mkApp(Lazy.force coq_NEG,[| dump_positive p|]) let pp_z o x = Printf.fprintf o "%s" (Big_int.string_of_big_int (CoqToCaml.z_big_int x)) let dump_num bd1 = - Term.mkApp(Lazy.force coq_Qmake, - [|dump_z (CamlToCoq.bigint (numerator bd1)) ; - dump_positive (CamlToCoq.positive_big_int (denominator bd1)) |]) + EConstr.mkApp(Lazy.force coq_Qmake, + [|dump_z (CamlToCoq.bigint (numerator bd1)) ; + dump_positive (CamlToCoq.positive_big_int (denominator bd1)) |]) let dump_q q = - Term.mkApp(Lazy.force coq_Qmake, - [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|]) + EConstr.mkApp(Lazy.force coq_Qmake, + [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|]) let parse_q sigma term = match EConstr.kind sigma term with - | Constr.App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then + | Term.App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then {Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) } else raise ParseError | _ -> raise ParseError @@ -712,13 +709,13 @@ struct match cst with | Mc.C0 -> Lazy.force coq_C0 | Mc.C1 -> Lazy.force coq_C1 - | Mc.CQ q -> Term.mkApp(Lazy.force coq_CQ, [| dump_q q |]) - | Mc.CZ z -> Term.mkApp(Lazy.force coq_CZ, [| dump_z z |]) - | Mc.CPlus(x,y) -> Term.mkApp(Lazy.force coq_CPlus, [| dump_Rcst x ; dump_Rcst y |]) - | Mc.CMinus(x,y) -> Term.mkApp(Lazy.force coq_CMinus, [| dump_Rcst x ; dump_Rcst y |]) - | Mc.CMult(x,y) -> Term.mkApp(Lazy.force coq_CMult, [| dump_Rcst x ; dump_Rcst y |]) - | Mc.CInv t -> Term.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |]) - | Mc.COpp t -> Term.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |]) + | Mc.CQ q -> EConstr.mkApp(Lazy.force coq_CQ, [| dump_q q |]) + | Mc.CZ z -> EConstr.mkApp(Lazy.force coq_CZ, [| dump_z z |]) + | Mc.CPlus(x,y) -> EConstr.mkApp(Lazy.force coq_CPlus, [| dump_Rcst x ; dump_Rcst y |]) + | Mc.CMinus(x,y) -> EConstr.mkApp(Lazy.force coq_CMinus, [| dump_Rcst x ; dump_Rcst y |]) + | Mc.CMult(x,y) -> EConstr.mkApp(Lazy.force coq_CMult, [| dump_Rcst x ; dump_Rcst y |]) + | Mc.CInv t -> EConstr.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |]) + | Mc.COpp t -> EConstr.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |]) let rec parse_Rcst sigma term = let (i,c) = get_left_construct sigma term in @@ -745,8 +742,8 @@ struct let rec dump_list typ dump_elt l = match l with - | [] -> Term.mkApp(Lazy.force coq_nil,[| typ |]) - | e :: l -> Term.mkApp(Lazy.force coq_cons, + | [] -> EConstr.mkApp(Lazy.force coq_nil,[| typ |]) + | e :: l -> EConstr.mkApp(Lazy.force coq_cons, [| typ; dump_elt e;dump_list typ dump_elt l|]) let pp_list op cl elt o l = @@ -776,27 +773,27 @@ struct let dump_expr typ dump_z e = let rec dump_expr e = match e with - | Mc.PEX n -> mkApp(Lazy.force coq_PEX,[| typ; dump_var n |]) - | Mc.PEc z -> mkApp(Lazy.force coq_PEc,[| typ ; dump_z z |]) - | Mc.PEadd(e1,e2) -> mkApp(Lazy.force coq_PEadd, - [| typ; dump_expr e1;dump_expr e2|]) - | Mc.PEsub(e1,e2) -> mkApp(Lazy.force coq_PEsub, - [| typ; dump_expr e1;dump_expr e2|]) - | Mc.PEopp e -> mkApp(Lazy.force coq_PEopp, - [| typ; dump_expr e|]) - | Mc.PEmul(e1,e2) -> mkApp(Lazy.force coq_PEmul, - [| typ; dump_expr e1;dump_expr e2|]) - | Mc.PEpow(e,n) -> mkApp(Lazy.force coq_PEpow, - [| typ; dump_expr e; dump_n n|]) + | Mc.PEX n -> EConstr.mkApp(Lazy.force coq_PEX,[| typ; dump_var n |]) + | Mc.PEc z -> EConstr.mkApp(Lazy.force coq_PEc,[| typ ; dump_z z |]) + | Mc.PEadd(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEadd, + [| typ; dump_expr e1;dump_expr e2|]) + | Mc.PEsub(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEsub, + [| typ; dump_expr e1;dump_expr e2|]) + | Mc.PEopp e -> EConstr.mkApp(Lazy.force coq_PEopp, + [| typ; dump_expr e|]) + | Mc.PEmul(e1,e2) -> EConstr.mkApp(Lazy.force coq_PEmul, + [| typ; dump_expr e1;dump_expr e2|]) + | Mc.PEpow(e,n) -> EConstr.mkApp(Lazy.force coq_PEpow, + [| typ; dump_expr e; dump_n n|]) in dump_expr e let dump_pol typ dump_c e = let rec dump_pol e = match e with - | Mc.Pc n -> mkApp(Lazy.force coq_Pc, [|typ ; dump_c n|]) - | Mc.Pinj(p,pol) -> mkApp(Lazy.force coq_Pinj , [| typ ; dump_positive p ; dump_pol pol|]) - | Mc.PX(pol1,p,pol2) -> mkApp(Lazy.force coq_PX, [| typ ; dump_pol pol1 ; dump_positive p ; dump_pol pol2|]) in + | Mc.Pc n -> EConstr.mkApp(Lazy.force coq_Pc, [|typ ; dump_c n|]) + | Mc.Pinj(p,pol) -> EConstr.mkApp(Lazy.force coq_Pinj , [| typ ; dump_positive p ; dump_pol pol|]) + | Mc.PX(pol1,p,pol2) -> EConstr.mkApp(Lazy.force coq_PX, [| typ ; dump_pol pol1 ; dump_positive p ; dump_pol pol2|]) in dump_pol e let pp_pol pp_c o e = @@ -815,17 +812,17 @@ struct let z = Lazy.force typ in let rec dump_cone e = match e with - | Mc.PsatzIn n -> mkApp(Lazy.force coq_PsatzIn,[| z; dump_nat n |]) - | Mc.PsatzMulC(e,c) -> mkApp(Lazy.force coq_PsatzMultC, - [| z; dump_pol z dump_z e ; dump_cone c |]) - | Mc.PsatzSquare e -> mkApp(Lazy.force coq_PsatzSquare, - [| z;dump_pol z dump_z e|]) - | Mc.PsatzAdd(e1,e2) -> mkApp(Lazy.force coq_PsatzAdd, - [| z; dump_cone e1; dump_cone e2|]) - | Mc.PsatzMulE(e1,e2) -> mkApp(Lazy.force coq_PsatzMulE, - [| z; dump_cone e1; dump_cone e2|]) - | Mc.PsatzC p -> mkApp(Lazy.force coq_PsatzC,[| z; dump_z p|]) - | Mc.PsatzZ -> mkApp( Lazy.force coq_PsatzZ,[| z|]) in + | Mc.PsatzIn n -> EConstr.mkApp(Lazy.force coq_PsatzIn,[| z; dump_nat n |]) + | Mc.PsatzMulC(e,c) -> EConstr.mkApp(Lazy.force coq_PsatzMultC, + [| z; dump_pol z dump_z e ; dump_cone c |]) + | Mc.PsatzSquare e -> EConstr.mkApp(Lazy.force coq_PsatzSquare, + [| z;dump_pol z dump_z e|]) + | Mc.PsatzAdd(e1,e2) -> EConstr.mkApp(Lazy.force coq_PsatzAdd, + [| z; dump_cone e1; dump_cone e2|]) + | Mc.PsatzMulE(e1,e2) -> EConstr.mkApp(Lazy.force coq_PsatzMulE, + [| z; dump_cone e1; dump_cone e2|]) + | Mc.PsatzC p -> EConstr.mkApp(Lazy.force coq_PsatzC,[| z; dump_z p|]) + | Mc.PsatzZ -> EConstr.mkApp(Lazy.force coq_PsatzZ,[| z|]) in dump_cone e let pp_psatz pp_z o e = @@ -868,10 +865,10 @@ struct Printf.fprintf o"(%a %a %a)" (pp_expr pp_z) l pp_op op (pp_expr pp_z) r let dump_cstr typ dump_constant {Mc.flhs = e1 ; Mc.fop = o ; Mc.frhs = e2} = - Term.mkApp(Lazy.force coq_Build, - [| typ; dump_expr typ dump_constant e1 ; - dump_op o ; - dump_expr typ dump_constant e2|]) + EConstr.mkApp(Lazy.force coq_Build, + [| typ; dump_expr typ dump_constant e1 ; + dump_op o ; + dump_expr typ dump_constant e2|]) let assoc_const sigma x l = try @@ -905,8 +902,8 @@ struct let parse_zop gl (op,args) = let sigma = gl.sigma in match EConstr.kind sigma op with - | Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) - | Ind((n,0),_) -> + | Term.Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) + | Term.Ind((n,0),_) -> if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError @@ -915,8 +912,8 @@ struct let parse_rop gl (op,args) = let sigma = gl.sigma in match EConstr.kind sigma op with - | Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) - | Ind((n,0),_) -> + | Term.Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) + | Term.Ind((n,0),_) -> if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError @@ -927,7 +924,7 @@ struct let is_constant sigma t = (* This is an approx *) match EConstr.kind sigma t with - | Construct(i,_) -> true + | Term.Construct(i,_) -> true | _ -> false type 'a op = @@ -948,14 +945,14 @@ struct module Env = struct - type t = constr list + type t = EConstr.constr list let compute_rank_add env sigma v = let rec _add env n v = match env with | [] -> ([v],n) | e::l -> - if eq_constr sigma e v + if EConstr.eq_constr sigma e v then (env,n) else let (env,n) = _add l ( n+1) v in @@ -969,7 +966,7 @@ struct match env with | [] -> raise (Invalid_argument "get_rank") | e::l -> - if eq_constr sigma e v + if EConstr.eq_constr sigma e v then n else _get_rank l (n+1) in _get_rank env 1 @@ -1010,10 +1007,10 @@ struct try (Mc.PEc (parse_constant term) , env) with ParseError -> match EConstr.kind sigma term with - | App(t,args) -> + | Term.App(t,args) -> ( match EConstr.kind sigma t with - | Const c -> + | Term.Const c -> ( match assoc_ops sigma t ops_spec with | Binop f -> combine env f (args.(0),args.(1)) | Opp -> let (expr,env) = parse_expr env args.(0) in @@ -1076,13 +1073,13 @@ struct let rec rconstant sigma term = match EConstr.kind sigma term with - | Const x -> + | Term.Const x -> if EConstr.eq_constr sigma term (Lazy.force coq_R0) then Mc.C0 else if EConstr.eq_constr sigma term (Lazy.force coq_R1) then Mc.C1 else raise ParseError - | App(op,args) -> + | Term.App(op,args) -> begin try (* the evaluation order is important in the following *) @@ -1151,7 +1148,7 @@ struct if debug then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr cstr ++ fnl ()); match EConstr.kind sigma cstr with - | App(op,args) -> + | Term.App(op,args) -> let (op,lhs,rhs) = parse_op gl (op,args) in let (e1,env) = parse_expr sigma env lhs in let (e2,env) = parse_expr sigma env rhs in @@ -1206,29 +1203,29 @@ struct let rec xparse_formula env tg term = match EConstr.kind sigma term with - | App(l,rst) -> + | Term.App(l,rst) -> (match rst with - | [|a;b|] when eq_constr sigma l (Lazy.force coq_and) -> + | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_and) -> let f,env,tg = xparse_formula env tg a in let g,env, tg = xparse_formula env tg b in mkformula_binary mkC term f g,env,tg - | [|a;b|] when eq_constr sigma l (Lazy.force coq_or) -> + | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_or) -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkD term f g,env,tg - | [|a|] when eq_constr sigma l (Lazy.force coq_not) -> + | [|a|] when EConstr.eq_constr sigma l (Lazy.force coq_not) -> let (f,env,tg) = xparse_formula env tg a in (N(f), env,tg) - | [|a;b|] when eq_constr sigma l (Lazy.force coq_iff) -> + | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_iff) -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkIff term f g,env,tg | _ -> parse_atom env tg term) - | Prod(typ,a,b) when Vars.noccurn sigma 1 b -> + | Term.Prod(typ,a,b) when EConstr.Vars.noccurn sigma 1 b -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkI term f g,env,tg - | _ when eq_constr sigma term (Lazy.force coq_True) -> (TT,env,tg) - | _ when eq_constr sigma term (Lazy.force coq_False) -> (FF,env,tg) + | _ when EConstr.eq_constr sigma term (Lazy.force coq_True) -> (TT,env,tg) + | _ when EConstr.eq_constr sigma term (Lazy.force coq_False) -> (FF,env,tg) | _ when is_prop term -> X(term),env,tg | _ -> raise ParseError in @@ -1237,14 +1234,14 @@ struct let dump_formula typ dump_atom f = let rec xdump f = match f with - | TT -> mkApp(Lazy.force coq_TT,[|typ|]) - | FF -> mkApp(Lazy.force coq_FF,[|typ|]) - | C(x,y) -> mkApp(Lazy.force coq_And,[|typ ; xdump x ; xdump y|]) - | D(x,y) -> mkApp(Lazy.force coq_Or,[|typ ; xdump x ; xdump y|]) - | I(x,_,y) -> mkApp(Lazy.force coq_Impl,[|typ ; xdump x ; xdump y|]) - | N(x) -> mkApp(Lazy.force coq_Neg,[|typ ; xdump x|]) - | A(x,_,_) -> mkApp(Lazy.force coq_Atom,[|typ ; dump_atom x|]) - | X(t) -> mkApp(Lazy.force coq_X,[|typ ; t|]) in + | TT -> EConstr.mkApp(Lazy.force coq_TT,[|typ|]) + | FF -> EConstr.mkApp(Lazy.force coq_FF,[|typ|]) + | C(x,y) -> EConstr.mkApp(Lazy.force coq_And,[|typ ; xdump x ; xdump y|]) + | D(x,y) -> EConstr.mkApp(Lazy.force coq_Or,[|typ ; xdump x ; xdump y|]) + | I(x,_,y) -> EConstr.mkApp(Lazy.force coq_Impl,[|typ ; xdump x ; xdump y|]) + | N(x) -> EConstr.mkApp(Lazy.force coq_Neg,[|typ ; xdump x|]) + | A(x,_,_) -> EConstr.mkApp(Lazy.force coq_Atom,[|typ ; dump_atom x|]) + | X(t) -> EConstr.mkApp(Lazy.force coq_X,[|typ ; t|]) in xdump f @@ -1284,15 +1281,15 @@ struct type 'cst dump_expr = (* 'cst is the type of the syntactic constants *) { - interp_typ : constr; - dump_cst : 'cst -> constr; - dump_add : constr; - dump_sub : constr; - dump_opp : constr; - dump_mul : constr; - dump_pow : constr; - dump_pow_arg : Mc.n -> constr; - dump_op : (Mc.op2 * Term.constr) list + interp_typ : EConstr.constr; + dump_cst : 'cst -> EConstr.constr; + dump_add : EConstr.constr; + dump_sub : EConstr.constr; + dump_opp : EConstr.constr; + dump_mul : EConstr.constr; + dump_pow : EConstr.constr; + dump_pow_arg : Mc.n -> EConstr.constr; + dump_op : (Mc.op2 * EConstr.constr) list } let dump_zexpr = lazy @@ -1326,8 +1323,8 @@ let dump_qexpr = lazy let add = Lazy.force coq_Rplus in let one = Lazy.force coq_R1 in - let mk_add x y = mkApp(add,[|x;y|]) in - let mk_mult x y = mkApp(mult,[|x;y|]) in + let mk_add x y = EConstr.mkApp(add,[|x;y|]) in + let mk_mult x y = EConstr.mkApp(mult,[|x;y|]) in let two = mk_add one one in @@ -1350,13 +1347,13 @@ let rec dump_Rcst_as_R cst = match cst with | Mc.C0 -> Lazy.force coq_R0 | Mc.C1 -> Lazy.force coq_R1 - | Mc.CQ q -> Term.mkApp(Lazy.force coq_IQR, [| dump_q q |]) - | Mc.CZ z -> Term.mkApp(Lazy.force coq_IZR, [| dump_z z |]) - | Mc.CPlus(x,y) -> Term.mkApp(Lazy.force coq_Rplus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |]) - | Mc.CMinus(x,y) -> Term.mkApp(Lazy.force coq_Rminus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |]) - | Mc.CMult(x,y) -> Term.mkApp(Lazy.force coq_Rmult, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |]) - | Mc.CInv t -> Term.mkApp(Lazy.force coq_Rinv, [| dump_Rcst_as_R t |]) - | Mc.COpp t -> Term.mkApp(Lazy.force coq_Ropp, [| dump_Rcst_as_R t |]) + | Mc.CQ q -> EConstr.mkApp(Lazy.force coq_IQR, [| dump_q q |]) + | Mc.CZ z -> EConstr.mkApp(Lazy.force coq_IZR, [| dump_z z |]) + | Mc.CPlus(x,y) -> EConstr.mkApp(Lazy.force coq_Rplus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |]) + | Mc.CMinus(x,y) -> EConstr.mkApp(Lazy.force coq_Rminus, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |]) + | Mc.CMult(x,y) -> EConstr.mkApp(Lazy.force coq_Rmult, [| dump_Rcst_as_R x ; dump_Rcst_as_R y |]) + | Mc.CInv t -> EConstr.mkApp(Lazy.force coq_Rinv, [| dump_Rcst_as_R t |]) + | Mc.COpp t -> EConstr.mkApp(Lazy.force coq_Ropp, [| dump_Rcst_as_R t |]) let dump_rexpr = lazy @@ -1385,7 +1382,7 @@ let dump_rexpr = lazy let prodn n env b = let rec prodrec = function | (0, env, b) -> b - | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b)) + | (n, ((v,t)::l), b) -> prodrec (n-1, l, EConstr.mkProd (v,t,b)) | _ -> assert false in prodrec (n,env,b) @@ -1399,32 +1396,32 @@ let make_goal_of_formula sigma dexpr form = let props = prop_env_of_formula sigma form in - let vars_n = List.map (fun (_,i) -> (Names.id_of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in - let props_n = List.mapi (fun i _ -> (Names.id_of_string (Printf.sprintf "__p%i" (i+1))) , Term.mkProp) props in + let vars_n = List.map (fun (_,i) -> (Names.Id.of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in + let props_n = List.mapi (fun i _ -> (Names.Id.of_string (Printf.sprintf "__p%i" (i+1))) , EConstr.mkProp) props in let var_name_pos = List.map2 (fun (idx,_) (id,_) -> id,idx) vars_idx vars_n in let dump_expr i e = let rec dump_expr = function - | Mc.PEX n -> mkRel (i+(List.assoc (CoqToCaml.positive n) vars_idx)) + | Mc.PEX n -> EConstr.mkRel (i+(List.assoc (CoqToCaml.positive n) vars_idx)) | Mc.PEc z -> dexpr.dump_cst z - | Mc.PEadd(e1,e2) -> mkApp(dexpr.dump_add, + | Mc.PEadd(e1,e2) -> EConstr.mkApp(dexpr.dump_add, [| dump_expr e1;dump_expr e2|]) - | Mc.PEsub(e1,e2) -> mkApp(dexpr.dump_sub, + | Mc.PEsub(e1,e2) -> EConstr.mkApp(dexpr.dump_sub, [| dump_expr e1;dump_expr e2|]) - | Mc.PEopp e -> mkApp(dexpr.dump_opp, - [| dump_expr e|]) - | Mc.PEmul(e1,e2) -> mkApp(dexpr.dump_mul, - [| dump_expr e1;dump_expr e2|]) - | Mc.PEpow(e,n) -> mkApp(dexpr.dump_pow, - [| dump_expr e; dexpr.dump_pow_arg n|]) + | Mc.PEopp e -> EConstr.mkApp(dexpr.dump_opp, + [| dump_expr e|]) + | Mc.PEmul(e1,e2) -> EConstr.mkApp(dexpr.dump_mul, + [| dump_expr e1;dump_expr e2|]) + | Mc.PEpow(e,n) -> EConstr.mkApp(dexpr.dump_pow, + [| dump_expr e; dexpr.dump_pow_arg n|]) in dump_expr e in let mkop op e1 e2 = try - Term.mkApp(List.assoc op dexpr.dump_op, [| e1; e2|]) + EConstr.mkApp(List.assoc op dexpr.dump_op, [| e1; e2|]) with Not_found -> - Term.mkApp(Lazy.force coq_Eq,[|dexpr.interp_typ ; e1 ;e2|]) in + EConstr.mkApp(Lazy.force coq_Eq,[|dexpr.interp_typ ; e1 ;e2|]) in let dump_cstr i { Mc.flhs ; Mc.fop ; Mc.frhs } = mkop fop (dump_expr i flhs) (dump_expr i frhs) in @@ -1433,13 +1430,13 @@ let make_goal_of_formula sigma dexpr form = match f with | TT -> Lazy.force coq_True | FF -> Lazy.force coq_False - | C(x,y) -> mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|]) - | D(x,y) -> mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|]) - | I(x,_,y) -> mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y) - | N(x) -> mkArrow (xdump pi xi x) (Lazy.force coq_False) + | C(x,y) -> EConstr.mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|]) + | D(x,y) -> EConstr.mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|]) + | I(x,_,y) -> EConstr.mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y) + | N(x) -> EConstr.mkArrow (xdump pi xi x) (Lazy.force coq_False) | A(x,_,_) -> dump_cstr xi x | X(t) -> let idx = Env.get_rank props sigma t in - mkRel (pi+idx) in + EConstr.mkRel (pi+idx) in let nb_vars = List.length vars_n in let nb_props = List.length props_n in @@ -1448,12 +1445,12 @@ let make_goal_of_formula sigma dexpr form = let subst_prop p = let idx = Env.get_rank props sigma p in - mkVar (Names.id_of_string (Printf.sprintf "__p%i" idx)) in + EConstr.mkVar (Names.Id.of_string (Printf.sprintf "__p%i" idx)) in let form' = map_prop subst_prop form in - (prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n) - (prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n) + (prodn nb_props (List.map (fun (x,y) -> Name.Name x,y) props_n) + (prodn nb_vars (List.map (fun (x,y) -> Name.Name x,y) vars_n) (xdump (List.length vars_n) 0 form)), List.rev props_n, List.rev var_name_pos,form') @@ -1468,7 +1465,7 @@ let make_goal_of_formula sigma dexpr form = | [] -> acc | (e::l) -> let (name,expr,typ) = e in - xset (Term.mkNamedLetIn + xset (EConstr.mkNamedLetIn (Names.Id.of_string name) expr typ acc) l in xset concl l @@ -1544,10 +1541,10 @@ let coq_VarMap = let rec dump_varmap typ m = match m with - | Mc.Empty -> Term.mkApp(Lazy.force coq_Empty,[| typ |]) - | Mc.Leaf v -> Term.mkApp(Lazy.force coq_Leaf,[| typ; v|]) + | Mc.Empty -> EConstr.mkApp(Lazy.force coq_Empty,[| typ |]) + | Mc.Leaf v -> EConstr.mkApp(Lazy.force coq_Leaf,[| typ; v|]) | Mc.Node(l,o,r) -> - Term.mkApp (Lazy.force coq_Node, [| typ; dump_varmap typ l; o ; dump_varmap typ r |]) + EConstr.mkApp (Lazy.force coq_Node, [| typ; dump_varmap typ l; o ; dump_varmap typ r |]) let vm_of_list env = @@ -1569,15 +1566,15 @@ let rec pp_varmap o vm = let rec dump_proof_term = function | Micromega.DoneProof -> Lazy.force coq_doneProof | Micromega.RatProof(cone,rst) -> - Term.mkApp(Lazy.force coq_ratProof, [| dump_psatz coq_Z dump_z cone; dump_proof_term rst|]) + EConstr.mkApp(Lazy.force coq_ratProof, [| dump_psatz coq_Z dump_z cone; dump_proof_term rst|]) | Micromega.CutProof(cone,prf) -> - Term.mkApp(Lazy.force coq_cutProof, + EConstr.mkApp(Lazy.force coq_cutProof, [| dump_psatz coq_Z dump_z cone ; dump_proof_term prf|]) | Micromega.EnumProof(c1,c2,prfs) -> - Term.mkApp (Lazy.force coq_enumProof, - [| dump_psatz coq_Z dump_z c1 ; dump_psatz coq_Z dump_z c2 ; - dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |]) + EConstr.mkApp (Lazy.force coq_enumProof, + [| dump_psatz coq_Z dump_z c1 ; dump_psatz coq_Z dump_z c2 ; + dump_list (Lazy.force coq_proofTerm) dump_proof_term prfs |]) let rec size_of_psatz = function @@ -1637,11 +1634,11 @@ let parse_goal gl parse_arith env hyps term = * The datastructures that aggregate theory-dependent proof values. *) type ('synt_c, 'prf) domain_spec = { - typ : Term.constr; (* is the type of the interpretation domain - Z, Q, R*) - coeff : Term.constr ; (* is the type of the syntactic coeffs - Z , Q , Rcst *) - dump_coeff : 'synt_c -> Term.constr ; - proof_typ : Term.constr ; - dump_proof : 'prf -> Term.constr + typ : EConstr.constr; (* is the type of the interpretation domain - Z, Q, R*) + coeff : EConstr.constr ; (* is the type of the syntactic coeffs - Z , Q , Rcst *) + dump_coeff : 'synt_c -> EConstr.constr ; + proof_typ : EConstr.constr ; + dump_proof : 'prf -> EConstr.constr } let zz_domain_spec = lazy { @@ -1706,7 +1703,7 @@ let topo_sort_constr l = let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic*) = (* let ids = Util.List.map_i (fun i _ -> (Names.Id.of_string ("__v"^(string_of_int i)))) 0 env in *) - let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[|spec.coeff|])) in + let formula_typ = (EConstr.mkApp (Lazy.force coq_Cstr,[|spec.coeff|])) in let ff = dump_formula formula_typ (dump_cstr spec.coeff spec.dump_coeff) ff in let vm = dump_varmap (spec.typ) (vm_of_list env) in (* todo : directly generate the proof term - or generalize before conversion? *) @@ -1716,8 +1713,8 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* Tactics.change_concl (set [ - ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |])); - ("__varmap", vm, Term.mkApp(Lazy.force coq_VarMap, [|spec.typ|])); + ("__ff", ff, EConstr.mkApp(Lazy.force coq_Formula, [|formula_typ |])); + ("__varmap", vm, EConstr.mkApp(Lazy.force coq_VarMap, [|spec.typ|])); ("__wit", cert, cert_typ) ] (Tacmach.New.pf_concl gl)) @@ -1841,20 +1838,20 @@ let abstract_formula hyps f = | A(a,t,term) -> if TagSet.mem t hyps then A(a,t,term) else X(term) | C(f1,f2) -> (match xabs f1 , xabs f2 with - | X a1 , X a2 -> X (Term.mkApp(Lazy.force coq_and, [|a1;a2|])) + | X a1 , X a2 -> X (EConstr.mkApp(Lazy.force coq_and, [|a1;a2|])) | f1 , f2 -> C(f1,f2) ) | D(f1,f2) -> (match xabs f1 , xabs f2 with - | X a1 , X a2 -> X (Term.mkApp(Lazy.force coq_or, [|a1;a2|])) + | X a1 , X a2 -> X (EConstr.mkApp(Lazy.force coq_or, [|a1;a2|])) | f1 , f2 -> D(f1,f2) ) | N(f) -> (match xabs f with - | X a -> X (Term.mkApp(Lazy.force coq_not, [|a|])) + | X a -> X (EConstr.mkApp(Lazy.force coq_not, [|a|])) | f -> N f) | I(f1,hyp,f2) -> (match xabs f1 , hyp, xabs f2 with | X a1 , Some _ , af2 -> af2 - | X a1 , None , X a2 -> X (Term.mkArrow a1 a2) + | X a1 , None , X a2 -> X (EConstr.mkArrow a1 a2) | af1 , _ , af2 -> I(af1,hyp,af2) ) | FF -> FF @@ -1908,7 +1905,7 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2 if debug then begin Feedback.msg_notice (Pp.str "Formula....\n") ; - let formula_typ = (Term.mkApp(Lazy.force coq_Cstr, [|spec.coeff|])) in + let formula_typ = (EConstr.mkApp(Lazy.force coq_Cstr, [|spec.coeff|])) in let ff = dump_formula formula_typ (dump_cstr spec.typ spec.dump_coeff) ff in Feedback.msg_notice (Printer.pr_leconstr ff); @@ -1933,7 +1930,7 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2 if debug then begin Feedback.msg_notice (Pp.str "\nAFormula\n") ; - let formula_typ = (Term.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in + let formula_typ = (EConstr.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in let ff' = dump_formula formula_typ (dump_cstr spec.typ spec.dump_coeff) ff' in Feedback.msg_notice (Printer.pr_leconstr ff'); @@ -1991,11 +1988,11 @@ let micromega_gen let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in let goal_name = fresh_id [] (Names.Id.of_string "__arith") gl in - let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in + let env' = List.map (fun (id,i) -> EConstr.mkVar id,i) vars in let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; micromega_order_change spec res' - (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in + (EConstr.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in let goal_props = List.rev (prop_env_of_formula sigma ff') in @@ -2014,8 +2011,8 @@ let micromega_gen [ kill_arith; (Tacticals.New.tclTHENLIST - [(Tactics.generalize (List.map Term.mkVar ids)); - Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args)) + [(Tactics.generalize (List.map EConstr.mkVar ids)); + Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args)) ] ) ] with @@ -2043,9 +2040,9 @@ let micromega_order_changer cert env ff = let coeff = Lazy.force coq_Rcst in let dump_coeff = dump_Rcst in let typ = Lazy.force coq_R in - let cert_typ = (Term.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in + let cert_typ = (EConstr.mkApp(Lazy.force coq_list, [|Lazy.force coq_QWitness |])) in - let formula_typ = (Term.mkApp (Lazy.force coq_Cstr,[| coeff|])) in + let formula_typ = (EConstr.mkApp (Lazy.force coq_Cstr,[| coeff|])) in let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in let vm = dump_varmap (typ) (vm_of_list env) in Proofview.Goal.nf_enter begin fun gl -> @@ -2054,8 +2051,8 @@ let micromega_order_changer cert env ff = (Tactics.change_concl (set [ - ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |])); - ("__varmap", vm, Term.mkApp + ("__ff", ff, EConstr.mkApp(Lazy.force coq_Formula, [|formula_typ |])); + ("__varmap", vm, EConstr.mkApp (gen_constant_in_modules "VarMap" [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|])); ("__wit", cert, cert_typ) @@ -2106,7 +2103,7 @@ let micromega_genr prover tac = let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in let goal_name = fresh_id [] (Names.Id.of_string "__arith") gl in - let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in + let env' = List.map (fun (id,i) -> EConstr.mkVar id,i) vars in let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; micromega_order_changer res' env' ff_arith ] in @@ -2128,8 +2125,8 @@ let micromega_genr prover tac = [ kill_arith; (Tacticals.New.tclTHENLIST - [(Tactics.generalize (List.map Term.mkVar ids)); - Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args)) + [(Tactics.generalize (List.map EConstr.mkVar ids)); + Tactics.exact_check (EConstr.applist (EConstr.mkVar goal_name, arith_args)) ] ) ] diff --git a/plugins/micromega/g_micromega.ml4 b/plugins/micromega/g_micromega.ml4 index ccb6daa11..d803c7554 100644 --- a/plugins/micromega/g_micromega.ml4 +++ b/plugins/micromega/g_micromega.ml4 @@ -16,6 +16,7 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Ltac_plugin open Stdarg open Tacarg diff --git a/plugins/micromega/micromega.ml b/plugins/micromega/micromega.ml new file mode 100644 index 000000000..7da4a3b82 --- /dev/null +++ b/plugins/micromega/micromega.ml @@ -0,0 +1,1773 @@ + +(** val negb : bool -> bool **) + +let negb = function +| true -> false +| false -> true + +type nat = +| O +| S of nat + +(** val app : 'a1 list -> 'a1 list -> 'a1 list **) + +let rec app l m = + match l with + | [] -> m + | a::l1 -> a::(app l1 m) + +type comparison = +| Eq +| Lt +| Gt + +(** val compOpp : comparison -> comparison **) + +let compOpp = function +| Eq -> Eq +| Lt -> Gt +| Gt -> Lt + +module Coq__1 = struct + (** val add : nat -> nat -> nat **) + let rec add n0 m = + match n0 with + | O -> m + | S p -> S (add p m) +end +include Coq__1 + +type positive = +| XI of positive +| XO of positive +| XH + +type n = +| N0 +| Npos of positive + +type z = +| Z0 +| Zpos of positive +| Zneg of positive + +module Pos = + struct + type mask = + | IsNul + | IsPos of positive + | IsNeg + end + +module Coq_Pos = + struct + (** val succ : positive -> positive **) + + let rec succ = function + | XI p -> XO (succ p) + | XO p -> XI p + | XH -> XO XH + + (** val add : positive -> positive -> positive **) + + let rec add x y = + match x with + | XI p -> + (match y with + | XI q0 -> XO (add_carry p q0) + | XO q0 -> XI (add p q0) + | XH -> XO (succ p)) + | XO p -> + (match y with + | XI q0 -> XI (add p q0) + | XO q0 -> XO (add p q0) + | XH -> XI p) + | XH -> (match y with + | XI q0 -> XO (succ q0) + | XO q0 -> XI q0 + | XH -> XO XH) + + (** val add_carry : positive -> positive -> positive **) + + and add_carry x y = + match x with + | XI p -> + (match y with + | XI q0 -> XI (add_carry p q0) + | XO q0 -> XO (add_carry p q0) + | XH -> XI (succ p)) + | XO p -> + (match y with + | XI q0 -> XO (add_carry p q0) + | XO q0 -> XI (add p q0) + | XH -> XO (succ p)) + | XH -> + (match y with + | XI q0 -> XI (succ q0) + | XO q0 -> XO (succ q0) + | XH -> XI XH) + + (** val pred_double : positive -> positive **) + + let rec pred_double = function + | XI p -> XI (XO p) + | XO p -> XI (pred_double p) + | XH -> XH + + type mask = Pos.mask = + | IsNul + | IsPos of positive + | IsNeg + + (** val succ_double_mask : mask -> mask **) + + let succ_double_mask = function + | IsNul -> IsPos XH + | IsPos p -> IsPos (XI p) + | IsNeg -> IsNeg + + (** val double_mask : mask -> mask **) + + let double_mask = function + | IsPos p -> IsPos (XO p) + | x0 -> x0 + + (** val double_pred_mask : positive -> mask **) + + let double_pred_mask = function + | XI p -> IsPos (XO (XO p)) + | XO p -> IsPos (XO (pred_double p)) + | XH -> IsNul + + (** val sub_mask : positive -> positive -> mask **) + + let rec sub_mask x y = + match x with + | XI p -> + (match y with + | XI q0 -> double_mask (sub_mask p q0) + | XO q0 -> succ_double_mask (sub_mask p q0) + | XH -> IsPos (XO p)) + | XO p -> + (match y with + | XI q0 -> succ_double_mask (sub_mask_carry p q0) + | XO q0 -> double_mask (sub_mask p q0) + | XH -> IsPos (pred_double p)) + | XH -> (match y with + | XH -> IsNul + | _ -> IsNeg) + + (** val sub_mask_carry : positive -> positive -> mask **) + + and sub_mask_carry x y = + match x with + | XI p -> + (match y with + | XI q0 -> succ_double_mask (sub_mask_carry p q0) + | XO q0 -> double_mask (sub_mask p q0) + | XH -> IsPos (pred_double p)) + | XO p -> + (match y with + | XI q0 -> double_mask (sub_mask_carry p q0) + | XO q0 -> succ_double_mask (sub_mask_carry p q0) + | XH -> double_pred_mask p) + | XH -> IsNeg + + (** val sub : positive -> positive -> positive **) + + let sub x y = + match sub_mask x y with + | IsPos z0 -> z0 + | _ -> XH + + (** val mul : positive -> positive -> positive **) + + let rec mul x y = + match x with + | XI p -> add y (XO (mul p y)) + | XO p -> XO (mul p y) + | XH -> y + + (** val size_nat : positive -> nat **) + + let rec size_nat = function + | XI p2 -> S (size_nat p2) + | XO p2 -> S (size_nat p2) + | XH -> S O + + (** val compare_cont : comparison -> positive -> positive -> comparison **) + + let rec compare_cont r x y = + match x with + | XI p -> + (match y with + | XI q0 -> compare_cont r p q0 + | XO q0 -> compare_cont Gt p q0 + | XH -> Gt) + | XO p -> + (match y with + | XI q0 -> compare_cont Lt p q0 + | XO q0 -> compare_cont r p q0 + | XH -> Gt) + | XH -> (match y with + | XH -> r + | _ -> Lt) + + (** val compare : positive -> positive -> comparison **) + + let compare = + compare_cont Eq + + (** val gcdn : nat -> positive -> positive -> positive **) + + let rec gcdn n0 a b = + match n0 with + | O -> XH + | S n1 -> + (match a with + | XI a' -> + (match b with + | XI b' -> + (match compare a' b' with + | Eq -> a + | Lt -> gcdn n1 (sub b' a') a + | Gt -> gcdn n1 (sub a' b') b) + | XO b0 -> gcdn n1 a b0 + | XH -> XH) + | XO a0 -> + (match b with + | XI _ -> gcdn n1 a0 b + | XO b0 -> XO (gcdn n1 a0 b0) + | XH -> XH) + | XH -> XH) + + (** val gcd : positive -> positive -> positive **) + + let gcd a b = + gcdn (Coq__1.add (size_nat a) (size_nat b)) a b + + (** val of_succ_nat : nat -> positive **) + + let rec of_succ_nat = function + | O -> XH + | S x -> succ (of_succ_nat x) + end + +module N = + struct + (** val of_nat : nat -> n **) + + let of_nat = function + | O -> N0 + | S n' -> Npos (Coq_Pos.of_succ_nat n') + end + +(** val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1 **) + +let rec pow_pos rmul x = function +| XI i0 -> let p = pow_pos rmul x i0 in rmul x (rmul p p) +| XO i0 -> let p = pow_pos rmul x i0 in rmul p p +| XH -> x + +(** val nth : nat -> 'a1 list -> 'a1 -> 'a1 **) + +let rec nth n0 l default = + match n0 with + | O -> (match l with + | [] -> default + | x::_ -> x) + | S m -> (match l with + | [] -> default + | _::t0 -> nth m t0 default) + +(** val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list **) + +let rec map f = function +| [] -> [] +| a::t0 -> (f a)::(map f t0) + +(** val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1 **) + +let rec fold_right f a0 = function +| [] -> a0 +| b::t0 -> f b (fold_right f a0 t0) + +module Z = + struct + (** val double : z -> z **) + + let double = function + | Z0 -> Z0 + | Zpos p -> Zpos (XO p) + | Zneg p -> Zneg (XO p) + + (** val succ_double : z -> z **) + + let succ_double = function + | Z0 -> Zpos XH + | Zpos p -> Zpos (XI p) + | Zneg p -> Zneg (Coq_Pos.pred_double p) + + (** val pred_double : z -> z **) + + let pred_double = function + | Z0 -> Zneg XH + | Zpos p -> Zpos (Coq_Pos.pred_double p) + | Zneg p -> Zneg (XI p) + + (** val pos_sub : positive -> positive -> z **) + + let rec pos_sub x y = + match x with + | XI p -> + (match y with + | XI q0 -> double (pos_sub p q0) + | XO q0 -> succ_double (pos_sub p q0) + | XH -> Zpos (XO p)) + | XO p -> + (match y with + | XI q0 -> pred_double (pos_sub p q0) + | XO q0 -> double (pos_sub p q0) + | XH -> Zpos (Coq_Pos.pred_double p)) + | XH -> + (match y with + | XI q0 -> Zneg (XO q0) + | XO q0 -> Zneg (Coq_Pos.pred_double q0) + | XH -> Z0) + + (** val add : z -> z -> z **) + + let add x y = + match x with + | Z0 -> y + | Zpos x' -> + (match y with + | Z0 -> x + | Zpos y' -> Zpos (Coq_Pos.add x' y') + | Zneg y' -> pos_sub x' y') + | Zneg x' -> + (match y with + | Z0 -> x + | Zpos y' -> pos_sub y' x' + | Zneg y' -> Zneg (Coq_Pos.add x' y')) + + (** val opp : z -> z **) + + let opp = function + | Z0 -> Z0 + | Zpos x0 -> Zneg x0 + | Zneg x0 -> Zpos x0 + + (** val sub : z -> z -> z **) + + let sub m n0 = + add m (opp n0) + + (** val mul : z -> z -> z **) + + let mul x y = + match x with + | Z0 -> Z0 + | Zpos x' -> + (match y with + | Z0 -> Z0 + | Zpos y' -> Zpos (Coq_Pos.mul x' y') + | Zneg y' -> Zneg (Coq_Pos.mul x' y')) + | Zneg x' -> + (match y with + | Z0 -> Z0 + | Zpos y' -> Zneg (Coq_Pos.mul x' y') + | Zneg y' -> Zpos (Coq_Pos.mul x' y')) + + (** val compare : z -> z -> comparison **) + + let compare x y = + match x with + | Z0 -> (match y with + | Z0 -> Eq + | Zpos _ -> Lt + | Zneg _ -> Gt) + | Zpos x' -> (match y with + | Zpos y' -> Coq_Pos.compare x' y' + | _ -> Gt) + | Zneg x' -> + (match y with + | Zneg y' -> compOpp (Coq_Pos.compare x' y') + | _ -> Lt) + + (** val leb : z -> z -> bool **) + + let leb x y = + match compare x y with + | Gt -> false + | _ -> true + + (** val ltb : z -> z -> bool **) + + let ltb x y = + match compare x y with + | Lt -> true + | _ -> false + + (** val gtb : z -> z -> bool **) + + let gtb x y = + match compare x y with + | Gt -> true + | _ -> false + + (** val max : z -> z -> z **) + + let max n0 m = + match compare n0 m with + | Lt -> m + | _ -> n0 + + (** val abs : z -> z **) + + let abs = function + | Zneg p -> Zpos p + | x -> x + + (** val to_N : z -> n **) + + let to_N = function + | Zpos p -> Npos p + | _ -> N0 + + (** val pos_div_eucl : positive -> z -> z * z **) + + let rec pos_div_eucl a b = + match a with + | XI a' -> + let q0,r = pos_div_eucl a' b in + let r' = add (mul (Zpos (XO XH)) r) (Zpos XH) in + if ltb r' b + then (mul (Zpos (XO XH)) q0),r' + else (add (mul (Zpos (XO XH)) q0) (Zpos XH)),(sub r' b) + | XO a' -> + let q0,r = pos_div_eucl a' b in + let r' = mul (Zpos (XO XH)) r in + if ltb r' b + then (mul (Zpos (XO XH)) q0),r' + else (add (mul (Zpos (XO XH)) q0) (Zpos XH)),(sub r' b) + | XH -> if leb (Zpos (XO XH)) b then Z0,(Zpos XH) else (Zpos XH),Z0 + + (** val div_eucl : z -> z -> z * z **) + + let div_eucl a b = + match a with + | Z0 -> Z0,Z0 + | Zpos a' -> + (match b with + | Z0 -> Z0,Z0 + | Zpos _ -> pos_div_eucl a' b + | Zneg b' -> + let q0,r = pos_div_eucl a' (Zpos b') in + (match r with + | Z0 -> (opp q0),Z0 + | _ -> (opp (add q0 (Zpos XH))),(add b r))) + | Zneg a' -> + (match b with + | Z0 -> Z0,Z0 + | Zpos _ -> + let q0,r = pos_div_eucl a' b in + (match r with + | Z0 -> (opp q0),Z0 + | _ -> (opp (add q0 (Zpos XH))),(sub b r)) + | Zneg b' -> let q0,r = pos_div_eucl a' (Zpos b') in q0,(opp r)) + + (** val div : z -> z -> z **) + + let div a b = + let q0,_ = div_eucl a b in q0 + + (** val gcd : z -> z -> z **) + + let gcd a b = + match a with + | Z0 -> abs b + | Zpos a0 -> + (match b with + | Z0 -> abs a + | Zpos b0 -> Zpos (Coq_Pos.gcd a0 b0) + | Zneg b0 -> Zpos (Coq_Pos.gcd a0 b0)) + | Zneg a0 -> + (match b with + | Z0 -> abs a + | Zpos b0 -> Zpos (Coq_Pos.gcd a0 b0) + | Zneg b0 -> Zpos (Coq_Pos.gcd a0 b0)) + end + +(** val zeq_bool : z -> z -> bool **) + +let zeq_bool x y = + match Z.compare x y with + | Eq -> true + | _ -> false + +type 'c pol = +| Pc of 'c +| Pinj of positive * 'c pol +| PX of 'c pol * positive * 'c pol + +(** val p0 : 'a1 -> 'a1 pol **) + +let p0 cO = + Pc cO + +(** val p1 : 'a1 -> 'a1 pol **) + +let p1 cI = + Pc cI + +(** val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool **) + +let rec peq ceqb p p' = + match p with + | Pc c -> (match p' with + | Pc c' -> ceqb c c' + | _ -> false) + | Pinj (j, q0) -> + (match p' with + | Pinj (j', q') -> + (match Coq_Pos.compare j j' with + | Eq -> peq ceqb q0 q' + | _ -> false) + | _ -> false) + | PX (p2, i, q0) -> + (match p' with + | PX (p'0, i', q') -> + (match Coq_Pos.compare i i' with + | Eq -> if peq ceqb p2 p'0 then peq ceqb q0 q' else false + | _ -> false) + | _ -> false) + +(** val mkPinj : positive -> 'a1 pol -> 'a1 pol **) + +let mkPinj j p = match p with +| Pc _ -> p +| Pinj (j', q0) -> Pinj ((Coq_Pos.add j j'), q0) +| PX (_, _, _) -> Pinj (j, p) + +(** val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol **) + +let mkPinj_pred j p = + match j with + | XI j0 -> Pinj ((XO j0), p) + | XO j0 -> Pinj ((Coq_Pos.pred_double j0), p) + | XH -> p + +(** val mkPX : + 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let mkPX cO ceqb p i q0 = + match p with + | Pc c -> if ceqb c cO then mkPinj XH q0 else PX (p, i, q0) + | Pinj (_, _) -> PX (p, i, q0) + | PX (p', i', q') -> + if peq ceqb q' (p0 cO) + then PX (p', (Coq_Pos.add i' i), q0) + else PX (p, i, q0) + +(** val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol **) + +let mkXi cO cI i = + PX ((p1 cI), i, (p0 cO)) + +(** val mkX : 'a1 -> 'a1 -> 'a1 pol **) + +let mkX cO cI = + mkXi cO cI XH + +(** val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol **) + +let rec popp copp = function +| Pc c -> Pc (copp c) +| Pinj (j, q0) -> Pinj (j, (popp copp q0)) +| PX (p2, i, q0) -> PX ((popp copp p2), i, (popp copp q0)) + +(** val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol **) + +let rec paddC cadd p c = + match p with + | Pc c1 -> Pc (cadd c1 c) + | Pinj (j, q0) -> Pinj (j, (paddC cadd q0 c)) + | PX (p2, i, q0) -> PX (p2, i, (paddC cadd q0 c)) + +(** val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol **) + +let rec psubC csub p c = + match p with + | Pc c1 -> Pc (csub c1 c) + | Pinj (j, q0) -> Pinj (j, (psubC csub q0 c)) + | PX (p2, i, q0) -> PX (p2, i, (psubC csub q0 c)) + +(** val paddI : + ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> + positive -> 'a1 pol -> 'a1 pol **) + +let rec paddI cadd pop q0 j = function +| Pc c -> mkPinj j (paddC cadd q0 c) +| Pinj (j', q') -> + (match Z.pos_sub j' j with + | Z0 -> mkPinj j (pop q' q0) + | Zpos k -> mkPinj j (pop (Pinj (k, q')) q0) + | Zneg k -> mkPinj j' (paddI cadd pop q0 k q')) +| PX (p2, i, q') -> + (match j with + | XI j0 -> PX (p2, i, (paddI cadd pop q0 (XO j0) q')) + | XO j0 -> PX (p2, i, (paddI cadd pop q0 (Coq_Pos.pred_double j0) q')) + | XH -> PX (p2, i, (pop q' q0))) + +(** val psubI : + ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> + 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let rec psubI cadd copp pop q0 j = function +| Pc c -> mkPinj j (paddC cadd (popp copp q0) c) +| Pinj (j', q') -> + (match Z.pos_sub j' j with + | Z0 -> mkPinj j (pop q' q0) + | Zpos k -> mkPinj j (pop (Pinj (k, q')) q0) + | Zneg k -> mkPinj j' (psubI cadd copp pop q0 k q')) +| PX (p2, i, q') -> + (match j with + | XI j0 -> PX (p2, i, (psubI cadd copp pop q0 (XO j0) q')) + | XO j0 -> PX (p2, i, (psubI cadd copp pop q0 (Coq_Pos.pred_double j0) q')) + | XH -> PX (p2, i, (pop q' q0))) + +(** val paddX : + 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol + -> positive -> 'a1 pol -> 'a1 pol **) + +let rec paddX cO ceqb pop p' i' p = match p with +| Pc _ -> PX (p', i', p) +| Pinj (j, q') -> + (match j with + | XI j0 -> PX (p', i', (Pinj ((XO j0), q'))) + | XO j0 -> PX (p', i', (Pinj ((Coq_Pos.pred_double j0), q'))) + | XH -> PX (p', i', q')) +| PX (p2, i, q') -> + (match Z.pos_sub i i' with + | Z0 -> mkPX cO ceqb (pop p2 p') i q' + | Zpos k -> mkPX cO ceqb (pop (PX (p2, k, (p0 cO))) p') i' q' + | Zneg k -> mkPX cO ceqb (paddX cO ceqb pop p' k p2) i q') + +(** val psubX : + 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 + pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let rec psubX cO copp ceqb pop p' i' p = match p with +| Pc _ -> PX ((popp copp p'), i', p) +| Pinj (j, q') -> + (match j with + | XI j0 -> PX ((popp copp p'), i', (Pinj ((XO j0), q'))) + | XO j0 -> PX ((popp copp p'), i', (Pinj ((Coq_Pos.pred_double j0), q'))) + | XH -> PX ((popp copp p'), i', q')) +| PX (p2, i, q') -> + (match Z.pos_sub i i' with + | Z0 -> mkPX cO ceqb (pop p2 p') i q' + | Zpos k -> mkPX cO ceqb (pop (PX (p2, k, (p0 cO))) p') i' q' + | Zneg k -> mkPX cO ceqb (psubX cO copp ceqb pop p' k p2) i q') + +(** val padd : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol + -> 'a1 pol **) + +let rec padd cO cadd ceqb p = function +| Pc c' -> paddC cadd p c' +| Pinj (j', q') -> paddI cadd (padd cO cadd ceqb) q' j' p +| PX (p'0, i', q') -> + (match p with + | Pc c -> PX (p'0, i', (paddC cadd q' c)) + | Pinj (j, q0) -> + (match j with + | XI j0 -> PX (p'0, i', (padd cO cadd ceqb (Pinj ((XO j0), q0)) q')) + | XO j0 -> + PX (p'0, i', + (padd cO cadd ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) q')) + | XH -> PX (p'0, i', (padd cO cadd ceqb q0 q'))) + | PX (p2, i, q0) -> + (match Z.pos_sub i i' with + | Z0 -> + mkPX cO ceqb (padd cO cadd ceqb p2 p'0) i (padd cO cadd ceqb q0 q') + | Zpos k -> + mkPX cO ceqb (padd cO cadd ceqb (PX (p2, k, (p0 cO))) p'0) i' + (padd cO cadd ceqb q0 q') + | Zneg k -> + mkPX cO ceqb (paddX cO ceqb (padd cO cadd ceqb) p'0 k p2) i + (padd cO cadd ceqb q0 q'))) + +(** val psub : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 + -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) + +let rec psub cO cadd csub copp ceqb p = function +| Pc c' -> psubC csub p c' +| Pinj (j', q') -> psubI cadd copp (psub cO cadd csub copp ceqb) q' j' p +| PX (p'0, i', q') -> + (match p with + | Pc c -> PX ((popp copp p'0), i', (paddC cadd (popp copp q') c)) + | Pinj (j, q0) -> + (match j with + | XI j0 -> + PX ((popp copp p'0), i', + (psub cO cadd csub copp ceqb (Pinj ((XO j0), q0)) q')) + | XO j0 -> + PX ((popp copp p'0), i', + (psub cO cadd csub copp ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) + q')) + | XH -> PX ((popp copp p'0), i', (psub cO cadd csub copp ceqb q0 q'))) + | PX (p2, i, q0) -> + (match Z.pos_sub i i' with + | Z0 -> + mkPX cO ceqb (psub cO cadd csub copp ceqb p2 p'0) i + (psub cO cadd csub copp ceqb q0 q') + | Zpos k -> + mkPX cO ceqb (psub cO cadd csub copp ceqb (PX (p2, k, (p0 cO))) p'0) + i' (psub cO cadd csub copp ceqb q0 q') + | Zneg k -> + mkPX cO ceqb + (psubX cO copp ceqb (psub cO cadd csub copp ceqb) p'0 k p2) i + (psub cO cadd csub copp ceqb q0 q'))) + +(** val pmulC_aux : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> + 'a1 pol **) + +let rec pmulC_aux cO cmul ceqb p c = + match p with + | Pc c' -> Pc (cmul c' c) + | Pinj (j, q0) -> mkPinj j (pmulC_aux cO cmul ceqb q0 c) + | PX (p2, i, q0) -> + mkPX cO ceqb (pmulC_aux cO cmul ceqb p2 c) i (pmulC_aux cO cmul ceqb q0 c) + +(** val pmulC : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> + 'a1 -> 'a1 pol **) + +let pmulC cO cI cmul ceqb p c = + if ceqb c cO + then p0 cO + else if ceqb c cI then p else pmulC_aux cO cmul ceqb p c + +(** val pmulI : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> + 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol **) + +let rec pmulI cO cI cmul ceqb pmul0 q0 j = function +| Pc c -> mkPinj j (pmulC cO cI cmul ceqb q0 c) +| Pinj (j', q') -> + (match Z.pos_sub j' j with + | Z0 -> mkPinj j (pmul0 q' q0) + | Zpos k -> mkPinj j (pmul0 (Pinj (k, q')) q0) + | Zneg k -> mkPinj j' (pmulI cO cI cmul ceqb pmul0 q0 k q')) +| PX (p', i', q') -> + (match j with + | XI j' -> + mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 j p') i' + (pmulI cO cI cmul ceqb pmul0 q0 (XO j') q') + | XO j' -> + mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 j p') i' + (pmulI cO cI cmul ceqb pmul0 q0 (Coq_Pos.pred_double j') q') + | XH -> + mkPX cO ceqb (pmulI cO cI cmul ceqb pmul0 q0 XH p') i' (pmul0 q' q0)) + +(** val pmul : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) + +let rec pmul cO cI cadd cmul ceqb p p'' = match p'' with +| Pc c -> pmulC cO cI cmul ceqb p c +| Pinj (j', q') -> pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' j' p +| PX (p', i', q') -> + (match p with + | Pc c -> pmulC cO cI cmul ceqb p'' c + | Pinj (j, q0) -> + let qQ' = + match j with + | XI j0 -> pmul cO cI cadd cmul ceqb (Pinj ((XO j0), q0)) q' + | XO j0 -> + pmul cO cI cadd cmul ceqb (Pinj ((Coq_Pos.pred_double j0), q0)) q' + | XH -> pmul cO cI cadd cmul ceqb q0 q' + in + mkPX cO ceqb (pmul cO cI cadd cmul ceqb p p') i' qQ' + | PX (p2, i, q0) -> + let qQ' = pmul cO cI cadd cmul ceqb q0 q' in + let pQ' = pmulI cO cI cmul ceqb (pmul cO cI cadd cmul ceqb) q' XH p2 in + let qP' = pmul cO cI cadd cmul ceqb (mkPinj XH q0) p' in + let pP' = pmul cO cI cadd cmul ceqb p2 p' in + padd cO cadd ceqb + (mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb pP' i (p0 cO)) qP') i' + (p0 cO)) (mkPX cO ceqb pQ' i qQ')) + +(** val psquare : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> 'a1 pol -> 'a1 pol **) + +let rec psquare cO cI cadd cmul ceqb = function +| Pc c -> Pc (cmul c c) +| Pinj (j, q0) -> Pinj (j, (psquare cO cI cadd cmul ceqb q0)) +| PX (p2, i, q0) -> + let twoPQ = + pmul cO cI cadd cmul ceqb p2 + (mkPinj XH (pmulC cO cI cmul ceqb q0 (cadd cI cI))) + in + let q2 = psquare cO cI cadd cmul ceqb q0 in + let p3 = psquare cO cI cadd cmul ceqb p2 in + mkPX cO ceqb (padd cO cadd ceqb (mkPX cO ceqb p3 i (p0 cO)) twoPQ) i q2 + +type 'c pExpr = +| PEc of 'c +| PEX of positive +| PEadd of 'c pExpr * 'c pExpr +| PEsub of 'c pExpr * 'c pExpr +| PEmul of 'c pExpr * 'c pExpr +| PEopp of 'c pExpr +| PEpow of 'c pExpr * n + +(** val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol **) + +let mk_X cO cI j = + mkPinj_pred j (mkX cO cI) + +(** val ppow_pos : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 + pol **) + +let rec ppow_pos cO cI cadd cmul ceqb subst_l res p = function +| XI p3 -> + subst_l + (pmul cO cI cadd cmul ceqb + (ppow_pos cO cI cadd cmul ceqb subst_l + (ppow_pos cO cI cadd cmul ceqb subst_l res p p3) p p3) p) +| XO p3 -> + ppow_pos cO cI cadd cmul ceqb subst_l + (ppow_pos cO cI cadd cmul ceqb subst_l res p p3) p p3 +| XH -> subst_l (pmul cO cI cadd cmul ceqb res p) + +(** val ppow_N : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol **) + +let ppow_N cO cI cadd cmul ceqb subst_l p = function +| N0 -> p1 cI +| Npos p2 -> ppow_pos cO cI cadd cmul ceqb subst_l (p1 cI) p p2 + +(** val norm_aux : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **) + +let rec norm_aux cO cI cadd cmul csub copp ceqb = function +| PEc c -> Pc c +| PEX j -> mk_X cO cI j +| PEadd (pe1, pe2) -> + (match pe1 with + | PEopp pe3 -> + psub cO cadd csub copp ceqb + (norm_aux cO cI cadd cmul csub copp ceqb pe2) + (norm_aux cO cI cadd cmul csub copp ceqb pe3) + | _ -> + (match pe2 with + | PEopp pe3 -> + psub cO cadd csub copp ceqb + (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe3) + | _ -> + padd cO cadd ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe2))) +| PEsub (pe1, pe2) -> + psub cO cadd csub copp ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe2) +| PEmul (pe1, pe2) -> + pmul cO cI cadd cmul ceqb (norm_aux cO cI cadd cmul csub copp ceqb pe1) + (norm_aux cO cI cadd cmul csub copp ceqb pe2) +| PEopp pe1 -> popp copp (norm_aux cO cI cadd cmul csub copp ceqb pe1) +| PEpow (pe1, n0) -> + ppow_N cO cI cadd cmul ceqb (fun p -> p) + (norm_aux cO cI cadd cmul csub copp ceqb pe1) n0 + +type 'a bFormula = +| TT +| FF +| X +| A of 'a +| Cj of 'a bFormula * 'a bFormula +| D of 'a bFormula * 'a bFormula +| N of 'a bFormula +| I of 'a bFormula * 'a bFormula + +(** val map_bformula : ('a1 -> 'a2) -> 'a1 bFormula -> 'a2 bFormula **) + +let rec map_bformula fct = function +| TT -> TT +| FF -> FF +| X -> X +| A a -> A (fct a) +| Cj (f1, f2) -> Cj ((map_bformula fct f1), (map_bformula fct f2)) +| D (f1, f2) -> D ((map_bformula fct f1), (map_bformula fct f2)) +| N f0 -> N (map_bformula fct f0) +| I (f1, f2) -> I ((map_bformula fct f1), (map_bformula fct f2)) + +type 'x clause = 'x list + +type 'x cnf = 'x clause list + +(** val tt : 'a1 cnf **) + +let tt = + [] + +(** val ff : 'a1 cnf **) + +let ff = + []::[] + +(** val add_term : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 -> 'a1 clause -> 'a1 + clause option **) + +let rec add_term unsat deduce t0 = function +| [] -> + (match deduce t0 t0 with + | Some u -> if unsat u then None else Some (t0::[]) + | None -> Some (t0::[])) +| t'::cl0 -> + (match deduce t0 t' with + | Some u -> + if unsat u + then None + else (match add_term unsat deduce t0 cl0 with + | Some cl' -> Some (t'::cl') + | None -> None) + | None -> + (match add_term unsat deduce t0 cl0 with + | Some cl' -> Some (t'::cl') + | None -> None)) + +(** val or_clause : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause + -> 'a1 clause option **) + +let rec or_clause unsat deduce cl1 cl2 = + match cl1 with + | [] -> Some cl2 + | t0::cl -> + (match add_term unsat deduce t0 cl2 with + | Some cl' -> or_clause unsat deduce cl cl' + | None -> None) + +(** val or_clause_cnf : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf -> + 'a1 cnf **) + +let or_clause_cnf unsat deduce t0 f = + fold_right (fun e acc -> + match or_clause unsat deduce t0 e with + | Some cl -> cl::acc + | None -> acc) [] f + +(** val or_cnf : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1 + cnf **) + +let rec or_cnf unsat deduce f f' = + match f with + | [] -> tt + | e::rst -> + app (or_cnf unsat deduce rst f') (or_clause_cnf unsat deduce e f') + +(** val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf **) + +let and_cnf f1 f2 = + app f1 f2 + +(** val xcnf : + ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 + -> 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf **) + +let rec xcnf unsat deduce normalise0 negate0 pol0 = function +| TT -> if pol0 then tt else ff +| FF -> if pol0 then ff else tt +| X -> ff +| A x -> if pol0 then normalise0 x else negate0 x +| Cj (e1, e2) -> + if pol0 + then and_cnf (xcnf unsat deduce normalise0 negate0 pol0 e1) + (xcnf unsat deduce normalise0 negate0 pol0 e2) + else or_cnf unsat deduce (xcnf unsat deduce normalise0 negate0 pol0 e1) + (xcnf unsat deduce normalise0 negate0 pol0 e2) +| D (e1, e2) -> + if pol0 + then or_cnf unsat deduce (xcnf unsat deduce normalise0 negate0 pol0 e1) + (xcnf unsat deduce normalise0 negate0 pol0 e2) + else and_cnf (xcnf unsat deduce normalise0 negate0 pol0 e1) + (xcnf unsat deduce normalise0 negate0 pol0 e2) +| N e -> xcnf unsat deduce normalise0 negate0 (negb pol0) e +| I (e1, e2) -> + if pol0 + then or_cnf unsat deduce + (xcnf unsat deduce normalise0 negate0 (negb pol0) e1) + (xcnf unsat deduce normalise0 negate0 pol0 e2) + else and_cnf (xcnf unsat deduce normalise0 negate0 (negb pol0) e1) + (xcnf unsat deduce normalise0 negate0 pol0 e2) + +(** val cnf_checker : + ('a1 list -> 'a2 -> bool) -> 'a1 cnf -> 'a2 list -> bool **) + +let rec cnf_checker checker f l = + match f with + | [] -> true + | e::f0 -> + (match l with + | [] -> false + | c::l0 -> if checker e c then cnf_checker checker f0 l0 else false) + +(** val tauto_checker : + ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 + -> 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list -> + bool **) + +let tauto_checker unsat deduce normalise0 negate0 checker f w = + cnf_checker checker (xcnf unsat deduce normalise0 negate0 true f) w + +(** val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool **) + +let cneqb ceqb x y = + negb (ceqb x y) + +(** val cltb : + ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool **) + +let cltb ceqb cleb x y = + (&&) (cleb x y) (cneqb ceqb x y) + +type 'c polC = 'c pol + +type op1 = +| Equal +| NonEqual +| Strict +| NonStrict + +type 'c nFormula = 'c polC * op1 + +(** val opMult : op1 -> op1 -> op1 option **) + +let opMult o o' = + match o with + | Equal -> Some Equal + | NonEqual -> + (match o' with + | Equal -> Some Equal + | NonEqual -> Some NonEqual + | _ -> None) + | Strict -> (match o' with + | NonEqual -> None + | _ -> Some o') + | NonStrict -> + (match o' with + | Equal -> Some Equal + | NonEqual -> None + | _ -> Some NonStrict) + +(** val opAdd : op1 -> op1 -> op1 option **) + +let opAdd o o' = + match o with + | Equal -> Some o' + | NonEqual -> (match o' with + | Equal -> Some NonEqual + | _ -> None) + | Strict -> (match o' with + | NonEqual -> None + | _ -> Some Strict) + | NonStrict -> + (match o' with + | Equal -> Some NonStrict + | NonEqual -> None + | x -> Some x) + +type 'c psatz = +| PsatzIn of nat +| PsatzSquare of 'c polC +| PsatzMulC of 'c polC * 'c psatz +| PsatzMulE of 'c psatz * 'c psatz +| PsatzAdd of 'c psatz * 'c psatz +| PsatzC of 'c +| PsatzZ + +(** val map_option : ('a1 -> 'a2 option) -> 'a1 option -> 'a2 option **) + +let map_option f = function +| Some x -> f x +| None -> None + +(** val map_option2 : + ('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option **) + +let map_option2 f o o' = + match o with + | Some x -> (match o' with + | Some x' -> f x x' + | None -> None) + | None -> None + +(** val pexpr_times_nformula : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option **) + +let pexpr_times_nformula cO cI cplus ctimes ceqb e = function +| ef,o -> + (match o with + | Equal -> Some ((pmul cO cI cplus ctimes ceqb e ef),Equal) + | _ -> None) + +(** val nformula_times_nformula : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option **) + +let nformula_times_nformula cO cI cplus ctimes ceqb f1 f2 = + let e1,o1 = f1 in + let e2,o2 = f2 in + map_option (fun x -> Some ((pmul cO cI cplus ctimes ceqb e1 e2),x)) + (opMult o1 o2) + +(** val nformula_plus_nformula : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1 + nFormula -> 'a1 nFormula option **) + +let nformula_plus_nformula cO cplus ceqb f1 f2 = + let e1,o1 = f1 in + let e2,o2 = f2 in + map_option (fun x -> Some ((padd cO cplus ceqb e1 e2),x)) (opAdd o1 o2) + +(** val eval_Psatz : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1 + nFormula option **) + +let rec eval_Psatz cO cI cplus ctimes ceqb cleb l = function +| PsatzIn n0 -> Some (nth n0 l ((Pc cO),Equal)) +| PsatzSquare e0 -> Some ((psquare cO cI cplus ctimes ceqb e0),NonStrict) +| PsatzMulC (re, e0) -> + map_option (pexpr_times_nformula cO cI cplus ctimes ceqb re) + (eval_Psatz cO cI cplus ctimes ceqb cleb l e0) +| PsatzMulE (f1, f2) -> + map_option2 (nformula_times_nformula cO cI cplus ctimes ceqb) + (eval_Psatz cO cI cplus ctimes ceqb cleb l f1) + (eval_Psatz cO cI cplus ctimes ceqb cleb l f2) +| PsatzAdd (f1, f2) -> + map_option2 (nformula_plus_nformula cO cplus ceqb) + (eval_Psatz cO cI cplus ctimes ceqb cleb l f1) + (eval_Psatz cO cI cplus ctimes ceqb cleb l f2) +| PsatzC c -> if cltb ceqb cleb cO c then Some ((Pc c),Strict) else None +| PsatzZ -> Some ((Pc cO),Equal) + +(** val check_inconsistent : + 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> + bool **) + +let check_inconsistent cO ceqb cleb = function +| e,op -> + (match e with + | Pc c -> + (match op with + | Equal -> cneqb ceqb c cO + | NonEqual -> ceqb c cO + | Strict -> cleb c cO + | NonStrict -> cltb ceqb cleb c cO) + | _ -> false) + +(** val check_normalised_formulas : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool **) + +let check_normalised_formulas cO cI cplus ctimes ceqb cleb l cm = + match eval_Psatz cO cI cplus ctimes ceqb cleb l cm with + | Some f -> check_inconsistent cO ceqb cleb f + | None -> false + +type op2 = +| OpEq +| OpNEq +| OpLe +| OpGe +| OpLt +| OpGt + +type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr } + +(** val norm : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol **) + +let norm cO cI cplus ctimes cminus copp ceqb = + norm_aux cO cI cplus ctimes cminus copp ceqb + +(** val psub0 : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 + -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol **) + +let psub0 cO cplus cminus copp ceqb = + psub cO cplus cminus copp ceqb + +(** val padd0 : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol + -> 'a1 pol **) + +let padd0 cO cplus ceqb = + padd cO cplus ceqb + +(** val xnormalise : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 + nFormula list **) + +let xnormalise cO cI cplus ctimes cminus copp ceqb t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + let lhs0 = norm cO cI cplus ctimes cminus copp ceqb lhs in + let rhs0 = norm cO cI cplus ctimes cminus copp ceqb rhs in + (match o with + | OpEq -> + ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO cplus + cminus copp + ceqb rhs0 lhs0),Strict)::[]) + | OpNEq -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Equal)::[] + | OpLe -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::[] + | OpGe -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),Strict)::[] + | OpLt -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),NonStrict)::[] + | OpGt -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),NonStrict)::[]) + +(** val cnf_normalise : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 + nFormula cnf **) + +let cnf_normalise cO cI cplus ctimes cminus copp ceqb t0 = + map (fun x -> x::[]) (xnormalise cO cI cplus ctimes cminus copp ceqb t0) + +(** val xnegate : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 + nFormula list **) + +let xnegate cO cI cplus ctimes cminus copp ceqb t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + let lhs0 = norm cO cI cplus ctimes cminus copp ceqb lhs in + let rhs0 = norm cO cI cplus ctimes cminus copp ceqb rhs in + (match o with + | OpEq -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Equal)::[] + | OpNEq -> + ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::(((psub0 cO cplus + cminus copp + ceqb rhs0 lhs0),Strict)::[]) + | OpLe -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),NonStrict)::[] + | OpGe -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),NonStrict)::[] + | OpLt -> ((psub0 cO cplus cminus copp ceqb rhs0 lhs0),Strict)::[] + | OpGt -> ((psub0 cO cplus cminus copp ceqb lhs0 rhs0),Strict)::[]) + +(** val cnf_negate : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 + -> 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 + nFormula cnf **) + +let cnf_negate cO cI cplus ctimes cminus copp ceqb t0 = + map (fun x -> x::[]) (xnegate cO cI cplus ctimes cminus copp ceqb t0) + +(** val xdenorm : positive -> 'a1 pol -> 'a1 pExpr **) + +let rec xdenorm jmp = function +| Pc c -> PEc c +| Pinj (j, p2) -> xdenorm (Coq_Pos.add j jmp) p2 +| PX (p2, j, q0) -> + PEadd ((PEmul ((xdenorm jmp p2), (PEpow ((PEX jmp), (Npos j))))), + (xdenorm (Coq_Pos.succ jmp) q0)) + +(** val denorm : 'a1 pol -> 'a1 pExpr **) + +let denorm p = + xdenorm XH p + +(** val map_PExpr : ('a2 -> 'a1) -> 'a2 pExpr -> 'a1 pExpr **) + +let rec map_PExpr c_of_S = function +| PEc c -> PEc (c_of_S c) +| PEX p -> PEX p +| PEadd (e1, e2) -> PEadd ((map_PExpr c_of_S e1), (map_PExpr c_of_S e2)) +| PEsub (e1, e2) -> PEsub ((map_PExpr c_of_S e1), (map_PExpr c_of_S e2)) +| PEmul (e1, e2) -> PEmul ((map_PExpr c_of_S e1), (map_PExpr c_of_S e2)) +| PEopp e0 -> PEopp (map_PExpr c_of_S e0) +| PEpow (e0, n0) -> PEpow ((map_PExpr c_of_S e0), n0) + +(** val map_Formula : ('a2 -> 'a1) -> 'a2 formula -> 'a1 formula **) + +let map_Formula c_of_S f = + let { flhs = l; fop = o; frhs = r } = f in + { flhs = (map_PExpr c_of_S l); fop = o; frhs = (map_PExpr c_of_S r) } + +(** val simpl_cone : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz -> + 'a1 psatz **) + +let simpl_cone cO cI ctimes ceqb e = match e with +| PsatzSquare t0 -> + (match t0 with + | Pc c -> if ceqb cO c then PsatzZ else PsatzC (ctimes c c) + | _ -> PsatzSquare t0) +| PsatzMulE (t1, t2) -> + (match t1 with + | PsatzMulE (x, x0) -> + (match x with + | PsatzC p2 -> + (match t2 with + | PsatzC c -> PsatzMulE ((PsatzC (ctimes c p2)), x0) + | PsatzZ -> PsatzZ + | _ -> e) + | _ -> + (match x0 with + | PsatzC p2 -> + (match t2 with + | PsatzC c -> PsatzMulE ((PsatzC (ctimes c p2)), x) + | PsatzZ -> PsatzZ + | _ -> e) + | _ -> + (match t2 with + | PsatzC c -> if ceqb cI c then t1 else PsatzMulE (t1, t2) + | PsatzZ -> PsatzZ + | _ -> e))) + | PsatzC c -> + (match t2 with + | PsatzMulE (x, x0) -> + (match x with + | PsatzC p2 -> PsatzMulE ((PsatzC (ctimes c p2)), x0) + | _ -> + (match x0 with + | PsatzC p2 -> PsatzMulE ((PsatzC (ctimes c p2)), x) + | _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2))) + | PsatzAdd (y, z0) -> + PsatzAdd ((PsatzMulE ((PsatzC c), y)), (PsatzMulE ((PsatzC c), z0))) + | PsatzC c0 -> PsatzC (ctimes c c0) + | PsatzZ -> PsatzZ + | _ -> if ceqb cI c then t2 else PsatzMulE (t1, t2)) + | PsatzZ -> PsatzZ + | _ -> + (match t2 with + | PsatzC c -> if ceqb cI c then t1 else PsatzMulE (t1, t2) + | PsatzZ -> PsatzZ + | _ -> e)) +| PsatzAdd (t1, t2) -> + (match t1 with + | PsatzZ -> t2 + | _ -> (match t2 with + | PsatzZ -> t1 + | _ -> PsatzAdd (t1, t2))) +| _ -> e + +type q = { qnum : z; qden : positive } + +(** val qnum : q -> z **) + +let qnum x = x.qnum + +(** val qden : q -> positive **) + +let qden x = x.qden + +(** val qeq_bool : q -> q -> bool **) + +let qeq_bool x y = + zeq_bool (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)) + +(** val qle_bool : q -> q -> bool **) + +let qle_bool x y = + Z.leb (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden)) + +(** val qplus : q -> q -> q **) + +let qplus x y = + { qnum = (Z.add (Z.mul x.qnum (Zpos y.qden)) (Z.mul y.qnum (Zpos x.qden))); + qden = (Coq_Pos.mul x.qden y.qden) } + +(** val qmult : q -> q -> q **) + +let qmult x y = + { qnum = (Z.mul x.qnum y.qnum); qden = (Coq_Pos.mul x.qden y.qden) } + +(** val qopp : q -> q **) + +let qopp x = + { qnum = (Z.opp x.qnum); qden = x.qden } + +(** val qminus : q -> q -> q **) + +let qminus x y = + qplus x (qopp y) + +(** val qinv : q -> q **) + +let qinv x = + match x.qnum with + | Z0 -> { qnum = Z0; qden = XH } + | Zpos p -> { qnum = (Zpos x.qden); qden = p } + | Zneg p -> { qnum = (Zneg x.qden); qden = p } + +(** val qpower_positive : q -> positive -> q **) + +let qpower_positive = + pow_pos qmult + +(** val qpower : q -> z -> q **) + +let qpower q0 = function +| Z0 -> { qnum = (Zpos XH); qden = XH } +| Zpos p -> qpower_positive q0 p +| Zneg p -> qinv (qpower_positive q0 p) + +type 'a t = +| Empty +| Leaf of 'a +| Node of 'a t * 'a * 'a t + +(** val find : 'a1 -> 'a1 t -> positive -> 'a1 **) + +let rec find default vm p = + match vm with + | Empty -> default + | Leaf i -> i + | Node (l, e, r) -> + (match p with + | XI p2 -> find default r p2 + | XO p2 -> find default l p2 + | XH -> e) + +(** val singleton : 'a1 -> positive -> 'a1 -> 'a1 t **) + +let rec singleton default x v = + match x with + | XI p -> Node (Empty, default, (singleton default p v)) + | XO p -> Node ((singleton default p v), default, Empty) + | XH -> Leaf v + +(** val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t **) + +let rec vm_add default x v = function +| Empty -> singleton default x v +| Leaf vl -> + (match x with + | XI p -> Node (Empty, vl, (singleton default p v)) + | XO p -> Node ((singleton default p v), vl, Empty) + | XH -> Leaf v) +| Node (l, o, r) -> + (match x with + | XI p -> Node (l, o, (vm_add default p v r)) + | XO p -> Node ((vm_add default p v l), o, r) + | XH -> Node (l, v, r)) + +type zWitness = z psatz + +(** val zWeakChecker : z nFormula list -> z psatz -> bool **) + +let zWeakChecker = + check_normalised_formulas Z0 (Zpos XH) Z.add Z.mul zeq_bool Z.leb + +(** val psub1 : z pol -> z pol -> z pol **) + +let psub1 = + psub0 Z0 Z.add Z.sub Z.opp zeq_bool + +(** val padd1 : z pol -> z pol -> z pol **) + +let padd1 = + padd0 Z0 Z.add zeq_bool + +(** val norm0 : z pExpr -> z pol **) + +let norm0 = + norm Z0 (Zpos XH) Z.add Z.mul Z.sub Z.opp zeq_bool + +(** val xnormalise0 : z formula -> z nFormula list **) + +let xnormalise0 t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + let lhs0 = norm0 lhs in + let rhs0 = norm0 rhs in + (match o with + | OpEq -> + ((psub1 lhs0 (padd1 rhs0 (Pc (Zpos XH)))),NonStrict)::(((psub1 rhs0 + (padd1 lhs0 + (Pc (Zpos + XH)))),NonStrict)::[]) + | OpNEq -> ((psub1 lhs0 rhs0),Equal)::[] + | OpLe -> ((psub1 lhs0 (padd1 rhs0 (Pc (Zpos XH)))),NonStrict)::[] + | OpGe -> ((psub1 rhs0 (padd1 lhs0 (Pc (Zpos XH)))),NonStrict)::[] + | OpLt -> ((psub1 lhs0 rhs0),NonStrict)::[] + | OpGt -> ((psub1 rhs0 lhs0),NonStrict)::[]) + +(** val normalise : z formula -> z nFormula cnf **) + +let normalise t0 = + map (fun x -> x::[]) (xnormalise0 t0) + +(** val xnegate0 : z formula -> z nFormula list **) + +let xnegate0 t0 = + let { flhs = lhs; fop = o; frhs = rhs } = t0 in + let lhs0 = norm0 lhs in + let rhs0 = norm0 rhs in + (match o with + | OpEq -> ((psub1 lhs0 rhs0),Equal)::[] + | OpNEq -> + ((psub1 lhs0 (padd1 rhs0 (Pc (Zpos XH)))),NonStrict)::(((psub1 rhs0 + (padd1 lhs0 + (Pc (Zpos + XH)))),NonStrict)::[]) + | OpLe -> ((psub1 rhs0 lhs0),NonStrict)::[] + | OpGe -> ((psub1 lhs0 rhs0),NonStrict)::[] + | OpLt -> ((psub1 rhs0 (padd1 lhs0 (Pc (Zpos XH)))),NonStrict)::[] + | OpGt -> ((psub1 lhs0 (padd1 rhs0 (Pc (Zpos XH)))),NonStrict)::[]) + +(** val negate : z formula -> z nFormula cnf **) + +let negate t0 = + map (fun x -> x::[]) (xnegate0 t0) + +(** val zunsat : z nFormula -> bool **) + +let zunsat = + check_inconsistent Z0 zeq_bool Z.leb + +(** val zdeduce : z nFormula -> z nFormula -> z nFormula option **) + +let zdeduce = + nformula_plus_nformula Z0 Z.add zeq_bool + +(** val ceiling : z -> z -> z **) + +let ceiling a b = + let q0,r = Z.div_eucl a b in + (match r with + | Z0 -> q0 + | _ -> Z.add q0 (Zpos XH)) + +type zArithProof = +| DoneProof +| RatProof of zWitness * zArithProof +| CutProof of zWitness * zArithProof +| EnumProof of zWitness * zWitness * zArithProof list + +(** val zgcdM : z -> z -> z **) + +let zgcdM x y = + Z.max (Z.gcd x y) (Zpos XH) + +(** val zgcd_pol : z polC -> z * z **) + +let rec zgcd_pol = function +| Pc c -> Z0,c +| Pinj (_, p2) -> zgcd_pol p2 +| PX (p2, _, q0) -> + let g1,c1 = zgcd_pol p2 in + let g2,c2 = zgcd_pol q0 in (zgcdM (zgcdM g1 c1) g2),c2 + +(** val zdiv_pol : z polC -> z -> z polC **) + +let rec zdiv_pol p x = + match p with + | Pc c -> Pc (Z.div c x) + | Pinj (j, p2) -> Pinj (j, (zdiv_pol p2 x)) + | PX (p2, j, q0) -> PX ((zdiv_pol p2 x), j, (zdiv_pol q0 x)) + +(** val makeCuttingPlane : z polC -> z polC * z **) + +let makeCuttingPlane p = + let g,c = zgcd_pol p in + if Z.gtb g Z0 + then (zdiv_pol (psubC Z.sub p c) g),(Z.opp (ceiling (Z.opp c) g)) + else p,Z0 + +(** val genCuttingPlane : z nFormula -> ((z polC * z) * op1) option **) + +let genCuttingPlane = function +| e,op -> + (match op with + | Equal -> + let g,c = zgcd_pol e in + if (&&) (Z.gtb g Z0) + ((&&) (negb (zeq_bool c Z0)) (negb (zeq_bool (Z.gcd g c) g))) + then None + else Some ((makeCuttingPlane e),Equal) + | NonEqual -> Some ((e,Z0),op) + | Strict -> Some ((makeCuttingPlane (psubC Z.sub e (Zpos XH))),NonStrict) + | NonStrict -> Some ((makeCuttingPlane e),NonStrict)) + +(** val nformula_of_cutting_plane : ((z polC * z) * op1) -> z nFormula **) + +let nformula_of_cutting_plane = function +| e_z,o -> let e,z0 = e_z in (padd1 e (Pc z0)),o + +(** val is_pol_Z0 : z polC -> bool **) + +let is_pol_Z0 = function +| Pc z0 -> (match z0 with + | Z0 -> true + | _ -> false) +| _ -> false + +(** val eval_Psatz0 : z nFormula list -> zWitness -> z nFormula option **) + +let eval_Psatz0 = + eval_Psatz Z0 (Zpos XH) Z.add Z.mul zeq_bool Z.leb + +(** val valid_cut_sign : op1 -> bool **) + +let valid_cut_sign = function +| Equal -> true +| NonStrict -> true +| _ -> false + +(** val zChecker : z nFormula list -> zArithProof -> bool **) + +let rec zChecker l = function +| DoneProof -> false +| RatProof (w, pf0) -> + (match eval_Psatz0 l w with + | Some f -> if zunsat f then true else zChecker (f::l) pf0 + | None -> false) +| CutProof (w, pf0) -> + (match eval_Psatz0 l w with + | Some f -> + (match genCuttingPlane f with + | Some cp -> zChecker ((nformula_of_cutting_plane cp)::l) pf0 + | None -> true) + | None -> false) +| EnumProof (w1, w2, pf0) -> + (match eval_Psatz0 l w1 with + | Some f1 -> + (match eval_Psatz0 l w2 with + | Some f2 -> + (match genCuttingPlane f1 with + | Some p -> + let p2,op3 = p in + let e1,z1 = p2 in + (match genCuttingPlane f2 with + | Some p3 -> + let p4,op4 = p3 in + let e2,z2 = p4 in + if (&&) ((&&) (valid_cut_sign op3) (valid_cut_sign op4)) + (is_pol_Z0 (padd1 e1 e2)) + then let rec label pfs lb ub = + match pfs with + | [] -> Z.gtb lb ub + | pf1::rsr -> + (&&) (zChecker (((psub1 e1 (Pc lb)),Equal)::l) pf1) + (label rsr (Z.add lb (Zpos XH)) ub) + in label pf0 (Z.opp z1) z2 + else false + | None -> true) + | None -> true) + | None -> false) + | None -> false) + +(** val zTautoChecker : z formula bFormula -> zArithProof list -> bool **) + +let zTautoChecker f w = + tauto_checker zunsat zdeduce normalise negate zChecker f w + +type qWitness = q psatz + +(** val qWeakChecker : q nFormula list -> q psatz -> bool **) + +let qWeakChecker = + check_normalised_formulas { qnum = Z0; qden = XH } { qnum = (Zpos XH); + qden = XH } qplus qmult qeq_bool qle_bool + +(** val qnormalise : q formula -> q nFormula cnf **) + +let qnormalise = + cnf_normalise { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } + qplus qmult qminus qopp qeq_bool + +(** val qnegate : q formula -> q nFormula cnf **) + +let qnegate = + cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus + qmult qminus qopp qeq_bool + +(** val qunsat : q nFormula -> bool **) + +let qunsat = + check_inconsistent { qnum = Z0; qden = XH } qeq_bool qle_bool + +(** val qdeduce : q nFormula -> q nFormula -> q nFormula option **) + +let qdeduce = + nformula_plus_nformula { qnum = Z0; qden = XH } qplus qeq_bool + +(** val qTautoChecker : q formula bFormula -> qWitness list -> bool **) + +let qTautoChecker f w = + tauto_checker qunsat qdeduce qnormalise qnegate qWeakChecker f w + +type rcst = +| C0 +| C1 +| CQ of q +| CZ of z +| CPlus of rcst * rcst +| CMinus of rcst * rcst +| CMult of rcst * rcst +| CInv of rcst +| COpp of rcst + +(** val q_of_Rcst : rcst -> q **) + +let rec q_of_Rcst = function +| C0 -> { qnum = Z0; qden = XH } +| C1 -> { qnum = (Zpos XH); qden = XH } +| CQ q0 -> q0 +| CZ z0 -> { qnum = z0; qden = XH } +| CPlus (r1, r2) -> qplus (q_of_Rcst r1) (q_of_Rcst r2) +| CMinus (r1, r2) -> qminus (q_of_Rcst r1) (q_of_Rcst r2) +| CMult (r1, r2) -> qmult (q_of_Rcst r1) (q_of_Rcst r2) +| CInv r0 -> qinv (q_of_Rcst r0) +| COpp r0 -> qopp (q_of_Rcst r0) + +type rWitness = q psatz + +(** val rWeakChecker : q nFormula list -> q psatz -> bool **) + +let rWeakChecker = + check_normalised_formulas { qnum = Z0; qden = XH } { qnum = (Zpos XH); + qden = XH } qplus qmult qeq_bool qle_bool + +(** val rnormalise : q formula -> q nFormula cnf **) + +let rnormalise = + cnf_normalise { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } + qplus qmult qminus qopp qeq_bool + +(** val rnegate : q formula -> q nFormula cnf **) + +let rnegate = + cnf_negate { qnum = Z0; qden = XH } { qnum = (Zpos XH); qden = XH } qplus + qmult qminus qopp qeq_bool + +(** val runsat : q nFormula -> bool **) + +let runsat = + check_inconsistent { qnum = Z0; qden = XH } qeq_bool qle_bool + +(** val rdeduce : q nFormula -> q nFormula -> q nFormula option **) + +let rdeduce = + nformula_plus_nformula { qnum = Z0; qden = XH } qplus qeq_bool + +(** val rTautoChecker : rcst formula bFormula -> rWitness list -> bool **) + +let rTautoChecker f w = + tauto_checker runsat rdeduce rnormalise rnegate rWeakChecker + (map_bformula (map_Formula q_of_Rcst) f) w diff --git a/plugins/micromega/micromega.mli b/plugins/micromega/micromega.mli new file mode 100644 index 000000000..961978178 --- /dev/null +++ b/plugins/micromega/micromega.mli @@ -0,0 +1,517 @@ + +val negb : bool -> bool + +type nat = +| O +| S of nat + +val app : 'a1 list -> 'a1 list -> 'a1 list + +type comparison = +| Eq +| Lt +| Gt + +val compOpp : comparison -> comparison + +val add : nat -> nat -> nat + +type positive = +| XI of positive +| XO of positive +| XH + +type n = +| N0 +| Npos of positive + +type z = +| Z0 +| Zpos of positive +| Zneg of positive + +module Pos : + sig + type mask = + | IsNul + | IsPos of positive + | IsNeg + end + +module Coq_Pos : + sig + val succ : positive -> positive + + val add : positive -> positive -> positive + + val add_carry : positive -> positive -> positive + + val pred_double : positive -> positive + + type mask = Pos.mask = + | IsNul + | IsPos of positive + | IsNeg + + val succ_double_mask : mask -> mask + + val double_mask : mask -> mask + + val double_pred_mask : positive -> mask + + val sub_mask : positive -> positive -> mask + + val sub_mask_carry : positive -> positive -> mask + + val sub : positive -> positive -> positive + + val mul : positive -> positive -> positive + + val size_nat : positive -> nat + + val compare_cont : comparison -> positive -> positive -> comparison + + val compare : positive -> positive -> comparison + + val gcdn : nat -> positive -> positive -> positive + + val gcd : positive -> positive -> positive + + val of_succ_nat : nat -> positive + end + +module N : + sig + val of_nat : nat -> n + end + +val pow_pos : ('a1 -> 'a1 -> 'a1) -> 'a1 -> positive -> 'a1 + +val nth : nat -> 'a1 list -> 'a1 -> 'a1 + +val map : ('a1 -> 'a2) -> 'a1 list -> 'a2 list + +val fold_right : ('a2 -> 'a1 -> 'a1) -> 'a1 -> 'a2 list -> 'a1 + +module Z : + sig + val double : z -> z + + val succ_double : z -> z + + val pred_double : z -> z + + val pos_sub : positive -> positive -> z + + val add : z -> z -> z + + val opp : z -> z + + val sub : z -> z -> z + + val mul : z -> z -> z + + val compare : z -> z -> comparison + + val leb : z -> z -> bool + + val ltb : z -> z -> bool + + val gtb : z -> z -> bool + + val max : z -> z -> z + + val abs : z -> z + + val to_N : z -> n + + val pos_div_eucl : positive -> z -> z * z + + val div_eucl : z -> z -> z * z + + val div : z -> z -> z + + val gcd : z -> z -> z + end + +val zeq_bool : z -> z -> bool + +type 'c pol = +| Pc of 'c +| Pinj of positive * 'c pol +| PX of 'c pol * positive * 'c pol + +val p0 : 'a1 -> 'a1 pol + +val p1 : 'a1 -> 'a1 pol + +val peq : ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> bool + +val mkPinj : positive -> 'a1 pol -> 'a1 pol + +val mkPinj_pred : positive -> 'a1 pol -> 'a1 pol + +val mkPX : + 'a1 -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val mkXi : 'a1 -> 'a1 -> positive -> 'a1 pol + +val mkX : 'a1 -> 'a1 -> 'a1 pol + +val popp : ('a1 -> 'a1) -> 'a1 pol -> 'a1 pol + +val paddC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol + +val psubC : ('a1 -> 'a1 -> 'a1) -> 'a1 pol -> 'a1 -> 'a1 pol + +val paddI : + ('a1 -> 'a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol -> + positive -> 'a1 pol -> 'a1 pol + +val psubI : + ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> + 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val paddX : + 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 pol) -> 'a1 pol + -> positive -> 'a1 pol -> 'a1 pol + +val psubX : + 'a1 -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> 'a1 pol -> 'a1 + pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val padd : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> + 'a1 pol + +val psub : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 + -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol + +val pmulC_aux : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 -> 'a1 + pol + +val pmulC : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 + -> 'a1 pol + +val pmulI : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> ('a1 pol -> + 'a1 pol -> 'a1 pol) -> 'a1 pol -> positive -> 'a1 pol -> 'a1 pol + +val pmul : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol + +val psquare : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> 'a1 pol -> 'a1 pol + +type 'c pExpr = +| PEc of 'c +| PEX of positive +| PEadd of 'c pExpr * 'c pExpr +| PEsub of 'c pExpr * 'c pExpr +| PEmul of 'c pExpr * 'c pExpr +| PEopp of 'c pExpr +| PEpow of 'c pExpr * n + +val mk_X : 'a1 -> 'a1 -> positive -> 'a1 pol + +val ppow_pos : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> 'a1 pol -> positive -> 'a1 pol + +val ppow_N : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> ('a1 pol -> 'a1 pol) -> 'a1 pol -> n -> 'a1 pol + +val norm_aux : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol + +type 'a bFormula = +| TT +| FF +| X +| A of 'a +| Cj of 'a bFormula * 'a bFormula +| D of 'a bFormula * 'a bFormula +| N of 'a bFormula +| I of 'a bFormula * 'a bFormula + +val map_bformula : ('a1 -> 'a2) -> 'a1 bFormula -> 'a2 bFormula + +type 'x clause = 'x list + +type 'x cnf = 'x clause list + +val tt : 'a1 cnf + +val ff : 'a1 cnf + +val add_term : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 -> 'a1 clause -> 'a1 + clause option + +val or_clause : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 clause -> + 'a1 clause option + +val or_clause_cnf : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 clause -> 'a1 cnf -> 'a1 + cnf + +val or_cnf : + ('a1 -> bool) -> ('a1 -> 'a1 -> 'a1 option) -> 'a1 cnf -> 'a1 cnf -> 'a1 cnf + +val and_cnf : 'a1 cnf -> 'a1 cnf -> 'a1 cnf + +val xcnf : + ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 -> + 'a2 cnf) -> bool -> 'a1 bFormula -> 'a2 cnf + +val cnf_checker : ('a1 list -> 'a2 -> bool) -> 'a1 cnf -> 'a2 list -> bool + +val tauto_checker : + ('a2 -> bool) -> ('a2 -> 'a2 -> 'a2 option) -> ('a1 -> 'a2 cnf) -> ('a1 -> + 'a2 cnf) -> ('a2 list -> 'a3 -> bool) -> 'a1 bFormula -> 'a3 list -> bool + +val cneqb : ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool + +val cltb : ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 -> 'a1 -> bool + +type 'c polC = 'c pol + +type op1 = +| Equal +| NonEqual +| Strict +| NonStrict + +type 'c nFormula = 'c polC * op1 + +val opMult : op1 -> op1 -> op1 option + +val opAdd : op1 -> op1 -> op1 option + +type 'c psatz = +| PsatzIn of nat +| PsatzSquare of 'c polC +| PsatzMulC of 'c polC * 'c psatz +| PsatzMulE of 'c psatz * 'c psatz +| PsatzAdd of 'c psatz * 'c psatz +| PsatzC of 'c +| PsatzZ + +val map_option : ('a1 -> 'a2 option) -> 'a1 option -> 'a2 option + +val map_option2 : + ('a1 -> 'a2 -> 'a3 option) -> 'a1 option -> 'a2 option -> 'a3 option + +val pexpr_times_nformula : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> 'a1 polC -> 'a1 nFormula -> 'a1 nFormula option + +val nformula_times_nformula : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> 'a1 nFormula -> 'a1 nFormula -> 'a1 nFormula option + +val nformula_plus_nformula : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> 'a1 + nFormula -> 'a1 nFormula option + +val eval_Psatz : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> 'a1 + nFormula option + +val check_inconsistent : + 'a1 -> ('a1 -> 'a1 -> bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula -> bool + +val check_normalised_formulas : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + bool) -> ('a1 -> 'a1 -> bool) -> 'a1 nFormula list -> 'a1 psatz -> bool + +type op2 = +| OpEq +| OpNEq +| OpLe +| OpGe +| OpLt +| OpGt + +type 't formula = { flhs : 't pExpr; fop : op2; frhs : 't pExpr } + +val norm : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pExpr -> 'a1 pol + +val psub0 : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1) -> ('a1 + -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> 'a1 pol + +val padd0 : + 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 pol -> 'a1 pol -> + 'a1 pol + +val xnormalise : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula + list + +val cnf_normalise : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula + cnf + +val xnegate : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula + list + +val cnf_negate : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> + 'a1) -> ('a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 formula -> 'a1 nFormula + cnf + +val xdenorm : positive -> 'a1 pol -> 'a1 pExpr + +val denorm : 'a1 pol -> 'a1 pExpr + +val map_PExpr : ('a2 -> 'a1) -> 'a2 pExpr -> 'a1 pExpr + +val map_Formula : ('a2 -> 'a1) -> 'a2 formula -> 'a1 formula + +val simpl_cone : + 'a1 -> 'a1 -> ('a1 -> 'a1 -> 'a1) -> ('a1 -> 'a1 -> bool) -> 'a1 psatz -> + 'a1 psatz + +type q = { qnum : z; qden : positive } + +val qnum : q -> z + +val qden : q -> positive + +val qeq_bool : q -> q -> bool + +val qle_bool : q -> q -> bool + +val qplus : q -> q -> q + +val qmult : q -> q -> q + +val qopp : q -> q + +val qminus : q -> q -> q + +val qinv : q -> q + +val qpower_positive : q -> positive -> q + +val qpower : q -> z -> q + +type 'a t = +| Empty +| Leaf of 'a +| Node of 'a t * 'a * 'a t + +val find : 'a1 -> 'a1 t -> positive -> 'a1 + +val singleton : 'a1 -> positive -> 'a1 -> 'a1 t + +val vm_add : 'a1 -> positive -> 'a1 -> 'a1 t -> 'a1 t + +type zWitness = z psatz + +val zWeakChecker : z nFormula list -> z psatz -> bool + +val psub1 : z pol -> z pol -> z pol + +val padd1 : z pol -> z pol -> z pol + +val norm0 : z pExpr -> z pol + +val xnormalise0 : z formula -> z nFormula list + +val normalise : z formula -> z nFormula cnf + +val xnegate0 : z formula -> z nFormula list + +val negate : z formula -> z nFormula cnf + +val zunsat : z nFormula -> bool + +val zdeduce : z nFormula -> z nFormula -> z nFormula option + +val ceiling : z -> z -> z + +type zArithProof = +| DoneProof +| RatProof of zWitness * zArithProof +| CutProof of zWitness * zArithProof +| EnumProof of zWitness * zWitness * zArithProof list + +val zgcdM : z -> z -> z + +val zgcd_pol : z polC -> z * z + +val zdiv_pol : z polC -> z -> z polC + +val makeCuttingPlane : z polC -> z polC * z + +val genCuttingPlane : z nFormula -> ((z polC * z) * op1) option + +val nformula_of_cutting_plane : ((z polC * z) * op1) -> z nFormula + +val is_pol_Z0 : z polC -> bool + +val eval_Psatz0 : z nFormula list -> zWitness -> z nFormula option + +val valid_cut_sign : op1 -> bool + +val zChecker : z nFormula list -> zArithProof -> bool + +val zTautoChecker : z formula bFormula -> zArithProof list -> bool + +type qWitness = q psatz + +val qWeakChecker : q nFormula list -> q psatz -> bool + +val qnormalise : q formula -> q nFormula cnf + +val qnegate : q formula -> q nFormula cnf + +val qunsat : q nFormula -> bool + +val qdeduce : q nFormula -> q nFormula -> q nFormula option + +val qTautoChecker : q formula bFormula -> qWitness list -> bool + +type rcst = +| C0 +| C1 +| CQ of q +| CZ of z +| CPlus of rcst * rcst +| CMinus of rcst * rcst +| CMult of rcst * rcst +| CInv of rcst +| COpp of rcst + +val q_of_Rcst : rcst -> q + +type rWitness = q psatz + +val rWeakChecker : q nFormula list -> q psatz -> bool + +val rnormalise : q formula -> q nFormula cnf + +val rnegate : q formula -> q nFormula cnf + +val runsat : q nFormula -> bool + +val rdeduce : q nFormula -> q nFormula -> q nFormula option + +val rTautoChecker : rcst formula bFormula -> rWitness list -> bool diff --git a/plugins/micromega/sos_types.mli b/plugins/micromega/sos_types.mli new file mode 100644 index 000000000..57c4e50ca --- /dev/null +++ b/plugins/micromega/sos_types.mli @@ -0,0 +1,40 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(* The type of positivstellensatz -- used to communicate with sos *) + +type vname = string;; + +type term = +| Zero +| Const of Num.num +| Var of vname +| Inv of term +| Opp of term +| Add of (term * term) +| Sub of (term * term) +| Mul of (term * term) +| Div of (term * term) +| Pow of (term * int);; + +val output_term : out_channel -> term -> unit + +type positivstellensatz = + Axiom_eq of int + | Axiom_le of int + | Axiom_lt of int + | Rational_eq of Num.num + | Rational_le of Num.num + | Rational_lt of Num.num + | Square of term + | Monoid of int list + | Eqmul of term * positivstellensatz + | Sum of positivstellensatz * positivstellensatz + | Product of positivstellensatz * positivstellensatz;; + +val output_psatz : out_channel -> positivstellensatz -> unit diff --git a/plugins/micromega/vo.itarget b/plugins/micromega/vo.itarget deleted file mode 100644 index a555d5ba1..000000000 --- a/plugins/micromega/vo.itarget +++ /dev/null @@ -1,16 +0,0 @@ -MExtraction.vo -EnvRing.vo -Env.vo -OrderedRing.vo -Psatz.vo -QMicromega.vo -Refl.vo -RingMicromega.vo -RMicromega.vo -Tauto.vo -VarMap.vo -ZCoeff.vo -ZMicromega.vo -Lia.vo -Lqa.vo -Lra.vo diff --git a/plugins/nsatz/g_nsatz.ml4 b/plugins/nsatz/g_nsatz.ml4 index 759885253..5a6d72036 100644 --- a/plugins/nsatz/g_nsatz.ml4 +++ b/plugins/nsatz/g_nsatz.ml4 @@ -8,8 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Ltac_plugin -open Names DECLARE PLUGIN "nsatz_plugin" diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml index 6ba4c0f93..dd1d8764a 100644 --- a/plugins/nsatz/nsatz.ml +++ b/plugins/nsatz/nsatz.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open CErrors open Util open Term diff --git a/plugins/nsatz/nsatz.mli b/plugins/nsatz/nsatz.mli index e876ccfa5..c0dad72ad 100644 --- a/plugins/nsatz/nsatz.mli +++ b/plugins/nsatz/nsatz.mli @@ -6,4 +6,5 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -val nsatz_compute : Constr.t -> unit Proofview.tactic +open API +val nsatz_compute : Term.constr -> unit Proofview.tactic diff --git a/plugins/nsatz/vo.itarget b/plugins/nsatz/vo.itarget deleted file mode 100644 index 06fc88343..000000000 --- a/plugins/nsatz/vo.itarget +++ /dev/null @@ -1 +0,0 @@ -Nsatz.vo diff --git a/plugins/omega/PreOmega.v b/plugins/omega/PreOmega.v index 5f5f548f8..2780be4aa 100644 --- a/plugins/omega/PreOmega.v +++ b/plugins/omega/PreOmega.v @@ -48,10 +48,13 @@ Ltac zify_unop_var_or_term t thm a := (remember a as za; zify_unop_core t thm za). Ltac zify_unop t thm a := - (* if a is a scalar, we can simply reduce the unop *) + (* If a is a scalar, we can simply reduce the unop. *) + (* Note that simpl wasn't enough to reduce [Z.max 0 0] (#5439) *) let isz := isZcst a in match isz with - | true => simpl (t a) in * + | true => + let u := eval compute in (t a) in + change (t a) with u in * | _ => zify_unop_var_or_term t thm a end. @@ -165,21 +168,31 @@ Ltac zify_nat_op := rewrite (Nat2Z.inj_mul a b) in * (* O -> Z0 *) - | H : context [ Z.of_nat O ] |- _ => simpl (Z.of_nat O) in H - | |- context [ Z.of_nat O ] => simpl (Z.of_nat O) + | H : context [ Z.of_nat O ] |- _ => change (Z.of_nat O) with Z0 in H + | |- context [ Z.of_nat O ] => change (Z.of_nat O) with Z0 (* S -> number or Z.succ *) | H : context [ Z.of_nat (S ?a) ] |- _ => let isnat := isnatcst a in match isnat with - | true => simpl (Z.of_nat (S a)) in H + | true => + let t := eval compute in (Z.of_nat (S a)) in + change (Z.of_nat (S a)) with t in H | _ => rewrite (Nat2Z.inj_succ a) in H + | _ => (* if the [rewrite] fails (most likely a dependent occurence of [Z.of_nat (S a)]), + hide [Z.of_nat (S a)] in this one hypothesis *) + change (Z.of_nat (S a)) with (Z_of_nat' (S a)) in H end | |- context [ Z.of_nat (S ?a) ] => let isnat := isnatcst a in match isnat with - | true => simpl (Z.of_nat (S a)) + | true => + let t := eval compute in (Z.of_nat (S a)) in + change (Z.of_nat (S a)) with t | _ => rewrite (Nat2Z.inj_succ a) + | _ => (* if the [rewrite] fails (most likely a dependent occurence of [Z.of_nat (S a)]), + hide [Z.of_nat (S a)] in the goal *) + change (Z.of_nat (S a)) with (Z_of_nat' (S a)) end (* atoms of type nat : we add a positivity condition (if not already there) *) @@ -258,8 +271,8 @@ Ltac zify_positive_op := | |- context [ Zpos (Pos.max ?a ?b) ] => rewrite (Pos2Z.inj_max a b) (* Pos.sub -> Z.max 1 (Z.sub ... ...) *) - | H : context [ Zpos (Pos.sub ?a ?b) ] |- _ => rewrite (Pos2Z.inj_sub a b) in H - | |- context [ Zpos (Pos.sub ?a ?b) ] => rewrite (Pos2Z.inj_sub a b) + | H : context [ Zpos (Pos.sub ?a ?b) ] |- _ => rewrite (Pos2Z.inj_sub_max a b) in H + | |- context [ Zpos (Pos.sub ?a ?b) ] => rewrite (Pos2Z.inj_sub_max a b) (* Pos.succ -> Z.succ *) | H : context [ Zpos (Pos.succ ?a) ] |- _ => rewrite (Pos2Z.inj_succ a) in H @@ -401,4 +414,3 @@ Ltac zify_N := repeat zify_N_rel; repeat zify_N_op; unfold Z_of_N' in *. (** The complete Z-ification tactic *) Ltac zify := repeat (zify_nat; zify_positive; zify_N); zify_op. - diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 465e77019..440a10bfb 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -13,6 +13,7 @@ (* *) (**************************************************************************) +open API open CErrors open Util open Names @@ -651,7 +652,7 @@ let clever_rewrite_base_poly typ p result theorem = let full = pf_concl gl in let env = pf_env gl in let (abstracted,occ) = abstract_path (project gl) typ (List.rev p) full in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let t = applist (mkLambda @@ -687,7 +688,7 @@ let clever_rewrite_gen_nat p result (t,args) = (** Solve using the term the term [t _] *) let refine_app gl t = let open Tacmach.New in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let env = pf_env gl in let ht = match EConstr.kind sigma (pf_get_type_of gl t) with | Prod (_, t, _) -> t @@ -707,6 +708,39 @@ let clever_rewrite p vpath t = refine_app gl t' end +(** simpl_coeffs : + The subterm at location [path_init] in the current goal should + look like [(v1*c1 + (v2*c2 + ... (vn*cn + k)))], and we reduce + via "simpl" each [ci] and the final constant [k]. + The path [path_k] gives the location of constant [k]. + Earlier, the whole was a mere call to [focused_simpl], + leading to reduction inside the atoms [vi], which is bad, + for instance when the atom is an evaluable definition + (see #4132). *) + +let simpl_coeffs path_init path_k = + Proofview.Goal.enter begin fun gl -> + let sigma = project gl in + let rec loop n t = + if Int.equal n 0 then pf_nf gl t + else + (* t should be of the form ((v * c) + ...) *) + match EConstr.kind sigma t with + | App(f,[|t1;t2|]) -> + (match EConstr.kind sigma t1 with + | App (g,[|v;c|]) -> + let c' = pf_nf gl c in + let t2' = loop (pred n) t2 in + mkApp (f,[|mkApp (g,[|v;c'|]);t2'|]) + | _ -> assert false) + | _ -> assert false + in + let n = Pervasives.(-) (List.length path_k) (List.length path_init) in + let newc = context sigma (fun _ t -> loop n t) (List.rev path_init) (pf_concl gl) + in + convert_concl_no_check newc DEFAULTcast + end + let rec shuffle p (t1,t2) = match t1,t2 with | Oplus(l1,r1), Oplus(l2,r2) -> @@ -769,7 +803,7 @@ let shuffle_mult p_init k1 e1 k2 e2 = let tac' = clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) in - tac :: focused_simpl (P_APP 1::P_APP 2:: p) :: tac' :: + tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: loop p (l1,l2) else tac :: loop (P_APP 2 :: p) (l1,l2) else if v1 > v2 then @@ -804,7 +838,7 @@ let shuffle_mult p_init k1 e1 k2 e2 = [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA12) :: loop (P_APP 2 :: p) ([],l2) - | [],[] -> [focused_simpl p_init] + | [],[] -> [simpl_coeffs p_init p] in loop p_init (e1,e2) @@ -827,7 +861,7 @@ let shuffle_mult_right p_init e1 k2 e2 = clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) in - tac :: focused_simpl (P_APP 1::P_APP 2:: p) :: tac' :: + tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: loop p (l1,l2) else tac :: loop (P_APP 2 :: p) (l1,l2) else if v1 > v2 then @@ -854,7 +888,7 @@ let shuffle_mult_right p_init e1 k2 e2 = [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA12) :: loop (P_APP 2 :: p) ([],l2) - | [],[] -> [focused_simpl p_init] + | [],[] -> [simpl_coeffs p_init p] in loop p_init (e1,e2) @@ -895,7 +929,7 @@ let rec scalar p n = function let scalar_norm p_init = let rec loop p = function - | [] -> [focused_simpl p_init] + | [] -> [simpl_coeffs p_init p] | (_::l) -> clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1];[P_APP 1; P_APP 1; P_APP 2]; @@ -906,7 +940,7 @@ let scalar_norm p_init = let norm_add p_init = let rec loop p = function - | [] -> [focused_simpl p_init] + | [] -> [simpl_coeffs p_init p] | _:: l -> clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: @@ -916,7 +950,7 @@ let norm_add p_init = let scalar_norm_add p_init = let rec loop p = function - | [] -> [focused_simpl p_init] + | [] -> [simpl_coeffs p_init p] | _ :: l -> clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; diff --git a/plugins/omega/g_omega.ml4 b/plugins/omega/g_omega.ml4 index ce7ffb1e7..2fcf076f1 100644 --- a/plugins/omega/g_omega.ml4 +++ b/plugins/omega/g_omega.ml4 @@ -15,6 +15,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API + DECLARE PLUGIN "omega_plugin" open Ltac_plugin @@ -24,7 +26,7 @@ open Stdarg let eval_tactic name = let dp = DirPath.make (List.map Id.of_string ["PreOmega"; "omega"; "Coq"]) in - let kn = KerName.make2 (MPfile dp) (Label.make name) in + let kn = KerName.make2 (ModPath.MPfile dp) (Label.make name) in let tac = Tacenv.interp_ltac kn in Tacinterp.eval_tactic tac diff --git a/plugins/omega/vo.itarget b/plugins/omega/vo.itarget deleted file mode 100644 index 842210e21..000000000 --- a/plugins/omega/vo.itarget +++ /dev/null @@ -1,5 +0,0 @@ -OmegaLemmas.vo -OmegaPlugin.vo -OmegaTactic.vo -Omega.vo -PreOmega.vo diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4 index 980f03db3..c43d7d0b5 100644 --- a/plugins/quote/g_quote.ml4 +++ b/plugins/quote/g_quote.ml4 @@ -8,6 +8,7 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API open Ltac_plugin open Names open Misctypes diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index ffacd8b36..15d0f5f37 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -101,6 +101,7 @@ (*i*) +open API open CErrors open Util open Names @@ -168,8 +169,8 @@ exchange ?1 and ?2 in the example above) module ConstrSet = Set.Make( struct - type t = Constr.constr - let compare = constr_ord + type t = Term.constr + let compare = Term.compare end) type inversion_scheme = { @@ -386,7 +387,7 @@ let rec sort_subterm gl l = | h::t -> insert h (sort_subterm gl t) module Constrhash = Hashtbl.Make - (struct type t = Constr.constr + (struct type t = Term.constr let equal = Term.eq_constr let hash = Term.hash_constr end) diff --git a/plugins/quote/vo.itarget b/plugins/quote/vo.itarget deleted file mode 100644 index 7a44fc5aa..000000000 --- a/plugins/quote/vo.itarget +++ /dev/null @@ -1 +0,0 @@ -Quote.vo
\ No newline at end of file diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml index d97dea039..06c80a825 100644 --- a/plugins/romega/const_omega.ml +++ b/plugins/romega/const_omega.ml @@ -6,6 +6,9 @@ *************************************************************************) +open API +open Names + let module_refl_name = "ReflOmegaCore" let module_refl_path = ["Coq"; "romega"; module_refl_name] @@ -37,7 +40,7 @@ let destructurate t = | Term.Ind (isp,_), args -> Kapp (string_of_global (Globnames.IndRef isp), args) | Term.Var id, [] -> Kvar(Names.Id.to_string id) - | Term.Prod (Names.Anonymous,typ,body), [] -> Kimp(typ,body) + | Term.Prod (Anonymous,typ,body), [] -> Kimp(typ,body) | _ -> Kufo exception DestConstApp @@ -242,7 +245,7 @@ let minus = lazy (z_constant "Z.sub") let recognize_pos t = let rec loop t = let f,l = dest_const_apply t in - match Names.Id.to_string f,l with + match Id.to_string f,l with | "xI",[t] -> Bigint.add Bigint.one (Bigint.mult Bigint.two (loop t)) | "xO",[t] -> Bigint.mult Bigint.two (loop t) | "xH",[] -> Bigint.one @@ -253,7 +256,7 @@ let recognize_pos t = let recognize_Z t = try let f,l = dest_const_apply t in - match Names.Id.to_string f,l with + match Id.to_string f,l with | "Zpos",[t] -> recognize_pos t | "Zneg",[t] -> Option.map Bigint.neg (recognize_pos t) | "Z0",[] -> Some Bigint.zero diff --git a/plugins/romega/const_omega.mli b/plugins/romega/const_omega.mli index a452b1a91..6dc5d9f7e 100644 --- a/plugins/romega/const_omega.mli +++ b/plugins/romega/const_omega.mli @@ -6,6 +6,7 @@ *************************************************************************) +open API (** Coq objects used in romega *) diff --git a/plugins/romega/g_romega.ml4 b/plugins/romega/g_romega.ml4 index 6479c683b..53f6f42c8 100644 --- a/plugins/romega/g_romega.ml4 +++ b/plugins/romega/g_romega.ml4 @@ -8,6 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API + DECLARE PLUGIN "romega_plugin" open Ltac_plugin @@ -17,7 +19,7 @@ open Stdarg let eval_tactic name = let dp = DirPath.make (List.map Id.of_string ["PreOmega"; "omega"; "Coq"]) in - let kn = KerName.make2 (MPfile dp) (Label.make name) in + let kn = KerName.make2 (ModPath.MPfile dp) (Label.make name) in let tac = Tacenv.interp_ltac kn in Tacinterp.eval_tactic tac diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml index 575634174..1a53862ec 100644 --- a/plugins/romega/refl_omega.ml +++ b/plugins/romega/refl_omega.ml @@ -6,6 +6,7 @@ *************************************************************************) +open API open Pp open Util open Const_omega diff --git a/plugins/romega/vo.itarget b/plugins/romega/vo.itarget deleted file mode 100644 index f7a3c41c7..000000000 --- a/plugins/romega/vo.itarget +++ /dev/null @@ -1,2 +0,0 @@ -ReflOmegaCore.vo -ROmega.vo diff --git a/plugins/rtauto/g_rtauto.ml4 b/plugins/rtauto/g_rtauto.ml4 index 7e58ef9a3..565308f72 100644 --- a/plugins/rtauto/g_rtauto.ml4 +++ b/plugins/rtauto/g_rtauto.ml4 @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + (*i camlp4deps: "grammar/grammar.cma" i*) open Ltac_plugin diff --git a/plugins/rtauto/proof_search.ml b/plugins/rtauto/proof_search.ml index 153a6a49a..8dd7a5e46 100644 --- a/plugins/rtauto/proof_search.ml +++ b/plugins/rtauto/proof_search.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open CErrors open Util open Goptions diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index 1b07a8ca8..f84eebadc 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + module Search = Explore.Make(Proof_search) open Ltac_plugin @@ -299,7 +301,7 @@ let rtauto_tac gls= build_form formula; build_proof [] 0 prf|]) in let term= - applist (main,List.rev_map (fun (id,_) -> mkVar id) hyps) in + applistc main (List.rev_map (fun (id,_) -> mkVar id) hyps) in let build_end_time=System.get_time () in let _ = if !verbose then begin diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli index 092552364..ac260e51a 100644 --- a/plugins/rtauto/refl_tauto.mli +++ b/plugins/rtauto/refl_tauto.mli @@ -7,16 +7,18 @@ (************************************************************************) (* raises Not_found if no proof is found *) +open API + type atom_env= {mutable next:int; mutable env:(Term.constr*int) list} val make_form : atom_env -> - Proof_type.goal Tacmach.sigma -> EConstr.types -> Proof_search.form + Proof_type.goal Evd.sigma -> EConstr.types -> Proof_search.form val make_hyps : atom_env -> - Proof_type.goal Tacmach.sigma -> + Proof_type.goal Evd.sigma -> EConstr.types list -> EConstr.named_context -> (Names.Id.t * Proof_search.form) list diff --git a/plugins/rtauto/vo.itarget b/plugins/rtauto/vo.itarget deleted file mode 100644 index 4c9364ad7..000000000 --- a/plugins/rtauto/vo.itarget +++ /dev/null @@ -1,2 +0,0 @@ -Bintree.vo -Rtauto.vo diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4 index 05ab8ab32..ada41274f 100644 --- a/plugins/setoid_ring/g_newring.ml4 +++ b/plugins/setoid_ring/g_newring.ml4 @@ -8,6 +8,8 @@ (*i camlp4deps: "grammar/grammar.cma" i*) +open API +open Grammar_API open Ltac_plugin open Pp open Util diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 85cbdc5a4..ee75d2908 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Ltac_plugin open Pp open Util @@ -151,7 +152,7 @@ let ic_unsafe c = (*FIXME remove *) EConstr.of_constr (fst (Constrintern.interp_constr env sigma c)) let decl_constant na ctx c = - let open Constr in + let open Term in let vars = Universes.universes_of_constr c in let ctx = Universes.restrict_universe_context (Univ.ContextSet.of_context ctx) vars in mkConst(declare_constant (Id.of_string na) @@ -282,7 +283,7 @@ let my_reference c = let znew_ring_path = DirPath.make (List.map Id.of_string ["InitialRing";plugin_dir;"Coq"]) let zltac s = - lazy(make_kn (MPfile znew_ring_path) DirPath.empty (Label.make s)) + lazy(KerName.make (ModPath.MPfile znew_ring_path) DirPath.empty (Label.make s)) let mk_cst l s = lazy (Coqlib.coq_reference "newring" l s);; let pol_cst s = mk_cst [plugin_dir;"Ring_polynom"] s ;; @@ -346,7 +347,11 @@ let _ = add_map "ring" let pr_constr c = pr_econstr c -module Cmap = Map.Make(Constr) +module M = struct + type t = Term.constr + let compare = Term.compare +end +module Cmap = Map.Make(M) let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table" let from_name = Summary.ref Spmap.empty ~name:"ring-tac-name-table" @@ -769,7 +774,7 @@ let new_field_path = DirPath.make (List.map Id.of_string ["Field_tac";plugin_dir;"Coq"]) let field_ltac s = - lazy(make_kn (MPfile new_field_path) DirPath.empty (Label.make s)) + lazy(KerName.make (ModPath.MPfile new_field_path) DirPath.empty (Label.make s)) let _ = add_map "field" @@ -929,7 +934,7 @@ let field_equality evd r inv req = inv_m_lem let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign odiv = - let open Constr in + let open Term in check_required_library (cdir@["Field_tac"]); let (sigma,fth) = ic fth in let env = Global.env() in diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli index d9d32c681..7f685063c 100644 --- a/plugins/setoid_ring/newring.mli +++ b/plugins/setoid_ring/newring.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Names open EConstr open Libnames diff --git a/plugins/setoid_ring/newring_ast.mli b/plugins/setoid_ring/newring_ast.mli index c26fcc8d1..b7afd2eff 100644 --- a/plugins/setoid_ring/newring_ast.mli +++ b/plugins/setoid_ring/newring_ast.mli @@ -6,7 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Constr +open API +open Term open Libnames open Constrexpr open Tacexpr diff --git a/plugins/setoid_ring/vo.itarget b/plugins/setoid_ring/vo.itarget deleted file mode 100644 index 595ba55ec..000000000 --- a/plugins/setoid_ring/vo.itarget +++ /dev/null @@ -1,24 +0,0 @@ -ArithRing.vo -BinList.vo -Field_tac.vo -Field_theory.vo -Field.vo -InitialRing.vo -NArithRing.vo -RealField.vo -Ring_base.vo -Ring_polynom.vo -Ring_tac.vo -Ring_theory.vo -Ring.vo -ZArithRing.vo -Algebra_syntax.vo -Cring.vo -Ncring.vo -Ncring_polynom.vo -Ncring_initial.vo -Ncring_tac.vo -Rings_Z.vo -Rings_R.vo -Rings_Q.vo -Integral_domain.vo
\ No newline at end of file diff --git a/plugins/ssr/ssrast.mli b/plugins/ssr/ssrast.mli index 69202ae2d..0f4b86d10 100644 --- a/plugins/ssr/ssrast.mli +++ b/plugins/ssr/ssrast.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Names open Ltac_plugin @@ -78,7 +79,7 @@ type ssripat = | IPatView of ssrterm list (* /view *) | IPatClear of ssrclear (* {H1 H2} *) | IPatSimpl of ssrsimpl - | IPatNewHidden of identifier list + | IPatNewHidden of Id.t list (* | IPatVarsForAbstract of Id.t list *) and ssripats = ssripat list @@ -93,10 +94,10 @@ type ssrintrosarg = Tacexpr.raw_tactic_expr * ssripats type ssrfwdid = Id.t (** Binders (for fwd tactics) *) type 'term ssrbind = - | Bvar of name - | Bdecl of name list * 'term - | Bdef of name * 'term option * 'term - | Bstruct of name + | Bvar of Name.t + | Bdecl of Name.t list * 'term + | Bdef of Name.t * 'term option * 'term + | Bstruct of Name.t | Bcast of 'term (* We use an intermediate structure to correctly render the binder list *) (* abbreviations. We use a list of hints to extract the binders and *) diff --git a/plugins/ssr/ssrbwd.ml b/plugins/ssr/ssrbwd.ml index cc0e86684..3988f00ba 100644 --- a/plugins/ssr/ssrbwd.ml +++ b/plugins/ssr/ssrbwd.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Printer open Pretyping open Globnames diff --git a/plugins/ssr/ssrbwd.mli b/plugins/ssr/ssrbwd.mli index 8bf785a21..b0e98bdb4 100644 --- a/plugins/ssr/ssrbwd.mli +++ b/plugins/ssr/ssrbwd.mli @@ -8,6 +8,8 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API + val apply_top_tac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma val inner_ssrapplytac : diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml index e90be92cf..490ded9d4 100644 --- a/plugins/ssr/ssrcommon.ml +++ b/plugins/ssr/ssrcommon.ml @@ -8,10 +8,12 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API +open Grammar_API open Util open Names open Evd -open Constr +open Term open Termops open Printer open Locusops @@ -131,7 +133,7 @@ let tac_on_all gl tac = (* Used to thread data between intro patterns at run time *) type tac_ctx = { - tmp_ids : (Id.t * name ref) list; + tmp_ids : (Id.t * Name.t ref) list; wild_ids : Id.t list; delayed_clears : Id.t list; } @@ -224,8 +226,8 @@ let isAppInd gl c = let interp_refine ist gl rc = let constrvars = Tacinterp.extract_ltac_constr_values ist (pf_env gl) in - let vars = { Pretyping.empty_lvar with - Pretyping.ltac_constrs = constrvars; ltac_genargs = ist.Tacinterp.lfun + let vars = { Glob_ops.empty_lvar with + Glob_term.ltac_constrs = constrvars; ltac_genargs = ist.Tacinterp.lfun } in let kind = Pretyping.OfType (pf_concl gl) in let flags = { @@ -306,7 +308,7 @@ let is_internal_name s = List.exists (fun p -> p s) !internal_names let tmp_tag = "_the_" let tmp_post = "_tmp_" let mk_tmp_id i = - id_of_string (Printf.sprintf "%s%s%s" tmp_tag (CString.ordinal i) tmp_post) + Id.of_string (Printf.sprintf "%s%s%s" tmp_tag (CString.ordinal i) tmp_post) let new_tmp_id ctx = let id = mk_tmp_id (1 + List.length ctx.tmp_ids) in let orig = ref Anonymous in @@ -316,7 +318,7 @@ let new_tmp_id ctx = let mk_internal_id s = let s' = Printf.sprintf "_%s_" s in let s' = String.map (fun c -> if c = ' ' then '_' else c) s' in - add_internal_name ((=) s'); id_of_string s' + add_internal_name ((=) s'); Id.of_string s' let same_prefix s t n = let rec loop i = i = n || s.[i] = t.[i] && loop (i + 1) in loop 0 @@ -325,7 +327,7 @@ let skip_digits s = let n = String.length s in let rec loop i = if i < n && is_digit s.[i] then loop (i + 1) else i in loop -let mk_tagged_id t i = id_of_string (Printf.sprintf "%s%d_" t i) +let mk_tagged_id t i = Id.of_string (Printf.sprintf "%s%d_" t i) let is_tagged t s = let n = String.length s - 1 and m = String.length t in m < n && s.[n] = '_' && same_prefix s t m && skip_digits s m = n @@ -339,7 +341,7 @@ let ssr_anon_hyp = "Hyp" let wildcard_tag = "_the_" let wildcard_post = "_wildcard_" let mk_wildcard_id i = - id_of_string (Printf.sprintf "%s%s%s" wildcard_tag (CString.ordinal i) wildcard_post) + Id.of_string (Printf.sprintf "%s%s%s" wildcard_tag (CString.ordinal i) wildcard_post) let has_wildcard_tag s = let n = String.length s in let m = String.length wildcard_tag in let m' = String.length wildcard_post in @@ -355,15 +357,15 @@ let new_wild_id ctx = let discharged_tag = "_discharged_" let mk_discharged_id id = - id_of_string (Printf.sprintf "%s%s_" discharged_tag (string_of_id id)) + Id.of_string (Printf.sprintf "%s%s_" discharged_tag (Id.to_string id)) let has_discharged_tag s = let m = String.length discharged_tag and n = String.length s - 1 in m < n && s.[n] = '_' && same_prefix s discharged_tag m let _ = add_internal_name has_discharged_tag -let is_discharged_id id = has_discharged_tag (string_of_id id) +let is_discharged_id id = has_discharged_tag (Id.to_string id) let max_suffix m (t, j0 as tj0) id = - let s = string_of_id id in let n = String.length s - 1 in + let s = Id.to_string id in let n = String.length s - 1 in let dn = String.length t - 1 - n in let i0 = j0 - dn in if not (i0 >= m && s.[n] = '_' && same_prefix s t m) then tj0 else let rec loop i = @@ -383,7 +385,7 @@ let mk_anon_id t gl = let rec loop i j = let d = !s.[i] in if not (is_digit d) then i + 1, j else loop (i - 1) (if d = '0' then j else i) in - let m, j = loop (n - 1) n in m, (!s, j), id_of_string !s in + let m, j = loop (n - 1) n in m, (!s, j), Id.of_string !s in let gl_ids = pf_ids_of_hyps gl in if not (List.mem id0 gl_ids) then id0 else let s, i = List.fold_left (max_suffix m) si0 gl_ids in @@ -401,7 +403,7 @@ let convert_concl t = Tactics.convert_concl t Term.DEFAULTcast let rename_hd_prod orig_name_ref gl = match EConstr.kind (project gl) (pf_concl gl) with - | Constr.Prod(_,src,tgt) -> + | Term.Prod(_,src,tgt) -> Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkProd (!orig_name_ref,src,tgt))) gl | _ -> CErrors.anomaly (str "gentac creates no product") @@ -600,7 +602,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let rec loopP evlist c i = function | (_, (n, t, _)) :: evl -> let t = get evlist (i - 1) t in - let n = Name (id_of_string (ssr_anon_hyp ^ string_of_int n)) in + let n = Name (Id.of_string (ssr_anon_hyp ^ string_of_int n)) in loopP evlist (mkProd (n, t, c)) (i - 1) evl | [] -> c in let rec loop c i = function @@ -624,13 +626,13 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let nb_evar_deps = function | Name id -> - let s = string_of_id id in + let s = Id.to_string id in if not (is_tagged evar_tag s) then 0 else let m = String.length evar_tag in (try int_of_string (String.sub s m (String.length s - 1 - m)) with _ -> 0) | _ -> 0 -let pf_type_id gl t = id_of_string (Namegen.hdchar (pf_env gl) (project gl) t) +let pf_type_id gl t = Id.of_string (Namegen.hdchar (pf_env gl) (project gl) t) let pfe_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty @@ -689,7 +691,7 @@ let pf_merge_uc_of sigma gl = let rec constr_name sigma c = match EConstr.kind sigma c with | Var id -> Name id | Cast (c', _, _) -> constr_name sigma c' - | Const (cn,_) -> Name (id_of_label (con_label cn)) + | Const (cn,_) -> Name (Label.to_id (Constant.label cn)) | App (c', _) -> constr_name sigma c' | _ -> Anonymous @@ -701,9 +703,9 @@ let pf_mkprod gl c ?(name=constr_name (project gl) c) cl = let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (Termops.subst_term (project gl) c cl) (** look up a name in the ssreflect internals module *) -let ssrdirpath = make_dirpath [id_of_string "ssreflect"] -let ssrqid name = Libnames.make_qualid ssrdirpath (id_of_string name) -let ssrtopqid name = Libnames.make_short_qualid (id_of_string name) +let ssrdirpath = DirPath.make [Id.of_string "ssreflect"] +let ssrqid name = Libnames.make_qualid ssrdirpath (Id.of_string name) +let ssrtopqid name = Libnames.qualid_of_ident (Id.of_string name) let locate_reference qid = Smartlocate.global_of_extended_global (Nametab.locate_extended qid) let mkSsrRef name = @@ -812,7 +814,7 @@ let ssr_n_tac seed n gl = let name = if n = -1 then seed else ("ssr" ^ seed ^ string_of_int n) in let fail msg = CErrors.user_err (Pp.str msg) in let tacname = - try Nametab.locate_tactic (Libnames.qualid_of_ident (id_of_string name)) + try Nametab.locate_tactic (Libnames.qualid_of_ident (Id.of_string name)) with Not_found -> try Nametab.locate_tactic (ssrqid name) with Not_found -> if n = -1 then fail "The ssreflect library was not loaded" @@ -1080,7 +1082,7 @@ let introid ?(orig=ref Anonymous) name = tclTHEN (fun gl -> let anontac decl gl = let id = match RelDecl.get_name decl with | Name id -> - if is_discharged_id id then id else mk_anon_id (string_of_id id) gl + if is_discharged_id id then id else mk_anon_id (Id.to_string id) gl | _ -> mk_anon_id ssr_anon_hyp gl in introid id gl diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli index 834b7b722..7a4b47a46 100644 --- a/plugins/ssr/ssrcommon.mli +++ b/plugins/ssr/ssrcommon.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Names open Environ open Proof_type @@ -56,7 +57,7 @@ type 'a tac_a = (goal * 'a) sigma -> (goal * 'a) list sigma (* Thread around names to be cleared or generalized back, and the speed *) type tac_ctx = { - tmp_ids : (Id.t * name ref) list; + tmp_ids : (Id.t * Name.t ref) list; wild_ids : Id.t list; (* List of variables to be cleared at the end of the sentence *) delayed_clears : Id.t list; @@ -173,18 +174,18 @@ val pf_type_id : Proof_type.goal Evd.sigma -> EConstr.types -> Id.t val pf_abs_evars : Proof_type.goal Evd.sigma -> evar_map * EConstr.t -> - int * EConstr.t * Constr.existential_key list * - Evd.evar_universe_context + int * EConstr.t * Evar.t list * + UState.t val pf_abs_evars2 : (* ssr2 *) Proof_type.goal Evd.sigma -> Evar.t list -> evar_map * EConstr.t -> - int * EConstr.t * Constr.existential_key list * - Evd.evar_universe_context + int * EConstr.t * Evar.t list * + UState.t val pf_abs_cterm : Proof_type.goal Evd.sigma -> int -> EConstr.t -> EConstr.t val pf_merge_uc : - Evd.evar_universe_context -> 'a Evd.sigma -> 'a Evd.sigma + UState.t -> 'a Evd.sigma -> 'a Evd.sigma val pf_merge_uc_of : evar_map -> 'a Evd.sigma -> 'a Evd.sigma val constr_name : evar_map -> EConstr.t -> Name.t @@ -195,14 +196,14 @@ val pfe_type_of : Proof_type.goal Evd.sigma -> EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types val pf_abs_prod : - Names.name -> + Name.t -> Proof_type.goal Evd.sigma -> EConstr.t -> EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types val pf_mkprod : Proof_type.goal Evd.sigma -> EConstr.t -> - ?name:Names.name -> + ?name:Name.t -> EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types val mkSsrRRef : string -> Glob_term.glob_constr * 'a option @@ -214,7 +215,7 @@ val pf_mkSsrConst : string -> Proof_type.goal Evd.sigma -> EConstr.t * Proof_type.goal Evd.sigma -val new_wild_id : tac_ctx -> Names.identifier * tac_ctx +val new_wild_id : tac_ctx -> Names.Id.t * tac_ctx val pf_fresh_global : @@ -228,7 +229,7 @@ val is_tagged : string -> string -> bool val has_discharged_tag : string -> bool val ssrqid : string -> Libnames.qualid val new_tmp_id : - tac_ctx -> (Names.identifier * Names.name ref) * tac_ctx + tac_ctx -> (Names.Id.t * Name.t ref) * tac_ctx val mk_anon_id : string -> Proof_type.goal Evd.sigma -> Id.t val pf_abs_evars_pirrel : Proof_type.goal Evd.sigma -> @@ -252,7 +253,7 @@ val red_product_skip_id : env -> evar_map -> EConstr.t -> EConstr.t val ssrautoprop_tac : - (Constr.existential_key Evd.sigma -> Constr.existential_key list Evd.sigma) ref + (Evar.t Evd.sigma -> Evar.t list Evd.sigma) ref val mkProt : EConstr.t -> @@ -285,7 +286,7 @@ val pf_abs_ssrterm : ist -> Proof_type.goal Evd.sigma -> ssrterm -> - evar_map * EConstr.t * Evd.evar_universe_context * int + evar_map * EConstr.t * UState.t * int val pf_interp_ty : ?resolve_typeclasses:bool -> @@ -293,7 +294,7 @@ val pf_interp_ty : Proof_type.goal Evd.sigma -> Ssrast.ssrtermkind * (Glob_term.glob_constr * Constrexpr.constr_expr option) -> - int * EConstr.t * EConstr.t * Evd.evar_universe_context + int * EConstr.t * EConstr.t * UState.t val ssr_n_tac : string -> int -> v82tac val donetac : int -> v82tac @@ -361,7 +362,7 @@ val pf_interp_gen_aux : (Ssrast.ssrhyp list option * Ssrmatching_plugin.Ssrmatching.occ) * Ssrmatching_plugin.Ssrmatching.cpattern -> bool * Ssrmatching_plugin.Ssrmatching.pattern * EConstr.t * - EConstr.t * Ssrast.ssrhyp list * Evd.evar_universe_context * + EConstr.t * Ssrast.ssrhyp list * UState.t * Proof_type.goal Evd.sigma val is_name_in_ipats : @@ -376,7 +377,7 @@ val mk_profiler : string -> profiler (** Basic tactics *) -val introid : ?orig:name ref -> Id.t -> v82tac +val introid : ?orig:Name.t ref -> Id.t -> v82tac val intro_anon : v82tac val intro_all : v82tac diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml index 832044909..bd9a05891 100644 --- a/plugins/ssr/ssrelim.ml +++ b/plugins/ssr/ssrelim.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Util open Names open Printer diff --git a/plugins/ssr/ssrelim.mli b/plugins/ssr/ssrelim.mli index fb1b58ac3..8dc28d8b7 100644 --- a/plugins/ssr/ssrelim.mli +++ b/plugins/ssr/ssrelim.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Ssrmatching_plugin val ssrelim : diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml index af315aac5..b0fe89897 100644 --- a/plugins/ssr/ssrequality.ml +++ b/plugins/ssr/ssrequality.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Ltac_plugin open Util open Names @@ -342,9 +343,9 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl = let elim, gl = pf_fresh_global (Indrec.lookup_eliminator ind sort) gl in if dir = R2L then elim, gl else (* taken from Coq's rewrite *) let elim, _ = Term.destConst elim in - let mp,dp,l = repr_con (constant_of_kn (canonical_con elim)) in - let l' = label_of_id (Nameops.add_suffix (id_of_label l) "_r") in - let c1' = Global.constant_of_delta_kn (canonical_con (make_con mp dp l')) in + let mp,dp,l = Constant.repr3 (Constant.make1 (Constant.canonical elim)) in + let l' = Label.of_id (Nameops.add_suffix (Label.to_id l) "_r") in + let c1' = Global.constant_of_delta_kn (Constant.canonical (Constant.make3 mp dp l')) in mkConst c1', gl in let elim = EConstr.of_constr elim in let proof = EConstr.mkApp (elim, [| rdx_ty; new_rdx; pred; p; rdx; c |]) in diff --git a/plugins/ssr/ssrequality.mli b/plugins/ssr/ssrequality.mli index 9c5fd4983..f712002c1 100644 --- a/plugins/ssr/ssrequality.mli +++ b/plugins/ssr/ssrequality.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Ssrmatching_plugin open Ssrast diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml index 663bca15e..660c2e776 100644 --- a/plugins/ssr/ssrfwd.ml +++ b/plugins/ssr/ssrfwd.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Names open Tacmach @@ -200,7 +201,7 @@ let havetac ist let assert_is_conv gl = try Proofview.V82.of_tactic (convert_concl (EConstr.it_mkProd_or_LetIn concl ctx)) gl with _ -> errorstrm (str "Given proof term is not of type " ++ - pr_econstr (EConstr.mkArrow (EConstr.mkVar (id_of_string "_")) concl)) in + pr_econstr (EConstr.mkArrow (EConstr.mkVar (Id.of_string "_")) concl)) in gl, ty, Tacticals.tclTHEN assert_is_conv (Proofview.V82.of_tactic (Tactics.apply t)), id, itac_c | FwdHave, false, false -> let skols = List.flatten (List.map (function diff --git a/plugins/ssr/ssrfwd.mli b/plugins/ssr/ssrfwd.mli index 6fb97d524..ead361745 100644 --- a/plugins/ssr/ssrfwd.mli +++ b/plugins/ssr/ssrfwd.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Names open Ltac_plugin diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml index b850b0e95..7ae9e3824 100644 --- a/plugins/ssr/ssripats.ml +++ b/plugins/ssr/ssripats.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Names open Pp open Term @@ -94,7 +95,7 @@ let ssrmkabs id gl = end in Proofview.V82.of_tactic (Proofview.tclTHEN - (Tactics.New.refine step) + (Tactics.New.refine ~typecheck:false step) (Proofview.tclFOCUS 1 3 Proofview.shelve)) gl let ssrmkabstac ids = @@ -116,7 +117,7 @@ let delayed_clear force rest clr gl = let ren_clr, ren = List.split (List.map (fun x -> let x = hyp_id x in - let x' = mk_anon_id (string_of_id x) gl in + let x' = mk_anon_id (Id.to_string x) gl in x', (x, x')) clr) in let ctx = { ctx with delayed_clears = ren_clr @ ctx.delayed_clears } in let gl = push_ctx ctx gl in @@ -132,7 +133,7 @@ let with_defective maintac deps clr ist gl = let top_id = match EConstr.kind_of_type (project gl) (pf_concl gl) with | ProdType (Name id, _, _) - when has_discharged_tag (string_of_id id) -> id + when has_discharged_tag (Id.to_string id) -> id | _ -> top_id in let top_gen = mkclr clr, cpattern_of_id top_id in tclTHEN (introid top_id) (maintac deps top_gen ist) gl @@ -142,7 +143,7 @@ let with_defective_a maintac deps clr ist gl = let top_id = match EConstr.kind_of_type sigma (without_ctx pf_concl gl) with | ProdType (Name id, _, _) - when has_discharged_tag (string_of_id id) -> id + when has_discharged_tag (Id.to_string id) -> id | _ -> top_id in let top_gen = mkclr clr, cpattern_of_id top_id in tclTHEN_a (tac_ctx (introid top_id)) (maintac deps top_gen ist) gl diff --git a/plugins/ssr/ssripats.mli b/plugins/ssr/ssripats.mli index e90e75552..5f5c7f34a 100644 --- a/plugins/ssr/ssripats.mli +++ b/plugins/ssr/ssripats.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Ssrmatching_plugin open Ssrast open Ssrcommon diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4 index 1fba39150..09917339a 100644 --- a/plugins/ssr/ssrparser.ml4 +++ b/plugins/ssr/ssrparser.ml4 @@ -8,6 +8,8 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API +open Grammar_API open Names open Pp open Pcoq @@ -344,7 +346,8 @@ let interp_index ist gl idx = | Some c -> let rc = Detyping.detype false [] (pf_env gl) (project gl) c in begin match Notation.uninterp_prim_token rc with - | _, Constrexpr.Numeral bigi -> int_of_string (Bigint.to_string bigi) + | _, Constrexpr.Numeral (s,b) -> + let n = int_of_string s in if b then n else -n | _ -> raise Not_found end | None -> raise Not_found @@ -1463,7 +1466,7 @@ let ssr_id_of_string loc s = else Feedback.msg_warning (str ( "The name " ^ s ^ " fits the _xxx_ format used for anonymous variables.\n" ^ "Scripts with explicit references to anonymous variables are fragile.")) - end; id_of_string s + end; Id.of_string s let ssr_null_entry = Gram.Entry.of_parser "ssr_null" (fun _ -> ()) @@ -1553,7 +1556,7 @@ END let ssrautoprop gl = try let tacname = - try Nametab.locate_tactic (qualid_of_ident (id_of_string "ssrautoprop")) + try Nametab.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop")) with Not_found -> Nametab.locate_tactic (ssrqid "ssrautoprop") in let tacexpr = Loc.tag @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in Proofview.V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl @@ -2318,7 +2321,7 @@ let test_idcomma = Gram.Entry.of_parser "test_idcomma" accept_idcomma GEXTEND Gram GLOBAL: ssr_idcomma; ssr_idcomma: [ [ test_idcomma; - ip = [ id = IDENT -> Some (id_of_string id) | "_" -> None ]; "," -> + ip = [ id = IDENT -> Some (Id.of_string id) | "_" -> None ]; "," -> Some ip ] ]; END diff --git a/plugins/ssr/ssrparser.mli b/plugins/ssr/ssrparser.mli index bf6f44f11..154820666 100644 --- a/plugins/ssr/ssrparser.mli +++ b/plugins/ssr/ssrparser.mli @@ -8,6 +8,9 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API +open Grammar_API + val ssrtacarg : Tacexpr.raw_tactic_expr Pcoq.Gram.entry val wit_ssrtacarg : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr, Geninterp.Val.t) Genarg.genarg_type val pr_ssrtacarg : 'a -> 'b -> (int * Ppextend.parenRelation -> 'c) -> 'c diff --git a/plugins/ssr/ssrprinters.ml b/plugins/ssr/ssrprinters.ml index e865ef706..427109c1b 100644 --- a/plugins/ssr/ssrprinters.ml +++ b/plugins/ssr/ssrprinters.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Pp open Names open Printer diff --git a/plugins/ssr/ssrprinters.mli b/plugins/ssr/ssrprinters.mli index 56ec145ad..9207b9e43 100644 --- a/plugins/ssr/ssrprinters.mli +++ b/plugins/ssr/ssrprinters.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Ssrast val pp_term : diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml index 0fe8fdc26..b586d05e1 100644 --- a/plugins/ssr/ssrtacticals.ml +++ b/plugins/ssr/ssrtacticals.ml @@ -8,8 +8,8 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Names -open Constr open Termops open Tacmach open Misctypes @@ -102,10 +102,10 @@ let endclausestac id_map clseq gl_id cl0 gl = | ids, dc' -> forced && ids = [] && (not hide_goal || dc' = [] && c_hidden) in let rec unmark c = match EConstr.kind (project gl) c with - | Var id when hidden_clseq clseq && id = gl_id -> cl0 - | Prod (Name id, t, c') when List.mem_assoc id id_map -> + | Term.Var id when hidden_clseq clseq && id = gl_id -> cl0 + | Term.Prod (Name id, t, c') when List.mem_assoc id id_map -> EConstr.mkProd (Name (orig_id id), unmark t, unmark c') - | LetIn (Name id, v, t, c') when List.mem_assoc id id_map -> + | Term.LetIn (Name id, v, t, c') when List.mem_assoc id id_map -> EConstr.mkLetIn (Name (orig_id id), unmark v, unmark t, unmark c') | _ -> EConstr.map (project gl) unmark c in let utac hyp = diff --git a/plugins/ssr/ssrtacticals.mli b/plugins/ssr/ssrtacticals.mli index b8e95b2b1..1d1887138 100644 --- a/plugins/ssr/ssrtacticals.mli +++ b/plugins/ssr/ssrtacticals.mli @@ -8,6 +8,8 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API + val tclSEQAT : Ltac_plugin.Tacinterp.interp_sign -> Ltac_plugin.Tacinterp.Value.t -> diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4 index b154cf217..4c8827bf8 100644 --- a/plugins/ssr/ssrvernac.ml4 +++ b/plugins/ssr/ssrvernac.ml4 @@ -8,6 +8,8 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API +open Grammar_API open Names open Term open Termops @@ -353,7 +355,7 @@ let coerce_search_pattern_to_sort hpat = let coerce hp coe_index = let coe = Classops.get_coercion_value coe_index in try - let coe_ref = reference_of_constr coe in + let coe_ref = global_of_constr coe in let n_imps = Option.get (Classops.hide_coercion coe_ref) in mkPApp (Pattern.PRef coe_ref) n_imps [|hp|] with _ -> diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml index 3c995b1bb..91e40f368 100644 --- a/plugins/ssr/ssrview.ml +++ b/plugins/ssr/ssrview.ml @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Util open Names open Term diff --git a/plugins/ssr/ssrview.mli b/plugins/ssr/ssrview.mli index 6fd906ff4..8a7bd5d6e 100644 --- a/plugins/ssr/ssrview.mli +++ b/plugins/ssr/ssrview.mli @@ -8,6 +8,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API open Ssrast open Ssrcommon diff --git a/plugins/ssr/vo.itarget b/plugins/ssr/vo.itarget deleted file mode 100644 index 99f9f160b..000000000 --- a/plugins/ssr/vo.itarget +++ /dev/null @@ -1,3 +0,0 @@ -ssreflect.vo -ssrfun.vo -ssrbool.vo diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 67e6c7e93..796b6f43e 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -8,6 +8,9 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +open API +open Grammar_API + (* Defining grammar rules with "xx" in it automatically declares keywords too, * we thus save the lexer to restore it at the end of the file *) let frozen_lexer = CLexer.get_keyword_state () ;; @@ -397,7 +400,7 @@ type pattern_class = | KpatLam | KpatRigid | KpatFlex - | KpatProj of constant + | KpatProj of Constant.t type tpattern = { up_k : pattern_class; @@ -418,7 +421,7 @@ let isRigid c = match kind_of_term c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false -let hole_var = mkVar (id_of_string "_") +let hole_var = mkVar (Id.of_string "_") let pr_constr_pat c0 = let rec wipe_evar c = if isEvar c then hole_var else map_constr wipe_evar c in @@ -445,7 +448,7 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = Context.Named.fold_inside abs_dc ~init:([], (put evi.evar_concl)) dc in let m = Evarutil.new_meta () in ise := meta_declare m t !ise; - sigma := Evd.define k (applist (mkMeta m, a)) !sigma; + sigma := Evd.define k (applistc (mkMeta m) a) !sigma; put (existential_value !sigma ex) end | _ -> map_constr put c in @@ -462,7 +465,7 @@ let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = | Const (p,_) -> let np = proj_nparams p in if np = 0 || np > List.length a then KpatConst, f, a else - let a1, a2 = List.chop np a in KpatProj p, applist(f, a1), a2 + let a1, a2 = List.chop np a in KpatProj p, (applistc f a1), a2 | Proj (p,arg) -> KpatProj (Projection.constant p), f, a | Var _ | Ind _ | Construct _ -> KpatFixed, f, a | Evar (k, _) -> @@ -568,7 +571,7 @@ let filter_upat_FO i0 f n u fpats = | KpatFlex -> i0 := n; true in if ok then begin if !i0 < np then i0 := np; (u, np) :: fpats end else fpats -exception FoundUnif of (evar_map * evar_universe_context * tpattern) +exception FoundUnif of (evar_map * UState.t * tpattern) (* Note: we don't update env as we descend into the term, as the primitive *) (* unification procedure always rejects subterms with bound variables. *) @@ -711,7 +714,7 @@ type find_P = k:subst -> constr type conclude = unit -> - constr * ssrdir * (Evd.evar_map * Evd.evar_universe_context * constr) + constr * ssrdir * (Evd.evar_map * UState.t * constr) (* upats_origin makes a better error message only *) let mk_tpattern_matcher ?(all_instances=false) @@ -902,7 +905,7 @@ let glob_cpattern gs p = pp(lazy(str"globbing pattern: " ++ pr_term p)); let glob x = snd (glob_ssrterm gs (mk_lterm x)) in let encode k s l = - let name = Name (id_of_string ("_ssrpat_" ^ s)) in + let name = Name (Id.of_string ("_ssrpat_" ^ s)) in k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in let bind_in t1 t2 = let mkCHole = mkCHole ~loc:None in let n = Name (destCVar t1) in @@ -1128,9 +1131,9 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = sigma in let red = let rec decode_red (ist,red) = let open CAst in match red with | T(k,({ v = GCast ({ v = GHole _ },CastConv({ v = GLambda (Name id,_,_,t)}))},None)) - when let id = string_of_id id in let len = String.length id in + when let id = Id.to_string id in let len = String.length id in (len > 8 && String.sub id 0 8 = "_ssrpat_") -> - let id = string_of_id id in let len = String.length id in + let id = Id.to_string id in let len = String.length id in (match String.sub id 8 (len - 8), t with | "In", { v = GApp( _, [t]) } -> decodeG t xInT (fun x -> T x) | "In", { v = GApp( _, [e; t]) } -> decodeG t (eInXInT (mkG e)) (bad_enc id) @@ -1374,7 +1377,7 @@ let ssrpatterntac _ist (arg_ist,arg) gl = let t = EConstr.of_constr t in let concl_x = EConstr.of_constr concl_x in let gl, tty = pf_type_of gl t in - let concl = EConstr.mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in + let concl = EConstr.mkLetIn (Name (Id.of_string "selected"), t, tty, concl_x) in Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl (* Register "ssrpattern" tactic *) diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 8be989de5..c2bf12cb6 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -1,6 +1,8 @@ (* (c) Copyright 2006-2015 Microsoft Corporation and Inria. *) (* Distributed under the terms of CeCILL-B. *) +open API +open Grammar_API open Genarg open Tacexpr open Environ @@ -152,7 +154,7 @@ type find_P = instantiation, the proof term and the ssrdit stored in the tpattern @raise UserEerror if too many occurrences were specified *) type conclude = - unit -> constr * ssrdir * (evar_map * Evd.evar_universe_context * constr) + unit -> constr * ssrdir * (evar_map * UState.t * constr) (** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair a function [find_P] and [conclude] with the behaviour explained above. @@ -222,12 +224,12 @@ val pf_unify_HO : goal sigma -> EConstr.constr -> EConstr.constr -> goal sigma on top of the former APIs *) val tag_of_cpattern : cpattern -> char val loc_of_cpattern : cpattern -> Loc.t option -val id_of_pattern : pattern -> Names.variable option +val id_of_pattern : pattern -> Names.Id.t option val is_wildcard : cpattern -> bool -val cpattern_of_id : Names.variable -> cpattern +val cpattern_of_id : Names.Id.t -> cpattern val pr_constr_pat : constr -> Pp.std_ppcmds -val pf_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma -val pf_unsafe_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma +val pf_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma +val pf_unsafe_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma (* One can also "Set SsrMatchingDebug" from a .v *) val debug : bool -> unit diff --git a/plugins/ssrmatching/vo.itarget b/plugins/ssrmatching/vo.itarget deleted file mode 100644 index b0eb38834..000000000 --- a/plugins/ssrmatching/vo.itarget +++ /dev/null @@ -1 +0,0 @@ -ssrmatching.vo diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml index e7eea0284..6bf5b8cfc 100644 --- a/plugins/syntax/ascii_syntax.ml +++ b/plugins/syntax/ascii_syntax.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) +open API + (* Poor's man DECLARE PLUGIN *) let __coq_plugin_name = "ascii_syntax_plugin" let () = Mltop.add_known_module __coq_plugin_name 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/nat_syntax.ml b/plugins/syntax/nat_syntax.ml index 9a4cd6c25..a3d13c407 100644 --- a/plugins/syntax/nat_syntax.ml +++ b/plugins/syntax/nat_syntax.ml @@ -6,6 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API + (* Poor's man DECLARE PLUGIN *) let __coq_plugin_name = "nat_syntax_plugin" let () = Mltop.add_known_module __coq_plugin_name diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml deleted file mode 100644 index e23852bf8..000000000 --- a/plugins/syntax/numbers_syntax.ml +++ /dev/null @@ -1,311 +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 *) -(************************************************************************) - -(* 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 (MPfile (make_dir dir)) id -let make_mind_mpdot dir modname id = - let mp = MPdot (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/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml index 7ce066c59..a73468123 100644 --- a/plugins/syntax/r_syntax.ml +++ b/plugins/syntax/r_syntax.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Util open Names open Globnames diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml index b7f13b040..a4335a508 100644 --- a/plugins/syntax/string_syntax.ml +++ b/plugins/syntax/string_syntax.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) +open API open Globnames open Ascii_syntax_plugin.Ascii_syntax open Glob_term diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml index 479448e06..dfff8d9df 100644 --- a/plugins/syntax/z_syntax.ml +++ b/plugins/syntax/z_syntax.ml @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open API open Pp open CErrors open Util diff --git a/pretyping/cases.ml b/pretyping/cases.ml index c3f392980..b88532e1b 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -245,6 +245,7 @@ let push_history_pattern n pci cont = type 'a pattern_matching_problem = { env : env; + lvar : Glob_term.ltac_var_map; evdref : evar_map ref; pred : constr; tomatch : tomatch_stack; @@ -346,25 +347,45 @@ let find_tomatch_tycon evdref env loc = function | None -> empty_tycon,None -let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = +let make_return_predicate_ltac_lvar sigma na tm c lvar = + match na, tm.CAst.v with + | Name id, (GVar id' | GRef (Globnames.VarRef id', _)) when Id.equal id id' -> + if Id.Map.mem id lvar.ltac_genargs then + let ltac_genargs = Id.Map.remove id lvar.ltac_genargs in + let ltac_idents = match kind sigma c with + | Var id' -> Id.Map.add id id' lvar.ltac_idents + | _ -> lvar.ltac_idents in + { lvar with ltac_genargs; ltac_idents } + else lvar + | _ -> lvar + +let ltac_interp_realnames lvar = function + | t, IsInd (ty,ind,realnal) -> t, IsInd (ty,ind,List.map (ltac_interp_name lvar) realnal) + | _ as x -> x + +let coerce_row typing_fun evdref env lvar pats (tomatch,(na,indopt)) = let loc = loc_of_glob_constr tomatch in let tycon,realnames = find_tomatch_tycon evdref env loc indopt in - let j = typing_fun tycon env evdref tomatch in + let j = typing_fun tycon env evdref !lvar tomatch in let evd, j = Coercion.inh_coerce_to_base ?loc:(loc_of_glob_constr tomatch) env !evdref j in evdref := evd; let typ = nf_evar !evdref j.uj_type in + lvar := make_return_predicate_ltac_lvar !evdref na tomatch j.uj_val !lvar; let t = try try_find_ind env !evdref typ realnames with Not_found -> unify_tomatch_with_patterns evdref env loc typ pats realnames in (j.uj_val,t) -let coerce_to_indtype typing_fun evdref env matx tomatchl = +let coerce_to_indtype typing_fun evdref env lvar matx tomatchl = let pats = List.map (fun r -> r.patterns) matx in let matx' = match matrix_transpose pats with | [] -> List.map (fun _ -> []) tomatchl (* no patterns at all *) | m -> m in - List.map2 (coerce_row typing_fun evdref env) matx' tomatchl + let lvar = ref lvar in + let tms = List.map2 (coerce_row typing_fun evdref env lvar) matx' tomatchl in + let tms = List.map (ltac_interp_realnames !lvar) tms in + !lvar,tms (************************************************************************) (* Utils *) @@ -1392,6 +1413,7 @@ and match_current pb (initial,tomatch) = postprocess_dependencies !(pb.evdref) depstocheck brvals pb.tomatch pb.pred deps cstrs in let brvals = Array.map (fun (sign,body) -> + let sign = List.map (map_name (ltac_interp_name pb.lvar)) sign in it_mkLambda_or_LetIn body sign) brvals in let (pred,typ) = find_predicate pb.caseloc pb.env pb.evdref @@ -1824,6 +1846,7 @@ let build_inversion_problem loc env sigma tms t = let evdref = ref sigma in let pb = { env = pb_env; + lvar = empty_lvar; evdref = evdref; pred = (*ty *) mkSort s; tomatch = sub_tms; @@ -1847,15 +1870,15 @@ let build_initial_predicate arsign pred = | _ -> assert false in buildrec 0 pred [] (List.rev arsign) -let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = +let extract_arity_signature ?(dolift=true) env0 lvar tomatchl tmsign = let lift = if dolift then lift else fun n t -> t in let get_one_sign n tm (na,t) = match tm with | NotInd (bo,typ) -> (match t with - | None -> (match bo with + | None -> let sign = match bo with | None -> [LocalAssum (na, lift n typ)] - | Some b -> [LocalDef (na, lift n b, lift n typ)]) + | Some b -> [LocalDef (na, lift n b, lift n typ)] in sign,sign | Some (loc,_) -> user_err ?loc (str"Unexpected type annotation for a term of non inductive type.")) @@ -1865,22 +1888,31 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = let nrealargs_ctxt = inductive_nrealdecls_env env0 ind in let arsign = fst (get_arity env0 indf') in let arsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) arsign in - let realnal = + let realnal, realnal' = match t with | Some (loc,(ind',realnal)) -> if not (eq_ind ind ind') then user_err ?loc (str "Wrong inductive type."); if not (Int.equal nrealargs_ctxt (List.length realnal)) then anomaly (Pp.str "Ill-formed 'in' clause in cases."); - List.rev realnal - | None -> List.make nrealargs_ctxt Anonymous in - LocalAssum (na, EConstr.of_constr (build_dependent_inductive env0 indf')) - ::(List.map2 RelDecl.set_name realnal arsign) in + let realnal = List.rev realnal in + let realnal' = List.map (ltac_interp_name lvar) realnal in + realnal,realnal' + | None -> + let realnal = List.make nrealargs_ctxt Anonymous in + realnal, realnal in + let na' = ltac_interp_name lvar na in + let t = EConstr.of_constr (build_dependent_inductive env0 indf') in + (* Context with names for typing *) + let arsign1 = LocalAssum (na, t) :: List.map2 RelDecl.set_name realnal arsign in + (* Context with names for building the term *) + let arsign2 = LocalAssum (na', t) :: List.map2 RelDecl.set_name realnal' arsign in + arsign1,arsign2 in let rec buildrec n = function | [],[] -> [] | (_,tm)::ltm, (_,x)::tmsign -> let l = get_one_sign n tm x in - l :: buildrec (n + List.length l) (ltm,tmsign) + l :: buildrec (n + List.length (fst l)) (ltm,tmsign) | _ -> assert false in List.rev (buildrec 0 (tomatchl,tmsign)) @@ -1970,7 +2002,7 @@ let noccur_with_meta sigma n m term = in try (occur_rec n term; true) with LocalOccur -> false -let prepare_predicate ?loc typing_fun env sigma tomatchs arsign tycon pred = +let prepare_predicate ?loc typing_fun env sigma lvar tomatchs arsign tycon pred = let refresh_tycon sigma t = (** If we put the typing constraint in the term, it has to be refreshed to preserve the invariant that no algebraic universe @@ -1978,6 +2010,7 @@ let prepare_predicate ?loc typing_fun env sigma tomatchs arsign tycon pred = refresh_universes ~status:Evd.univ_flexible ~onlyalg:true (Some true) env sigma t in + let typing_arsign,building_arsign = List.split arsign in let preds = match pred, tycon with (* No return clause *) @@ -1987,7 +2020,7 @@ let prepare_predicate ?loc typing_fun env sigma tomatchs arsign tycon pred = (* First strategy: we abstract the tycon wrt to the dependencies *) let sigma, t = refresh_tycon sigma t in let p1 = - prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign t in + prepare_predicate_from_arsign_tycon env sigma loc tomatchs typing_arsign t in (* Second strategy: we build an "inversion" predicate *) let sigma2,pred2 = build_inversion_problem loc env sigma tomatchs t in (match p1 with @@ -2006,22 +2039,22 @@ let prepare_predicate ?loc typing_fun env sigma tomatchs arsign tycon pred = (* First strategy: we build an "inversion" predicate *) let sigma1,pred1 = build_inversion_problem loc env sigma tomatchs t in (* Second strategy: we directly use the evar as a non dependent pred *) - let pred2 = lift (List.length (List.flatten arsign)) t in + let pred2 = lift (List.length (List.flatten typing_arsign)) t in [sigma1, pred1; sigma, pred2] (* Some type annotation *) | Some rtntyp, _ -> (* We extract the signature of the arity *) - let envar = List.fold_right push_rel_context arsign env in + let envar = List.fold_right push_rel_context typing_arsign env in let sigma, newt = new_sort_variable univ_flexible_alg sigma in let evdref = ref sigma in - let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref rtntyp in + let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref lvar rtntyp in let sigma = !evdref in let predccl = nf_evar sigma predcclj.uj_val in [sigma, predccl] in List.map (fun (sigma,pred) -> - let (nal,pred) = build_initial_predicate arsign pred in + let (nal,pred) = build_initial_predicate building_arsign pred in sigma,nal,pred) preds @@ -2441,10 +2474,10 @@ let context_of_arsign l = l ([], 0) in x -let compile_program_cases ?loc style (typing_function, evdref) tycon env +let compile_program_cases ?loc style (typing_function, evdref) tycon env lvar (predopt, tomatchl, eqns) = let typing_fun tycon env = function - | Some t -> typing_function tycon env evdref t + | Some t -> typing_function tycon env evdref lvar t | None -> Evarutil.evd_comb0 use_unit_judge evdref in (* We build the matrix of patterns and right-hand side *) @@ -2452,14 +2485,15 @@ let compile_program_cases ?loc style (typing_function, evdref) tycon env (* We build the vector of terms to match consistently with the *) (* constructors found in patterns *) - let tomatchs = coerce_to_indtype typing_function evdref env matx tomatchl in + let predlvar,tomatchs = coerce_to_indtype typing_function evdref env lvar matx tomatchl in let tycon = valcon_of_tycon tycon in let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env !evdref tomatchs tycon in let env = push_rel_context tomatchs_lets env in let len = List.length eqns in let sign, allnames, signlen, eqs, neqs, args = (* The arity signature *) - let arsign = extract_arity_signature ~dolift:false env tomatchs tomatchl in + let arsign = extract_arity_signature ~dolift:false env predlvar tomatchs tomatchl in + let arsign = List.map fst arsign in (* Because no difference between the arity for typing and the arity for building *) (* Build the dependent arity signature, the equalities which makes the first part of the predicate and their instantiations. *) let avoid = [] in @@ -2522,11 +2556,12 @@ let compile_program_cases ?loc style (typing_function, evdref) tycon env let initial_pushed = List.map (fun x -> Pushed (true,x)) typs' in let typing_function tycon env evdref = function - | Some t -> typing_function tycon env evdref t + | Some t -> typing_function tycon env evdref lvar t | None -> evd_comb0 use_unit_judge evdref in let pb = { env = env; + lvar = lvar; evdref = evdref; pred = pred; tomatch = initial_pushed; @@ -2548,10 +2583,10 @@ let compile_program_cases ?loc style (typing_function, evdref) tycon env (**************************************************************************) (* Main entry of the matching compilation *) -let compile_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) = +let compile_cases ?loc style (typing_fun, evdref) tycon env lvar (predopt, tomatchl, eqns) = if predopt == None && Flags.is_program_mode () && Program.is_program_cases () then compile_program_cases ?loc style (typing_fun, evdref) - tycon env (predopt, tomatchl, eqns) + tycon env lvar (predopt, tomatchl, eqns) else (* We build the matrix of patterns and right-hand side *) @@ -2559,15 +2594,15 @@ let compile_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, (* We build the vector of terms to match consistently with the *) (* constructors found in patterns *) - let tomatchs = coerce_to_indtype typing_fun evdref env matx tomatchl in + let predlvar,tomatchs = coerce_to_indtype typing_fun evdref env lvar matx tomatchl in (* If an elimination predicate is provided, we check it is compatible with the type of arguments to match; if none is provided, we build alternative possible predicates *) - let arsign = extract_arity_signature env tomatchs tomatchl in - let preds = prepare_predicate ?loc typing_fun env !evdref tomatchs arsign tycon predopt in + let arsign = extract_arity_signature env predlvar tomatchs tomatchl in + let preds = prepare_predicate ?loc typing_fun env !evdref predlvar tomatchs arsign tycon predopt in let compile_for_one_predicate (sigma,nal,pred) = (* We push the initial terms to match and push their alias to rhs' envs *) @@ -2598,13 +2633,14 @@ let compile_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, (* A typing function that provides with a canonical term for absurd cases*) let typing_fun tycon env evdref = function - | Some t -> typing_fun tycon env evdref t + | Some t -> typing_fun tycon env evdref lvar t | None -> evd_comb0 use_unit_judge evdref in let myevdref = ref sigma in let pb = { env = env; + lvar = lvar; evdref = myevdref; pred = pred; tomatch = initial_pushed; diff --git a/pretyping/cases.mli b/pretyping/cases.mli index b16342db4..4b1fde25a 100644 --- a/pretyping/cases.mli +++ b/pretyping/cases.mli @@ -39,9 +39,9 @@ val irrefutable : env -> cases_pattern -> bool val compile_cases : ?loc:Loc.t -> case_style -> - (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref -> + (type_constraint -> env -> evar_map ref -> ltac_var_map -> glob_constr -> unsafe_judgment) * evar_map ref -> type_constraint -> - env -> glob_constr option * tomatch_tuples * cases_clauses -> + env -> ltac_var_map -> glob_constr option * tomatch_tuples * cases_clauses -> unsafe_judgment val constr_of_pat : @@ -101,6 +101,7 @@ and pattern_continuation = type 'a pattern_matching_problem = { env : env; + lvar : Glob_term.ltac_var_map; evdref : evar_map ref; pred : constr; tomatch : tomatch_stack; @@ -115,10 +116,14 @@ val compile : 'a pattern_matching_problem -> unsafe_judgment val prepare_predicate : ?loc:Loc.t -> (Evarutil.type_constraint -> - Environ.env -> Evd.evar_map ref -> glob_constr -> unsafe_judgment) -> + Environ.env -> Evd.evar_map ref -> ltac_var_map -> glob_constr -> unsafe_judgment) -> Environ.env -> Evd.evar_map -> + Glob_term.ltac_var_map -> (types * tomatch_type) list -> - rel_context list -> + (rel_context * rel_context) list -> constr option -> glob_constr option -> (Evd.evar_map * Names.name list * constr) list + +val make_return_predicate_ltac_lvar : Evd.evar_map -> Names.name -> + Glob_term.glob_constr -> constr -> Glob_term.ltac_var_map -> Glob_term.ltac_var_map diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index 782552583..6c2086f3e 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -132,6 +132,7 @@ let mkSTACK = function | STACK(0,v0,stk0), stk -> STACK(0,v0,stack_concat stk0 stk) | v,stk -> STACK(0,v,stk) +type cbv_infos = { infos : cbv_value infos; sigma : Evd.evar_map } (* Change: zeta reduction cannot be avoided in CBV *) @@ -189,6 +190,43 @@ let pr_key = function | VarKey id -> Names.Id.print id | RelKey n -> Pp.(str "REL_" ++ int n) +let rec reify_stack t = function + | TOP -> t + | APP (args,st) -> + reify_stack (mkApp(t,Array.map reify_value args)) st + | CASE (ty,br,ci,env,st) -> + reify_stack + (mkCase (ci, ty, t,br)) + st + | PROJ (p, pinfo, st) -> + reify_stack (mkProj (p, t)) st + +and reify_value = function (* reduction under binders *) + | VAL (n,t) -> lift n t + | STACK (0,v,stk) -> + reify_stack (reify_value v) stk + | STACK (n,v,stk) -> + lift n (reify_stack (reify_value v) stk) + | CBN(t,env) -> + t + (* map_constr_with_binders subs_lift (cbv_norm_term) env t *) + | LAM (n,ctxt,b,env) -> + List.fold_left (fun c (n,t) -> Term.mkLambda (n, t, c)) b ctxt + | FIXP ((lij,(names,lty,bds)),env,args) -> + mkApp + (mkFix (lij, + (names, + lty, + bds)), + Array.map reify_value args) + | COFIXP ((j,(names,lty,bds)),env,args) -> + mkApp + (mkCoFix (j, + (names,lty,bds)), + Array.map reify_value args) + | CONSTR (c,args) -> + mkApp(mkConstructU c, Array.map reify_value args) + (* The main recursive functions * * Go under applications and cases/projections (pushed in the stack), @@ -213,12 +251,12 @@ let rec norm_head info env t stack = | Proj (p, c) -> let p' = - if red_set (info_flags info) (fCONST (Projection.constant p)) - && red_set (info_flags info) fBETA + if red_set (info_flags info.infos) (fCONST (Projection.constant p)) + && red_set (info_flags info.infos) fBETA then Projection.unfold p else p in - let pinfo = Environ.lookup_projection p (info_env info) in + let pinfo = Environ.lookup_projection p (info_env info.infos) in norm_head info env c (PROJ (p', pinfo, stack)) (* constants, axioms @@ -233,14 +271,16 @@ let rec norm_head info env t stack = | Var id -> norm_head_ref 0 info env stack (VarKey id) - | Const sp -> norm_head_ref 0 info env stack (ConstKey sp) + | Const sp -> + Reductionops.reduction_effect_hook (env_of_infos info.infos) info.sigma t (lazy (reify_stack t stack)); + norm_head_ref 0 info env stack (ConstKey sp) | LetIn (_, b, _, c) -> (* zeta means letin are contracted; delta without zeta means we *) (* allow bindings but leave let's in place *) - if red_set (info_flags info) fZETA then + if red_set (info_flags info.infos) fZETA then (* New rule: for Cbv, Delta does not apply to locally bound variables - or red_set (info_flags info) fDELTA + or red_set (info_flags info.infos) fDELTA *) let env' = subs_cons ([|cbv_stack_term info TOP env b|],env) in norm_head info env' c stack @@ -248,7 +288,7 @@ let rec norm_head info env t stack = (CBN(t,env), stack) (* Should we consider a commutative cut ? *) | Evar ev -> - (match evar_value info.i_cache ev with + (match evar_value info.infos.i_cache ev with Some c -> norm_head info env c stack | None -> (VAL(0, t), stack)) @@ -265,8 +305,8 @@ let rec norm_head info env t stack = | Prod _ -> (CBN(t,env), stack) and norm_head_ref k info env stack normt = - if red_set_ref (info_flags info) normt then - match ref_value_cache info normt with + if red_set_ref (info_flags info.infos) normt then + match ref_value_cache info.infos normt with | Some body -> if !debug_cbv then Feedback.msg_debug Pp.(str "Unfolding " ++ pr_key normt); strip_appl (shift_value k body) stack @@ -291,7 +331,7 @@ and cbv_stack_term info stack env t = and cbv_stack_value info env = function (* a lambda meets an application -> BETA *) | (LAM (nlams,ctxt,b,env), APP (args, stk)) - when red_set (info_flags info) fBETA -> + when red_set (info_flags info.infos) fBETA -> let nargs = Array.length args in if nargs == nlams then cbv_stack_term info stk (subs_cons(args,env)) b @@ -305,31 +345,31 @@ and cbv_stack_value info env = function (* a Fix applied enough -> IOTA *) | (FIXP(fix,env,[||]), stk) - when fixp_reducible (info_flags info) fix stk -> + when fixp_reducible (info_flags info.infos) fix stk -> let (envf,redfix) = contract_fixp env fix in cbv_stack_term info stk envf redfix (* constructor guard satisfied or Cofix in a Case -> IOTA *) | (COFIXP(cofix,env,[||]), stk) - when cofixp_reducible (info_flags info) cofix stk-> + when cofixp_reducible (info_flags info.infos) cofix stk-> let (envf,redfix) = contract_cofixp env cofix in cbv_stack_term info stk envf redfix (* constructor in a Case -> IOTA *) | (CONSTR(((sp,n),u),[||]), APP(args,CASE(_,br,ci,env,stk))) - when red_set (info_flags info) fMATCH -> + when red_set (info_flags info.infos) fMATCH -> let cargs = Array.sub args ci.ci_npar (Array.length args - ci.ci_npar) in cbv_stack_term info (stack_app cargs stk) env br.(n-1) (* constructor of arity 0 in a Case -> IOTA *) | (CONSTR(((_,n),u),[||]), CASE(_,br,_,env,stk)) - when red_set (info_flags info) fMATCH -> + when red_set (info_flags info.infos) fMATCH -> cbv_stack_term info stk env br.(n-1) (* constructor in a Projection -> IOTA *) | (CONSTR(((sp,n),u),[||]), APP(args,PROJ(p,pi,stk))) - when red_set (info_flags info) fMATCH && Projection.unfolded p -> + when red_set (info_flags info.infos) fMATCH && Projection.unfolded p -> let arg = args.(pi.Declarations.proj_npars + pi.Declarations.proj_arg) in cbv_stack_value info env (strip_appl arg stk) @@ -337,7 +377,7 @@ and cbv_stack_value info env = function | (FIXP(fix,env,[||]), APP(appl,TOP)) -> FIXP(fix,env,appl) | (COFIXP(cofix,env,[||]), APP(appl,TOP)) -> COFIXP(cofix,env,appl) | (CONSTR(c,[||]), APP(appl,TOP)) -> CONSTR(c,appl) - + (* definitely a value *) | (head,stk) -> mkSTACK(head, stk) @@ -400,12 +440,11 @@ let cbv_norm infos constr = let constr = EConstr.Unsafe.to_constr constr in EConstr.of_constr (with_stats (lazy (cbv_norm_term infos (subs_id 0) constr))) -type cbv_infos = cbv_value infos - (* constant bodies are normalized at the first expansion *) let create_cbv_infos flgs env sigma = - create - (fun old_info c -> cbv_stack_term old_info TOP (subs_id 0) c) + let infos = create + (fun old_info c -> cbv_stack_term { infos = old_info; sigma } TOP (subs_id 0) c) flgs env - (Reductionops.safe_evar_value sigma) + (Reductionops.safe_evar_value sigma) in + { infos; sigma } diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 9a973cff5..8d87f6e99 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -428,7 +428,7 @@ let automatically_import_coercions = ref false open Goptions let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "automatic import of coercions"; optkey = ["Automatic";"Coercions";"Import"]; optread = (fun () -> !automatically_import_coercions); @@ -454,15 +454,11 @@ let cache_coercion (_, c) = add_coercion_in_graph (xf,is,it) let load_coercion _ o = - if - !automatically_import_coercions || Flags.version_less_or_equal Flags.V8_2 - then + if !automatically_import_coercions then cache_coercion o let open_coercion i o = - if Int.equal i 1 && not - (!automatically_import_coercions || Flags.version_less_or_equal Flags.V8_2) - then + if Int.equal i 1 && not !automatically_import_coercions then cache_coercion o let subst_coercion (subst, c) = diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index e53d19b59..9c09396cc 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -457,11 +457,44 @@ let rec cases_pattern_of_glob_constr na = CAst.map (function | _ -> raise Not_found ) +open Declarations +open Term +open Context + +(* Keep only patterns which are not bound to a local definitions *) +let drop_local_defs typi args = + let (decls,_) = decompose_prod_assum typi in + let rec aux decls args = + match decls, args with + | [], [] -> [] + | Rel.Declaration.LocalDef _ :: decls, pat :: args -> + begin + match pat.CAst.v with + | PatVar Anonymous -> aux decls args + | _ -> raise Not_found (* The pattern is used, one cannot drop it *) + end + | Rel.Declaration.LocalAssum _ :: decls, a :: args -> a :: aux decls args + | _ -> assert false in + aux (List.rev decls) args + +let add_patterns_for_params_remove_local_defs (ind,j) l = + let (mib,mip) = Global.lookup_inductive ind in + let nparams = mib.Declarations.mind_nparams in + let l = + if mip.mind_consnrealdecls.(j-1) = mip.mind_consnrealargs.(j-1) then + (* Optimisation *) l + else + let typi = mip.mind_nf_lc.(j-1) in + let (_,typi) = decompose_prod_n_assum (Rel.length mib.mind_params_ctxt) typi in + drop_local_defs typi l in + Util.List.addn nparams (CAst.make @@ PatVar Anonymous) l + (* Turn a closed cases pattern into a glob_constr *) let rec glob_constr_of_closed_cases_pattern_aux x = CAst.map_with_loc (fun ?loc -> function | PatCstr (cstr,[],Anonymous) -> GRef (ConstructRef cstr,None) | PatCstr (cstr,l,Anonymous) -> let ref = CAst.make ?loc @@ GRef (ConstructRef cstr,None) in + let l = add_patterns_for_params_remove_local_defs cstr l in GApp (ref, List.map glob_constr_of_closed_cases_pattern_aux l) | _ -> raise Not_found ) x @@ -471,3 +504,27 @@ let glob_constr_of_closed_cases_pattern = function na,glob_constr_of_closed_cases_pattern_aux (CAst.make ?loc @@ PatCstr (cstr,l,Anonymous)) | _ -> raise Not_found + +(**********************************************************************) +(* Interpreting ltac variables *) + +open Pp +open CErrors + +let ltac_interp_name { ltac_idents ; ltac_genargs } = function + | Anonymous -> Anonymous + | Name id as n -> + try Name (Id.Map.find id ltac_idents) + with Not_found -> + if Id.Map.mem id ltac_genargs then + user_err (str"Ltac variable"++spc()++ pr_id id ++ + spc()++str"is not bound to an identifier."++spc()++ + str"It cannot be used in a binder.") + else n + +let empty_lvar : ltac_var_map = { + ltac_constrs = Id.Map.empty; + ltac_uconstrs = Id.Map.empty; + ltac_idents = Id.Map.empty; + ltac_genargs = Id.Map.empty; +} diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index f7cc08ca2..6bb421e73 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -81,3 +81,8 @@ val map_pattern : (glob_constr -> glob_constr) -> val cases_pattern_of_glob_constr : Name.t -> glob_constr -> cases_pattern val glob_constr_of_closed_cases_pattern : cases_pattern -> Name.t * glob_constr + +val add_patterns_for_params_remove_local_defs : constructor -> cases_pattern list -> cases_pattern list + +val ltac_interp_name : Glob_term.ltac_var_map -> Names.name -> Names.name +val empty_lvar : Glob_term.ltac_var_map diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index db2e5da95..c36542aeb 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -364,9 +364,9 @@ let rec pat_of_raw metas vars = CAst.with_loc_val (fun ?loc -> function PIf (pat_of_raw metas vars c, pat_of_raw metas vars b1,pat_of_raw metas vars b2) | GLetTuple (nal,(_,None),b,c) -> - let mkGLambda c na = CAst.make ?loc @@ + let mkGLambda na c = CAst.make ?loc @@ GLambda (na,Explicit, CAst.make @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in - let c = List.fold_left mkGLambda c nal in + let c = List.fold_right mkGLambda nal c in let cip = { cip_style = LetStyle; cip_ind = None; diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 92e728683..7488f35bf 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -42,21 +42,11 @@ open Pretype_errors open Glob_term open Glob_ops open Evarconv -open Pattern open Misctypes module NamedDecl = Context.Named.Declaration type typing_constraint = OfType of types | IsType | WithoutTypeConstraint -type var_map = constr_under_binders Id.Map.t -type uconstr_var_map = Glob_term.closed_glob_constr Id.Map.t -type unbound_ltac_var_map = Geninterp.Val.t Id.Map.t -type ltac_var_map = { - ltac_constrs : var_map; - ltac_uconstrs : uconstr_var_map; - ltac_idents: Id.t Id.Map.t; - ltac_genargs : unbound_ltac_var_map; -} type glob_constr_ltac_closure = ltac_var_map * glob_constr type pure_open_constr = evar_map * EConstr.constr @@ -419,17 +409,6 @@ let orelse_name name name' = match name with | Anonymous -> name' | _ -> name -let ltac_interp_name { ltac_idents ; ltac_genargs } = function - | Anonymous -> Anonymous - | Name id as n -> - try Name (Id.Map.find id ltac_idents) - with Not_found -> - if Id.Map.mem id ltac_genargs then - user_err (str"Ltac variable"++spc()++ pr_id id ++ - spc()++str"is not bound to an identifier."++spc()++ - str"It cannot be used in a binder.") - else n - let ltac_interp_name_env k0 lvar env sigma = (* envhd is the initial part of the env when pretype was called first *) (* (in practice is is probably 0, but we have to grant the @@ -943,16 +922,20 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre List.map (set_name Anonymous) arsgn else arsgn in - let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in + let indt = build_dependent_inductive env.ExtraEnv.env indf in + let psign = LocalAssum (na, indt) :: arsgn in (* For locating names in [po] *) + let predlvar = Cases.make_return_predicate_ltac_lvar !evdref na c cj.uj_val lvar in + let psign' = LocalAssum (ltac_interp_name predlvar na, indt) :: arsgn in + let psign' = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign' in + let psign' = Namegen.name_context env.ExtraEnv.env !evdref psign' in (* For naming abstractions in [po] *) let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in let nar = List.length arsgn in (match po with | Some p -> let env_p = push_rel_context !evdref psign env in - let pj = pretype_type empty_valcon env_p evdref lvar p in + let pj = pretype_type empty_valcon env_p evdref predlvar p in let ccl = nf_evar !evdref pj.utj_val in - let psign = make_arity_signature env.ExtraEnv.env !evdref true indf in (* with names *) - let p = it_mkLambda_or_LetIn ccl psign in + let p = it_mkLambda_or_LetIn ccl psign' in let inst = (Array.map_to_list EConstr.of_constr cs.cs_concl_realargs) @[EConstr.of_constr (build_dependent_constructor cs)] in @@ -968,7 +951,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | None -> let tycon = lift_tycon cs.cs_nargs tycon in - let fj = pretype tycon env_f evdref lvar d in + let fj = pretype tycon env_f evdref predlvar d in let ccl = nf_evar !evdref fj.uj_type in let ccl = if noccur_between !evdref 1 cs.cs_nargs ccl then @@ -977,7 +960,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre error_cant_find_case_type ?loc env.ExtraEnv.env !evdref cj.uj_val in (* let ccl = refresh_universes ccl in *) - let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in + let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign' in let v = let ind,_ = dest_ind_family indf in Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p; @@ -1004,14 +987,19 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre else arsgn in let nar = List.length arsgn in - let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in + let indt = build_dependent_inductive env.ExtraEnv.env indf in + let psign = LocalAssum (na, indt) :: arsgn in (* For locating names in [po] *) + let predlvar = Cases.make_return_predicate_ltac_lvar !evdref na c cj.uj_val lvar in + let psign' = LocalAssum (ltac_interp_name predlvar na, indt) :: arsgn in + let psign' = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign' in + let psign' = Namegen.name_context env.ExtraEnv.env !evdref psign' in (* For naming abstractions in [po] *) let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in let pred,p = match po with | Some p -> let env_p = push_rel_context !evdref psign env in - let pj = pretype_type empty_valcon env_p evdref lvar p in + let pj = pretype_type empty_valcon env_p evdref predlvar p in let ccl = nf_evar !evdref pj.utj_val in - let pred = it_mkLambda_or_LetIn ccl psign in + let pred = it_mkLambda_or_LetIn ccl psign' in let typ = lift (- nar) (beta_applist !evdref (pred,[cj.uj_val])) in pred, typ | None -> @@ -1021,7 +1009,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let env = ltac_interp_name_env k0 lvar env !evdref in new_type_evar env evdref loc in - it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in + it_mkLambda_or_LetIn (lift (nar+1) p) psign', p in let pred = nf_evar !evdref pred in let p = nf_evar !evdref p in let f cs b = @@ -1054,8 +1042,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | GCases (sty,po,tml,eqns) -> Cases.compile_cases ?loc sty - ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref) - tycon env.ExtraEnv.env (* loc *) (po,tml,eqns) + ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref),evdref) + tycon env.ExtraEnv.env (* loc *) lvar (po,tml,eqns) | GCast (c,k) -> let cj = @@ -1198,13 +1186,6 @@ let no_classes_no_fail_inference_flags = { let all_and_fail_flags = default_inference_flags true let all_no_fail_flags = default_inference_flags false -let empty_lvar : ltac_var_map = { - ltac_constrs = Id.Map.empty; - ltac_uconstrs = Id.Map.empty; - ltac_idents = Id.Map.empty; - ltac_genargs = Id.Map.empty; -} - let on_judgment sigma f j = let c = mkCast(j.uj_val,DEFAULTcast, j.uj_type) in let (c,_,t) = destCast sigma (f c) in diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index dcacd0720..e17468ef8 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -12,7 +12,6 @@ into elementary ones, insertion of coercions and resolution of implicit arguments. *) -open Names open Term open Environ open Evd @@ -28,23 +27,6 @@ val search_guard : type typing_constraint = OfType of types | IsType | WithoutTypeConstraint -type var_map = Pattern.constr_under_binders Id.Map.t -type uconstr_var_map = Glob_term.closed_glob_constr Id.Map.t -type unbound_ltac_var_map = Geninterp.Val.t Id.Map.t - -type ltac_var_map = { - ltac_constrs : var_map; - (** Ltac variables bound to constrs *) - ltac_uconstrs : uconstr_var_map; - (** Ltac variables bound to untyped constrs *) - ltac_idents: Id.t Id.Map.t; - (** Ltac variables bound to identifiers *) - ltac_genargs : unbound_ltac_var_map; - (** Ltac variables bound to other kinds of arguments *) -} - -val empty_lvar : ltac_var_map - type glob_constr_ltac_closure = ltac_var_map * glob_constr type pure_open_constr = evar_map * constr diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index b4654bfb5..c2a648301 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -40,6 +40,54 @@ let _ = Goptions.declare_bool_option { let get_refolding_in_reduction () = !refolding_in_reduction let set_refolding_in_reduction = (:=) refolding_in_reduction +(** Support for reduction effects *) + +open Mod_subst +open Libobject + +type effect_name = string + +(** create a persistent set to store effect functions *) +module ConstrMap = Map.Make (Constr) + +(* Table bindings a constant to an effect *) +let constant_effect_table = Summary.ref ~name:"reduction-side-effect" ConstrMap.empty + +(* Table bindings function key to effective functions *) +let effect_table = Summary.ref ~name:"reduction-function-effect" String.Map.empty + +(** a test to know whether a constant is actually the effect function *) +let reduction_effect_hook env sigma termkey c = + try + let funkey = ConstrMap.find termkey !constant_effect_table in + let effect = String.Map.find funkey !effect_table in + effect env sigma (Lazy.force c) + with Not_found -> () + +let cache_reduction_effect (_,(termkey,funkey)) = + constant_effect_table := ConstrMap.add termkey funkey !constant_effect_table + +let subst_reduction_effect (subst,(termkey,funkey)) = + (subst_mps subst termkey,funkey) + +let inReductionEffect : Constr.constr * string -> obj = + declare_object {(default_object "REDUCTION-EFFECT") with + cache_function = cache_reduction_effect; + open_function = (fun i o -> if Int.equal i 1 then cache_reduction_effect o); + subst_function = subst_reduction_effect; + classify_function = (fun o -> Substitute o) } + +let declare_reduction_effect funkey f = + if String.Map.mem funkey !effect_table then + CErrors.anomaly Pp.(str "Cannot redeclare effect function " ++ qstring funkey ++ str "."); + effect_table := String.Map.add funkey f !effect_table + +(** A function to set the value of the print function *) +let set_reduction_effect x funkey = + let termkey = Universes.constr_of_global x in + Lib.add_anonymous_leaf (inReductionEffect (termkey,funkey)) + + (** Machinery to custom the behavior of the reduction *) module ReductionBehaviour = struct open Globnames @@ -777,7 +825,7 @@ let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies context" in contract_fix *) let reduce_and_refold_fix recfun env sigma refold cst_l fix sk = let raw_answer = - let env = if refold then None else Some env in + let env = if refold then Some env else None in contract_fix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) fix in apply_subst (fun sigma x (t,sk') -> @@ -859,9 +907,12 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = (match safe_meta_value sigma ev with | Some body -> whrec cst_l (EConstr.of_constr body, stack) | None -> fold ()) - | Const (c,u as const) when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) -> - let u' = EInstance.kind sigma u in - (match constant_opt_value_in env (fst const, u') with + | Const (c,u as const) -> + reduction_effect_hook env sigma (EConstr.to_constr sigma x) + (lazy (EConstr.to_constr sigma (Stack.zip sigma (x,stack)))); + if CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) then + let u' = EInstance.kind sigma u in + (match constant_opt_value_in env (c, u') with | None -> fold () | Some body -> let body = EConstr.of_constr body in @@ -901,7 +952,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Some (bef,arg,s') -> whrec Cst_stack.empty (arg,Stack.Cst(Stack.Cst_const (fst const, u'),curr,remains,bef,cst_l)::s') - ) + ) else fold () | Proj (p, c) when CClosure.RedFlags.red_projection flags p -> (let pb = lookup_projection p env in let kn = Projection.constant p in @@ -1035,7 +1086,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = |_ -> fold () else fold () - | Rel _ | Var _ | Const _ | LetIn _ | Proj _ -> fold () + | Rel _ | Var _ | LetIn _ | Proj _ -> fold () | Sort _ | Ind _ | Prod _ -> fold () in fun xs -> diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index af0e28cdd..af4ea3ac5 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -34,6 +34,22 @@ end val get_refolding_in_reduction : unit -> bool val set_refolding_in_reduction : bool -> unit +(** {6 Support for reduction effects } *) + +type effect_name = string + +(* [declare_reduction_effect name f] declares [f] under key [name]; + [name] must be a unique in "world". *) +val declare_reduction_effect : effect_name -> + (Environ.env -> Evd.evar_map -> Constr.constr -> unit) -> unit + +(* [set_reduction_effect cst name] declares effect [name] to be called when [cst] is found *) +val set_reduction_effect : Globnames.global_reference -> effect_name -> unit + +(* [effect_hook env sigma key term] apply effect associated to [key] on [term] *) +val reduction_effect_hook : Environ.env -> Evd.evar_map -> Constr.constr -> + Constr.constr Lazy.t -> unit + (** {6 Machinery about a stack of unfolded constant } cst applied to params must convertible to term of the state applied to args diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index ec3669bfe..62737b65e 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -536,20 +536,27 @@ let reduce_mind_case_use_function func env sigma mia = | _ -> assert false -let match_eval_ref env sigma constr = +let match_eval_ref env sigma constr stack = match EConstr.kind sigma constr with - | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> - Some (EvalConst sp, u) + | Const (sp, u) -> + reduction_effect_hook env sigma (EConstr.to_constr sigma constr) + (lazy (EConstr.to_constr sigma (applist (constr,stack)))); + if is_evaluable env (EvalConstRef sp) then Some (EvalConst sp, u) else None | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, EInstance.empty) | Rel i -> Some (EvalRel i, EInstance.empty) | Evar ev -> Some (EvalEvar ev, EInstance.empty) | _ -> None -let match_eval_ref_value env sigma constr = +let match_eval_ref_value env sigma constr stack = match EConstr.kind sigma constr with - | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> - let u = EInstance.kind sigma u in - Some (EConstr.of_constr (constant_value_in env (sp, u))) + | Const (sp, u) -> + reduction_effect_hook env sigma (EConstr.to_constr sigma constr) + (lazy (EConstr.to_constr sigma (applist (constr,stack)))); + if is_evaluable env (EvalConstRef sp) then + let u = EInstance.kind sigma u in + Some (EConstr.of_constr (constant_value_in env (sp, u))) + else + None | Var id when is_evaluable env (EvalVarRef id) -> env |> lookup_named id |> NamedDecl.get_value | Rel n -> @@ -559,7 +566,7 @@ let match_eval_ref_value env sigma constr = let special_red_case env sigma whfun (ci, p, c, lf) = let rec redrec s = let (constr, cargs) = whfun s in - match match_eval_ref env sigma constr with + match match_eval_ref env sigma constr cargs with | Some (ref, u) -> (match reference_opt_value env sigma ref u with | None -> raise Redelimination @@ -774,7 +781,7 @@ and whd_simpl_stack env sigma = with Redelimination -> s') | _ -> - match match_eval_ref env sigma x with + match match_eval_ref env sigma x stack with | Some (ref, u) -> (try let sapp, nocase = red_elim_const env sigma ref u stack in @@ -796,7 +803,7 @@ and whd_simpl_stack env sigma = and whd_construct_stack env sigma s = let (constr, cargs as s') = whd_simpl_stack env sigma s in if reducible_mind_case sigma constr then s' - else match match_eval_ref env sigma constr with + else match match_eval_ref env sigma constr cargs with | Some (ref, u) -> (match reference_opt_value env sigma ref u with | None -> raise Redelimination @@ -844,7 +851,7 @@ let try_red_product env sigma c = | Reduced s -> simpfun (applist s) | NotReducible -> raise Redelimination) | _ -> - (match match_eval_ref env sigma x with + (match match_eval_ref env sigma x [] with | Some (ref, u) -> (* TO DO: re-fold fixpoints after expansion *) (* to get true one-step reductions *) @@ -925,7 +932,7 @@ let whd_simpl_stack = let whd_simpl_orelse_delta_but_fix env sigma c = let rec redrec s = let (constr, stack as s') = whd_simpl_stack env sigma s in - match match_eval_ref_value env sigma constr with + match match_eval_ref_value env sigma constr stack with | Some c -> (match EConstr.kind sigma (snd (decompose_lam sigma c)) with | CoFix _ | Fix _ -> s' diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 0fb48ed8c..67c8b07e7 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -248,24 +248,13 @@ let sort_eqns = unify_r2l let global_pattern_unification_flag = ref true -(* Compatibility option introduced and activated in Coq 8.3 whose - syntax is now deprecated. *) - open Goptions -let _ = - declare_bool_option - { optdepr = true; - optname = "pattern-unification for existential variables in tactics"; - optkey = ["Tactic";"Evars";"Pattern";"Unification"]; - optread = (fun () -> !global_pattern_unification_flag); - optwrite = (:=) global_pattern_unification_flag } -(* Compatibility option superseding the previous one, introduced and - activated in Coq 8.4 *) +(* Compatibility option introduced and activated in Coq 8.4 *) let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "pattern-unification for existential variables in tactics"; optkey = ["Tactic";"Pattern";"Unification"]; optread = (fun () -> !global_pattern_unification_flag); @@ -481,12 +470,10 @@ let set_flags_for_type flags = { flags with let use_evars_pattern_unification flags = !global_pattern_unification_flag && flags.use_pattern_unification - && Flags.version_strictly_greater Flags.V8_2 let use_metas_pattern_unification sigma flags nb l = !global_pattern_unification_flag && flags.use_pattern_unification - || (Flags.version_less_or_equal Flags.V8_3 || - flags.use_meta_bound_pattern_unification) && + || flags.use_meta_bound_pattern_unification && Array.for_all (fun c -> isRel sigma c && destRel sigma c <= nb) l type key = @@ -609,9 +596,6 @@ let do_reduce ts (env, nb) sigma c = Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma Cst_stack.empty (c, Stack.empty))) -let use_full_betaiota flags = - flags.modulo_betaiota && Flags.version_strictly_greater Flags.V8_3 - let isAllowedEvar sigma flags c = match EConstr.kind sigma c with | Evar (evk,_) -> not (Evar.Set.mem evk flags.frozen_evars) | _ -> false @@ -949,7 +933,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e expand curenvnb pb opt substn cM f1 l1 cN f2 l2 and reduce curenvnb pb opt (sigma, metas, evars as substn) cM cN = - if use_full_betaiota flags && not (subterm_restriction opt flags) then + if flags.modulo_betaiota && not (subterm_restriction opt flags) then let cM' = do_reduce flags.modulo_delta curenvnb sigma cM in if not (EConstr.eq_constr sigma cM cM') then unirec_rec curenvnb pb opt substn cM' cN diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 626464b96..49eedb767 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -80,7 +80,7 @@ let tag_var = tag Tag.variable | Any -> true let prec_of_prim_token = function - | Numeral p -> if Bigint.is_pos_or_zero p then lposint else lnegint + | Numeral (_,b) -> if b then lposint else lnegint | String _ -> latom open Notation @@ -231,7 +231,7 @@ let tag_var = tag Tag.variable | ArgVar (loc,s) -> pr_lident (loc,s) let pr_prim_token = function - | Numeral n -> str (Bigint.to_string n) + | Numeral (n,s) -> str (if s then n else "-"^n) | String s -> qs s let pr_evar pr id l = diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index 781af4789..9d28bc4f8 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -561,17 +561,13 @@ open Decl_kinds | GoalUid n -> spc () ++ str n in let pr_showable = function | ShowGoal n -> keyword "Show" ++ pr_goal_reference n - | ShowGoalImplicitly n -> keyword "Show Implicit Arguments" ++ pr_opt int n | ShowProof -> keyword "Show Proof" - | ShowNode -> keyword "Show Node" | ShowScript -> keyword "Show Script" | ShowExistentials -> keyword "Show Existentials" | ShowUniverses -> keyword "Show Universes" - | ShowTree -> keyword "Show Tree" | ShowProofNames -> keyword "Show Conjectures" | ShowIntros b -> keyword "Show " ++ (if b then keyword "Intros" else keyword "Intro") | ShowMatch id -> keyword "Show Match " ++ pr_reference id - | ShowThesis -> keyword "Show Thesis" in return (pr_showable s) | VernacCheckGuard -> diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 2b21b3f9e..3ae7da8fc 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -587,8 +587,6 @@ let gallina_print_library_entry with_values ent = Some (str " >>>>>>> Module " ++ pr_name oname) | (oname,Lib.ClosedModule _) -> Some (str " >>>>>>> Closed Module " ++ pr_name oname) - | (_,Lib.FrozenState _) -> - None let gallina_print_context with_values = let rec prec n = function diff --git a/printing/printer.ml b/printing/printer.ml index 3c31dd96b..d6f0778f7 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -17,7 +17,6 @@ open Nametab open Evd open Proof_type open Refiner -open Pfedit open Constrextern open Ppconstr open Declarations @@ -812,7 +811,7 @@ let pr_open_subgoals ?(proof=Proof_global.give_me_the_proof ()) () = end let pr_nth_open_subgoal n = - let pf = get_pftreestate () in + let pf = Proof_global.give_me_the_proof () in let { it=gls ; sigma=sigma } = Proof.V82.subgoals pf in pr_subgoal n sigma gls diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 79d2e4694..34875cbcd 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -662,7 +662,8 @@ let evar_of_binder holes = function | NamedHyp s -> evar_with_name holes s | AnonHyp n -> try - let h = List.nth holes (pred n) in + let nondeps = List.filter (fun hole -> not hole.hole_deps) holes in + let h = List.nth nondeps (pred n) in h.hole_evar with e when CErrors.noncritical e -> user_err (str "No such binder.") diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index 3fb66d1b8..b28234a50 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -24,19 +24,6 @@ let _ = Goptions.declare_bool_option { let use_unification_heuristics () = !use_unification_heuristics_ref -let refining = Proof_global.there_are_pending_proofs -let check_no_pending_proofs = Proof_global.check_no_pending_proof - -let get_current_proof_name = Proof_global.get_current_proof_name -let get_all_proof_names = Proof_global.get_all_proof_names - -type lemma_possible_guards = Proof_global.lemma_possible_guards -type universe_binders = Proof_global.universe_binders - -let delete_proof = Proof_global.discard -let delete_current_proof = Proof_global.discard_current -let delete_all_proofs = Proof_global.discard_all - let start_proof (id : Id.t) ?pl str sigma hyps c ?init_tac terminator = let goals = [ (Global.env_of_context hyps , c) ] in Proof_global.start_proof sigma id ?pl str goals terminator; @@ -55,32 +42,20 @@ let cook_this_proof p = let cook_proof () = cook_this_proof (fst (Proof_global.close_proof ~keep_body_ucst_separate:false (fun x -> x))) -let get_pftreestate () = - Proof_global.give_me_the_proof () - -let set_end_tac tac = - Proof_global.set_endline_tactic tac - -let set_used_variables l = - Proof_global.set_used_variables l -let get_used_variables () = - Proof_global.get_used_variables () - -let get_universe_binders () = - Proof_global.get_universe_binders () exception NoSuchGoal let _ = CErrors.register_handler begin function | NoSuchGoal -> CErrors.user_err Pp.(str "No such goal.") | _ -> raise CErrors.Unhandled end + let get_nth_V82_goal i = let p = Proof_global.give_me_the_proof () in let { it=goals ; sigma = sigma; } = Proof.V82.subgoals p in try { it=(List.nth goals (i-1)) ; sigma=sigma; } with Failure _ -> raise NoSuchGoal - + let get_goal_context_gen i = let { it=goal ; sigma=sigma; } = get_nth_V82_goal i in (sigma, Refiner.pf_env { it=goal ; sigma=sigma; }) @@ -106,7 +81,7 @@ let get_current_context () = (Evd.from_env env, env) | NoSuchGoal -> (* No more focused goals ? *) - let p = get_pftreestate () in + let p = Proof_global.give_me_the_proof () in let evd = Proof.in_proof p (fun x -> x) in (evd, Global.env ()) @@ -165,11 +140,11 @@ let build_constant_by_tactic id ctx sign ?(goal_kind = Global, false, Proof Theo try let status = by tac in let _,(const,univs,_) = cook_proof () in - delete_current_proof (); + Proof_global.discard_current (); const, status, fst univs with reraise -> let reraise = CErrors.push reraise in - delete_current_proof (); + Proof_global.discard_current (); iraise reraise let build_by_tactic ?(side_eff=true) env sigma ?(poly=false) typ tac = @@ -257,4 +232,32 @@ let solve_by_implicit_tactic () = match !implicit_tactic with | None -> None | Some tac -> Some (apply_implicit_tactic tac) +(** Deprecated functions *) +let refining = Proof_global.there_are_pending_proofs +let check_no_pending_proofs = Proof_global.check_no_pending_proof + +let get_current_proof_name = Proof_global.get_current_proof_name +let get_all_proof_names = Proof_global.get_all_proof_names + +type lemma_possible_guards = Proof_global.lemma_possible_guards +type universe_binders = Proof_global.universe_binders + +let delete_proof = Proof_global.discard +let delete_current_proof = Proof_global.discard_current +let delete_all_proofs = Proof_global.discard_all + +let get_pftreestate () = + Proof_global.give_me_the_proof () + +let set_end_tac tac = + Proof_global.set_endline_tactic tac + +let set_used_variables l = + Proof_global.set_used_variables l + +let get_used_variables () = + Proof_global.get_used_variables () + +let get_universe_binders () = + Proof_global.get_universe_binders () diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 1bf65b8ae..f009593e9 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -14,39 +14,6 @@ open Term open Environ open Decl_kinds -(** Several proofs can be opened simultaneously but at most one is - focused at some time. The following functions work by side-effect - on current set of open proofs. In this module, ``proofs'' means an - open proof (something started by vernacular command [Goal], [Lemma] - or [Theorem]), and ``goal'' means a subgoal of the current focused - proof *) - -(** {6 ... } *) -(** [refining ()] tells if there is some proof in progress, even if a not - focused one *) - -val refining : unit -> bool - -(** [check_no_pending_proofs ()] fails if there is still some proof in - progress *) - -val check_no_pending_proofs : unit -> unit - -(** {6 ... } *) -(** [delete_proof name] deletes proof of name [name] or fails if no proof - has this name *) - -val delete_proof : Id.t located -> unit - -(** [delete_current_proof ()] deletes current focused proof or fails if - no proof is focused *) - -val delete_current_proof : unit -> unit - -(** [delete_all_proofs ()] deletes all open proofs if any *) - -val delete_all_proofs : unit -> unit - (** {6 ... } *) (** [start_proof s str env t hook tac] starts a proof of name [s] and conclusion [t]; [hook] is optionally a function to be applied at @@ -55,12 +22,8 @@ val delete_all_proofs : unit -> unit systematically apply at initialization time (e.g. to start the proof of mutually dependent theorems) *) -type lemma_possible_guards = Proof_global.lemma_possible_guards - -type universe_binders = Id.t Loc.located list - val start_proof : - Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> named_context_val -> EConstr.constr -> + Id.t -> ?pl:Proof_global.universe_binders -> goal_kind -> Evd.evar_map -> named_context_val -> EConstr.constr -> ?init_tac:unit Proofview.tactic -> Proof_global.proof_terminator -> unit @@ -80,11 +43,6 @@ val cook_proof : unit -> (Safe_typing.private_constants Entries.definition_entry * Proof_global.proof_universes * goal_kind)) (** {6 ... } *) -(** [get_pftreestate ()] returns the current focused pending proof. - @raise NoCurrentProof if there is no pending proof. *) - -val get_pftreestate : unit -> Proof.proof - (** [get_goal_context n] returns the context of the [n]th subgoal of the current focused proof or raises a [UserError] if there is no focused proof or if there is no more subgoals *) @@ -109,34 +67,6 @@ val current_proof_statement : unit -> Id.t * goal_kind * EConstr.types (** {6 ... } *) -(** [get_current_proof_name ()] return the name of the current focused - proof or failed if no proof is focused *) - -val get_current_proof_name : unit -> Id.t - -(** [get_all_proof_names ()] returns the list of all pending proof names. - The first name is the current proof, the other names may come in - any order. *) - -val get_all_proof_names : unit -> Id.t list - -(** {6 ... } *) -(** [set_end_tac tac] applies tactic [tac] to all subgoal generate - by [solve] *) - -val set_end_tac : Genarg.glob_generic_argument -> unit - -(** {6 ... } *) -(** [set_used_variables l] declares that section variables [l] will be - used in the proof *) -val set_used_variables : - Id.t list -> Context.Named.t * Names.Id.t Loc.located list -val get_used_variables : unit -> Context.Named.t option - -(** {6 Universe binders } *) -val get_universe_binders : unit -> universe_binders option - -(** {6 ... } *) (** [solve (SelectNth n) tac] applies tactic [tac] to the [n]th subgoal of the current focused proof or raises a [UserError] if no proof is focused or if there is no [n]th subgoal. [solve SelectAll @@ -191,3 +121,88 @@ val clear_implicit_tactic : unit -> unit (* Raise Exit if cannot solve *) val solve_by_implicit_tactic : unit -> Pretyping.inference_hook option + +(** {5 Deprecated functions in favor of [Proof_global]} *) + +(** {6 ... } *) +(** Several proofs can be opened simultaneously but at most one is + focused at some time. The following functions work by side-effect + on current set of open proofs. In this module, ``proofs'' means an + open proof (something started by vernacular command [Goal], [Lemma] + or [Theorem]), and ``goal'' means a subgoal of the current focused + proof *) + +(** [refining ()] tells if there is some proof in progress, even if a not + focused one *) + +val refining : unit -> bool +[@@ocaml.deprecated "use Proof_global.there_are_pending_proofs"] + +(** [check_no_pending_proofs ()] fails if there is still some proof in + progress *) + +val check_no_pending_proofs : unit -> unit +[@@ocaml.deprecated "use Proof_global.check_no_pending_proofs"] + +(** {6 ... } *) +(** [delete_proof name] deletes proof of name [name] or fails if no proof + has this name *) + +val delete_proof : Id.t located -> unit +[@@ocaml.deprecated "use Proof_global.discard"] + +(** [delete_current_proof ()] deletes current focused proof or fails if + no proof is focused *) + +val delete_current_proof : unit -> unit +[@@ocaml.deprecated "use Proof_global.discard_current"] + +(** [delete_all_proofs ()] deletes all open proofs if any *) +val delete_all_proofs : unit -> unit +[@@ocaml.deprecated "use Proof_global.discard_all"] + +(** [get_pftreestate ()] returns the current focused pending proof. + @raise NoCurrentProof if there is no pending proof. *) + +val get_pftreestate : unit -> Proof.proof +[@@ocaml.deprecated "use Proof_global.give_me_the_proof"] + +(** {6 ... } *) +(** [set_end_tac tac] applies tactic [tac] to all subgoal generate + by [solve] *) + +val set_end_tac : Genarg.glob_generic_argument -> unit +[@@ocaml.deprecated "use Proof_global.set_endline_tactic"] + +(** {6 ... } *) +(** [set_used_variables l] declares that section variables [l] will be + used in the proof *) +val set_used_variables : + Id.t list -> Context.Named.t * Names.Id.t Loc.located list +[@@ocaml.deprecated "use Proof_global.set_used_variables"] + +val get_used_variables : unit -> Context.Named.t option +[@@ocaml.deprecated "use Proof_global.get_used_variables"] + +(** {6 Universe binders } *) +val get_universe_binders : unit -> Proof_global.universe_binders option +[@@ocaml.deprecated "use Proof_global.get_universe_binders"] + +(** {6 ... } *) +(** [get_current_proof_name ()] return the name of the current focused + proof or failed if no proof is focused *) +val get_current_proof_name : unit -> Id.t +[@@ocaml.deprecated "use Proof_global.get_current_proof_name"] + +(** [get_all_proof_names ()] returns the list of all pending proof names. + The first name is the current proof, the other names may come in + any order. *) +val get_all_proof_names : unit -> Id.t list +[@@ocaml.deprecated "use Proof_global.get_all_proof_names"] + +type lemma_possible_guards = Proof_global.lemma_possible_guards +[@@ocaml.deprecated "use Proof_global.lemma_possible_guards"] + +type universe_binders = Proof_global.universe_binders +[@@ocaml.deprecated "use Proof_global.universe_binders"] + diff --git a/proofs/proof.ml b/proofs/proof.ml index 2aa620c1d..ef454299e 100644 --- a/proofs/proof.ml +++ b/proofs/proof.ml @@ -428,7 +428,7 @@ module V82 = struct in let env = Evd.evar_filtered_env evi in let rawc = Constrintern.intern_constr env com in - let ltac_vars = Pretyping.empty_lvar in + let ltac_vars = Glob_ops.empty_lvar in let sigma = Evar_refiner.w_refine (evk, evi) (ltac_vars, rawc) sigma in Proofview.Unsafe.tclEVARS sigma end in diff --git a/proofs/refine.ml b/proofs/refine.ml index caa6b9fb3..796b4b837 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -69,7 +69,7 @@ let add_side_effect env = function let add_side_effects env effects = List.fold_left (fun env eff -> add_side_effect env eff) env effects -let generic_refine ?(unsafe = true) f gl = +let generic_refine ~typecheck f gl = let gl = Proofview.Goal.assume gl in let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in @@ -91,9 +91,9 @@ let generic_refine ?(unsafe = true) f gl = let env = add_side_effects env sideff in (** Check that the introduced evars are well-typed *) let fold accu ev = typecheck_evar ev env accu in - let sigma = if unsafe then sigma else CList.fold_left fold sigma evs in + let sigma = if typecheck then CList.fold_left fold sigma evs else sigma in (** Check that the refined term is typesafe *) - let sigma = if unsafe then sigma else typecheck_proof c concl env sigma in + let sigma = if typecheck then typecheck_proof c concl env sigma else sigma in (** Check that the goal itself does not appear in the refined term *) let self = Proofview.Goal.goal gl in let _ = @@ -132,16 +132,16 @@ let lift c = Proofview.tclUNIT c end -let make_refine_enter ?unsafe f gl = generic_refine ?unsafe (lift f) gl +let make_refine_enter ~typecheck f gl = generic_refine ~typecheck (lift f) gl -let refine_one ?(unsafe = true) f = - Proofview.Goal.enter_one (make_refine_enter ~unsafe f) +let refine_one ~typecheck f = + Proofview.Goal.enter_one (make_refine_enter ~typecheck f) -let refine ?(unsafe = true) f = +let refine ~typecheck f = let f evd = let (evd,c) = f evd in (evd,((), c)) in - Proofview.Goal.enter (make_refine_enter ~unsafe f) + Proofview.Goal.enter (make_refine_enter ~typecheck f) (** Useful definitions *) @@ -153,7 +153,7 @@ let with_type env evd c t = in evd , j'.Environ.uj_val -let refine_casted ?unsafe f = Proofview.Goal.enter begin fun gl -> +let refine_casted ~typecheck f = Proofview.Goal.enter begin fun gl -> let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in @@ -161,7 +161,7 @@ let refine_casted ?unsafe f = Proofview.Goal.enter begin fun gl -> let (h, c) = f h in with_type env h c concl in - refine ?unsafe f + refine ~typecheck f end (** {7 solve_constraints} diff --git a/proofs/refine.mli b/proofs/refine.mli index f1439f9a1..c1c57ecb8 100644 --- a/proofs/refine.mli +++ b/proofs/refine.mli @@ -21,19 +21,18 @@ val pr_constr : (** {7 Refinement primitives} *) -val refine : ?unsafe:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit tactic -(** In [refine ?unsafe t], [t] is a term with holes under some +val refine : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit tactic +(** In [refine ~typecheck t], [t] is a term with holes under some [evar_map] context. The term [t] is used as a partial solution for the current goal (refine is a goal-dependent tactic), the new holes created by [t] become the new subgoals. Exceptions raised during the interpretation of [t] are caught and result in - tactic failures. If [unsafe] is [false] (default is [true]) [t] is - type-checked beforehand. *) + tactic failures. If [typecheck] is [true] [t] is type-checked beforehand. *) -val refine_one : ?unsafe:bool -> (Evd.evar_map -> Evd.evar_map * ('a * EConstr.t)) -> 'a tactic +val refine_one : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * ('a * EConstr.t)) -> 'a tactic (** A variant of [refine] which assumes exactly one goal under focus *) -val generic_refine : ?unsafe:bool -> ('a * EConstr.t) tactic -> +val generic_refine : typecheck:bool -> ('a * EConstr.t) tactic -> [ `NF ] Proofview.Goal.t -> 'a tactic (** The general version of refine. *) @@ -44,7 +43,7 @@ val with_type : Environ.env -> Evd.evar_map -> (** [with_type env sigma c t] ensures that [c] is of type [t] inserting a coercion if needed. *) -val refine_casted : ?unsafe:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit tactic +val refine_casted : typecheck:bool -> (Evd.evar_map -> Evd.evar_map * EConstr.t) -> unit tactic (** Like {!refine} except the refined term is coerced to the conclusion of the current goal. *) diff --git a/stm/stm.ml b/stm/stm.ml index 11aca1627..8ca50e2d5 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -931,7 +931,7 @@ let show_script ?proof () = try let prf = try match proof with - | None -> Some (Pfedit.get_current_proof_name ()) + | None -> Some (Proof_global.get_current_proof_name ()) | Some (p,_) -> Some (p.Proof_global.id) with Proof_global.NoCurrentProof -> None in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index de49a521f..2faf1e0ec 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -250,7 +250,7 @@ let unify_resolve_refine poly flags gls ((c, t, ctx),n,clenv) = let open Clenv in let env = Proofview.Goal.env gls in let concl = Proofview.Goal.concl gls in - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let sigma, term, ty = if poly then let (subst, ctx) = Universes.fresh_universe_context_set_instance ctx in @@ -603,6 +603,7 @@ let make_hints g st only_classes sign = List.fold_left (fun hints hyp -> let consider = + not only_classes || try let t = hyp |> NamedDecl.get_id |> Global.lookup_named |> NamedDecl.get_type in (* Section variable, reindex only if the type changed *) not (EConstr.eq_constr (project g) (EConstr.of_constr t) (NamedDecl.get_type hyp)) diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index 0cee4b6ed..10bc6e3e2 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -72,7 +72,7 @@ let generalize_right mk typ c1 c2 = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let na = Name (next_name_away_with_default "x" Anonymous (Termops.ids_of_context env)) in let newconcl = mkProd (na, typ, mk typ c1 (mkRel 1)) in let (sigma, x) = Evarutil.new_evar env sigma ~principal:true ~store newconcl in diff --git a/tactics/equality.ml b/tactics/equality.ml index 05c5cd5ec..6e56dc48e 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -50,13 +50,12 @@ module NamedDecl = Context.Named.Declaration let discriminate_introduction = ref true -let discr_do_intro () = - !discriminate_introduction && Flags.version_strictly_greater Flags.V8_2 +let discr_do_intro () = !discriminate_introduction open Goptions let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "automatic introduction of hypotheses by discriminate"; optkey = ["Discriminate";"Introduction"]; optread = (fun () -> !discriminate_introduction); @@ -64,13 +63,11 @@ let _ = let injection_pattern_l2r_order = ref true -let use_injection_pattern_l2r_order () = - !injection_pattern_l2r_order - && Flags.version_strictly_greater Flags.V8_4 +let use_injection_pattern_l2r_order () = !injection_pattern_l2r_order let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "injection left-to-right pattern order and clear by default when with introduction pattern"; optkey = ["Injection";"L2R";"Pattern";"Order"]; optread = (fun () -> !injection_pattern_l2r_order) ; @@ -356,7 +353,6 @@ let find_elim hdcncl lft2rgt dep cls ot gl = if (is_global Coqlib.glob_eq hdcncl || (is_global Coqlib.glob_jmeq hdcncl && jmeq_same_dom gl ot)) && not dep - || Flags.version_less_or_equal Flags.V8_2 then let c = match EConstr.kind sigma hdcncl with @@ -1418,7 +1414,7 @@ let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause = "" else " You can try to use option Set Keep Proof Equalities." in tclZEROMSG (strbrk("No information can be deduced from this equality and the injectivity of constructors. This may be because the terms are convertible, or due to pattern matching restrictions in the sort Prop." ^ suggestion)) - | Inr [([],_,_)] when Flags.version_strictly_greater Flags.V8_3 -> + | Inr [([],_,_)] -> tclZEROMSG (str"Nothing to inject.") | Inr posns -> inject_at_positions env sigma l2r u eq_clause posns @@ -1769,13 +1765,10 @@ type subst_tactic_flags = { rewrite_dependent_proof : bool } -let default_subst_tactic_flags () = - if Flags.version_strictly_greater Flags.V8_2 then - { only_leibniz = false; rewrite_dependent_proof = true } - else - { only_leibniz = true; rewrite_dependent_proof = false } +let default_subst_tactic_flags = + { only_leibniz = false; rewrite_dependent_proof = true } -let subst_all ?(flags=default_subst_tactic_flags ()) () = +let subst_all ?(flags=default_subst_tactic_flags) () = if !regular_subst_tactic then diff --git a/tactics/hints.ml b/tactics/hints.ml index 773abb9f0..681db5d08 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -29,7 +29,6 @@ open Decl_kinds open Pattern open Patternops open Clenv -open Pfedit open Tacred open Printer open Vernacexpr @@ -1462,7 +1461,7 @@ let pr_hint_term sigma cl = (* print all hints that apply to the concl of the current goal *) let pr_applicable_hint () = - let pts = get_pftreestate () in + let pts = Proof_global.give_me_the_proof () in let glss = Proof.V82.subgoals pts in match glss.Evd.it with | [] -> CErrors.user_err Pp.(str "No focused goal.") diff --git a/tactics/inv.ml b/tactics/inv.ml index ec038f638..2bc9d9f78 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -460,7 +460,7 @@ let raw_inversion inv_kind id status names = in let refined id = let prf = mkApp (mkVar id, args) in - Refine.refine (fun h -> (h, prf)) + Refine.refine ~typecheck:false (fun h -> (h, prf)) in let neqns = List.length realargs in let as_mode = names != None in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index aa574e41c..4101dc23e 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -468,6 +468,7 @@ module New = struct let check_evars env sigma extsigma origsigma = let rec is_undefined_up_to_restriction sigma evk = + if Evd.mem origsigma evk then None else let evi = Evd.find sigma evk in match Evd.evar_body evi with | Evd.Evar_empty -> Some (evk,evi) @@ -481,7 +482,7 @@ module New = struct let rest = Evd.fold_undefined (fun evk evi acc -> match is_undefined_up_to_restriction sigma evk with - | Some (evk',evi) when not (Evd.mem origsigma evk) -> (evk',evi)::acc + | Some (evk',evi) -> (evk',evi)::acc | _ -> acc) extsigma [] in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index b553f316c..689cc48aa 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -25,7 +25,6 @@ open Inductiveops open Reductionops open Globnames open Evd -open Pfedit open Tacred open Genredexpr open Tacmach.New @@ -64,11 +63,10 @@ let dependent_propositions_elimination = ref true let use_dependent_propositions_elimination () = !dependent_propositions_elimination - && Flags.version_strictly_greater Flags.V8_2 let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "dependent-propositions-elimination tactic"; optkey = ["Dependent";"Propositions";"Elimination"]; optread = (fun () -> !dependent_propositions_elimination) ; @@ -142,11 +140,10 @@ let bracketing_last_or_and_intro_pattern = ref true let use_bracketing_last_or_and_intro_pattern () = !bracketing_last_or_and_intro_pattern - && Flags.version_strictly_greater Flags.V8_4 let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "bracketing last or-and introduction pattern"; optkey = ["Bracketing";"Last";"Introduction";"Pattern"]; optread = (fun () -> !bracketing_last_or_and_intro_pattern); @@ -163,7 +160,7 @@ let _ = (** This tactic creates a partial proof realizing the introduction rule, but does not check anything. *) let unsafe_intro env store decl b = - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let ctx = named_context_val env in let nctx = push_named_context_val decl ctx in let inst = List.map (NamedDecl.get_id %> mkVar) (named_context env) in @@ -200,7 +197,7 @@ let convert_concl ?(check=true) ty k = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in let conclty = Proofview.Goal.concl gl in - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let sigma = if check then begin ignore (Typing.unsafe_type_of env sigma ty); @@ -222,7 +219,7 @@ let convert_hyp ?(check=true) d = let store = Proofview.Goal.extra gl in let sign = convert_hyp check (named_context_val env) sigma d in let env = reset_with_named_context sign env in - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> Evarutil.new_evar env sigma ~principal:true ~store ty end end @@ -293,7 +290,7 @@ let clear_gen fail = function in let env = reset_with_named_context hyps env in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) - (Refine.refine ~unsafe:true begin fun sigma -> + (Refine.refine ~typecheck:false begin fun sigma -> Evarutil.new_evar env sigma ~principal:true concl end) end @@ -323,7 +320,7 @@ let move_hyp id dest = let sign = named_context_val env in let sign' = move_hyp_in_named_context sigma id dest sign in let env = reset_with_named_context sign' env in - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> Evarutil.new_evar env sigma ~principal:true ~store ty end end @@ -377,7 +374,7 @@ let rename_hyp repl = let nconcl = subst concl in let nctx = val_of_named_context nhyps in let instance = List.map (NamedDecl.get_id %> mkVar) hyps in - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> Evarutil.new_evar_instance nctx sigma nconcl ~principal:true ~store instance end end @@ -527,7 +524,7 @@ let mutual_fix f n rest j = Proofview.Goal.enter begin fun gl -> mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth in let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let (sigma, evs) = mk_holes nenv sigma (List.map pi3 all) in let ids = List.map pi1 all in let evs = List.map (Vars.subst_vars (List.rev ids)) evs in @@ -543,7 +540,7 @@ end let fix ido n = match ido with | None -> Proofview.Goal.enter begin fun gl -> - let name = Pfedit.get_current_proof_name () in + let name = Proof_global.get_current_proof_name () in let id = new_fresh_id [] name gl in mutual_fix id n [] 0 end @@ -579,7 +576,7 @@ let mutual_cofix f others j = Proofview.Goal.enter begin fun gl -> mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth in let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let (ids, types) = List.split all in let (sigma, evs) = mk_holes nenv sigma types in let evs = List.map (Vars.subst_vars (List.rev ids)) evs in @@ -594,7 +591,7 @@ end let cofix ido = match ido with | None -> Proofview.Goal.enter begin fun gl -> - let name = Pfedit.get_current_proof_name () in + let name = Proof_global.get_current_proof_name () in let id = new_fresh_id [] name gl in mutual_cofix id [] 0 end @@ -1140,7 +1137,7 @@ let rec intros_move = function let tactic_infer_flags with_evar = { Pretyping.use_typeclasses = true; Pretyping.solve_unification_constraints = true; - Pretyping.use_hook = solve_by_implicit_tactic (); + Pretyping.use_hook = Pfedit.solve_by_implicit_tactic (); Pretyping.fail_evar = not with_evar; Pretyping.expand_evars = true } @@ -1225,7 +1222,7 @@ let cut c = let id = next_name_away_with_default "H" Anonymous (Tacmach.New.pf_ids_of_hyps gl) in (** Backward compat: normalize [c]. *) let c = if normalize_cut then local_strong whd_betaiota sigma c else c in - Refine.refine ~unsafe:true begin fun h -> + Refine.refine ~typecheck:false begin fun h -> let (h, f) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in let (h, x) = Evarutil.new_evar env h c in let f = mkLetIn (Name id, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in @@ -1666,7 +1663,7 @@ let solve_remaining_apply_goals = if Typeclasses.is_class_type evd concl then let evd', c' = Typeclasses.resolve_one_typeclass env evd concl in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evd') - (Refine.refine ~unsafe:true (fun h -> (h,c'))) + (Refine.refine ~typecheck:false (fun h -> (h,c'))) else Proofview.tclUNIT () with Not_found -> Proofview.tclUNIT () else Proofview.tclUNIT () @@ -1914,7 +1911,7 @@ let cut_and_apply c = | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 -> let concl = Proofview.Goal.concl gl in let env = Tacmach.New.pf_env gl in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let typ = mkProd (Anonymous, c2, concl) in let (sigma, f) = Evarutil.new_evar env sigma typ in let (sigma, x) = Evarutil.new_evar env sigma c1 in @@ -1934,7 +1931,7 @@ let cut_and_apply c = (* let refine_no_check = Profile.profile2 refine_no_checkkey refine_no_check *) let exact_no_check c = - Refine.refine ~unsafe:true (fun h -> (h,c)) + Refine.refine ~typecheck:false (fun h -> (h,c)) let exact_check c = Proofview.Goal.enter begin fun gl -> @@ -1959,7 +1956,7 @@ let native_cast_no_check c = cast_no_check Term.NATIVEcast c let exact_proof c = let open Tacmach.New in Proofview.Goal.enter begin fun gl -> - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in let c = EConstr.of_constr c in let sigma = Evd.merge_universe_context sigma ctx in @@ -2076,7 +2073,7 @@ let clear_body ids = Tacticals.New.tclZEROMSG msg in check <*> - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> Evarutil.new_evar env sigma ~principal:true concl end end @@ -2128,7 +2125,7 @@ let apply_type newcl args = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let newcl = nf_betaiota sigma newcl (* As in former Logic.refine *) in let (sigma, ev) = Evarutil.new_evar env sigma ~principal:true ~store newcl in @@ -2149,7 +2146,7 @@ let bring_hyps hyps = let concl = Tacmach.New.pf_concl gl in let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in let args = Array.of_list (Context.Named.to_instance mkVar hyps) in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let (sigma, ev) = Evarutil.new_evar env sigma ~principal:true ~store newcl in (sigma, mkApp (ev, args)) @@ -2888,7 +2885,7 @@ let new_generalize_gen_let lconstr = 0 lconstr (concl, sigma, []) in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) - (Refine.refine begin fun sigma -> + (Refine.refine ~typecheck:false begin fun sigma -> let (sigma, ev) = Evarutil.new_evar env sigma ~principal:true newcl in (sigma, applist (ev, args)) end) @@ -3598,7 +3595,7 @@ let mk_term_eq homogeneous env sigma ty t ty' t' = let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = let open Context.Rel.Declaration in - Refine.refine begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let eqslen = List.length eqs in (* Abstract by the "generalized" hypothesis equality proof if necessary. *) let sigma, abshypeq, abshypt = @@ -4418,7 +4415,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim (* and destruct has side conditions first *) Tacticals.New.tclTHENLAST) (Tacticals.New.tclTHENLIST [ - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> let b = not with_evars && with_eq != None in let (sigma, c) = use_bindings env sigma elim b (c0,lbind) t0 in let t = Retyping.get_type_of env sigma c in @@ -4441,7 +4438,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim let env = reset_with_named_context sign env in let tac = Tacticals.New.tclTHENLIST [ - Refine.refine ~unsafe:true begin fun sigma -> + Refine.refine ~typecheck:false begin fun sigma -> mkletin_goal env sigma store with_eq true (id,lastlhyp,ccl,c) None end; tac @@ -5032,11 +5029,11 @@ let name_op_to_name name_op object_kind suffix = let default_gk = (Global, false, object_kind) in match name_op with | Some s -> - (try let _, gk, _ = current_proof_statement () in s, gk + (try let _, gk, _ = Pfedit.current_proof_statement () in s, gk with NoCurrentProof -> s, default_gk) | None -> let name, gk = - try let name, gk, _ = current_proof_statement () in name, gk + try let name, gk, _ = Pfedit.current_proof_statement () in name, gk with NoCurrentProof -> anon_id, default_gk in add_suffix name suffix, gk @@ -5101,7 +5098,7 @@ module New = struct rZeta=false;rDelta=false;rConst=[]}) {onhyps; concl_occs=AllOccurrences } - let refine ?unsafe c = - Refine.refine ?unsafe c <*> + let refine ~typecheck c = + Refine.refine ~typecheck c <*> reduce_after_refine end diff --git a/tactics/tactics.mli b/tactics/tactics.mli index ec8fe1145..2e17b8dd5 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -435,8 +435,8 @@ end module New : sig - val refine : ?unsafe:bool -> (evar_map -> evar_map * constr) -> unit Proofview.tactic - (** [refine ?unsafe c] is [Refine.refine ?unsafe c] + val refine : typecheck:bool -> (evar_map -> evar_map * constr) -> unit Proofview.tactic + (** [refine ~typecheck c] is [Refine.refine ~typecheck c] followed by beta-iota-reduction of the conclusion. *) val reduce_after_refine : unit Proofview.tactic diff --git a/test-suite/Makefile b/test-suite/Makefile index e15094ccf..5ab4cacda 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -169,9 +169,7 @@ summary.log: # local build, and downloadable on GitLab) report: summary.log $(HIDE)./save-logs.sh - $(HIDE)if [ -n "${TRAVIS}" ]; then echo 'travis_fold:start:coq.logs'; fi - $(HIDE)if [ -n "${TRAVIS}" ]; then find logs/ -name '*.log' -not -name 'summary.log' -exec cat '{}' ';'; fi - $(HIDE)if [ -n "${TRAVIS}" ]; then echo 'travis_fold:end:coq.logs'; fi + $(HIDE)if [ -n "${TRAVIS}" ]; then find logs/ -name '*.log' -not -name 'summary.log' -exec 'bash' '-c' 'echo "travis_fold:start:coq.logs.$$(echo '{}' | sed s,/,.,g)"' ';' -exec cat '{}' ';' -exec 'bash' '-c' 'echo "travis_fold:end:coq.logs.$$(echo '{}' | sed s,/,.,g)"' ';'; fi $(HIDE)if grep -q -F 'Error!' summary.log ; then echo FAILURES; grep -F 'Error!' summary.log; false; else echo NO FAILURES; fi ####################################################################### diff --git a/test-suite/bugs/closed/2141.v b/test-suite/bugs/closed/2141.v index 941ae530f..098a7e9e7 100644 --- a/test-suite/bugs/closed/2141.v +++ b/test-suite/bugs/closed/2141.v @@ -1,3 +1,4 @@ +Require Coq.extraction.Extraction. Require Import FSetList. Require Import OrderedTypeEx. diff --git a/test-suite/bugs/closed/3036.v b/test-suite/bugs/closed/3036.v index 451bec9b2..3b57310d6 100644 --- a/test-suite/bugs/closed/3036.v +++ b/test-suite/bugs/closed/3036.v @@ -15,11 +15,11 @@ Definition perm := Qc. Locate Qle_bool. Definition compatibleb (p1 p2 : perm) : bool := -let p1pos := Qle_bool 00 p1 in - let p2pos := Qle_bool 00 p2 in +let p1pos := Qle_bool 0 p1 in + let p2pos := Qle_bool 0 p2 in negb ( (p1pos && p2pos) - || ((p1pos || p2pos) && (negb (Qle_bool 00 ((p1 + p2)%Qc)))))%Qc. + || ((p1pos || p2pos) && (negb (Qle_bool 0 ((p1 + p2)%Qc)))))%Qc. Definition compatible (p1 p2 : perm) := compatibleb p1 p2 = true. diff --git a/test-suite/bugs/closed/3287.v b/test-suite/bugs/closed/3287.v index 7c7813125..1b758acd7 100644 --- a/test-suite/bugs/closed/3287.v +++ b/test-suite/bugs/closed/3287.v @@ -1,3 +1,5 @@ +Require Coq.extraction.Extraction. + Module Foo. (* Definition foo := (I,I). *) Definition bar := true. diff --git a/test-suite/bugs/closed/3923.v b/test-suite/bugs/closed/3923.v index 0aa029e73..2fb0a5439 100644 --- a/test-suite/bugs/closed/3923.v +++ b/test-suite/bugs/closed/3923.v @@ -1,3 +1,5 @@ +Require Coq.extraction.Extraction. + Module Type TRIVIAL. Parameter t:Type. End TRIVIAL. diff --git a/test-suite/bugs/closed/4132.v b/test-suite/bugs/closed/4132.v new file mode 100644 index 000000000..806ffb771 --- /dev/null +++ b/test-suite/bugs/closed/4132.v @@ -0,0 +1,31 @@ + +Require Import ZArith Omega. +Open Scope Z_scope. + +(** bug 4132: omega was using "simpl" either on whole equations, or on + delimited but wrong spots. This was leading to unexpected reductions + when one atom (here [b]) is an evaluable reference instead of a variable. *) + +Lemma foo + (x y x' zxy zxy' z : Z) + (b := 5) + (Ry : - b <= y < b) + (Bx : x' <= b) + (H : - zxy' <= zxy) + (H' : zxy' <= x') : - b <= zxy. +Proof. +omega. (* was: Uncaught exception Invalid_argument("index out of bounds"). *) +Qed. + +Lemma foo2 x y (b := 5) (H1 : x <= y) (H2 : y <= b) : x <= b. +omega. (* Pierre L: according to a comment of bug report #4132, + this might have triggered "index out of bounds" in the past, + but I never managed to reproduce that in any version, + even before my fix. *) +Qed. + +Lemma foo3 x y (b := 0) (H1 : x <= y) (H2 : y <= b) : x <= b. +omega. (* Pierre L: according to a comment of bug report #4132, + this might have triggered "Failure(occurence 2)" in the past, + but I never managed to reproduce that. *) +Qed. diff --git a/test-suite/bugs/closed/4394.v b/test-suite/bugs/closed/4394.v deleted file mode 100644 index 1ad81345d..000000000 --- a/test-suite/bugs/closed/4394.v +++ /dev/null @@ -1,19 +0,0 @@ -(* -*- coq-prog-args: ("-compat" "8.4") -*- *) - -Require Import Equality List. -Inductive Foo (I : Type -> Type) (A : Type) : Type := -| foo (B : Type) : A -> I B -> Foo I A. -Definition Family := Type -> Type. -Definition FooToo : Family -> Family := Foo. -Definition optionize (I : Type -> Type) (A : Type) := option (I A). -Definition bar (I : Type -> Type) (A : Type) : A -> option (I A) -> Foo (optionize I) A := foo (optionize I) A A. -Record Rec (I : Type -> Type) := { rec : forall A : Type, A -> I A -> Foo I A }. -Definition barRec : Rec (optionize id) := {| rec := bar id |}. -Inductive Empty {T} : T -> Prop := . -Theorem empty (family : Family) (a : fold_right prod unit (map (Foo family) nil)) (b : unit) : - Empty (a, b) -> False. -Proof. - intro e. - dependent induction e. -Qed. - diff --git a/test-suite/bugs/closed/4400.v b/test-suite/bugs/closed/4400.v deleted file mode 100644 index a89cf0cbc..000000000 --- a/test-suite/bugs/closed/4400.v +++ /dev/null @@ -1,19 +0,0 @@ -(* -*- coq-prog-args: ("-require" "Coq.Compat.Coq84" "-compat" "8.4") -*- *) -Require Import Coq.Lists.List Coq.Logic.JMeq Program.Equality. -Set Printing Universes. -Inductive Foo (I : Type -> Type) (A : Type) : Type := -| foo (B : Type) : A -> I B -> Foo I A. -Definition Family := Type -> Type. -Definition FooToo : Family -> Family := Foo. -Definition optionize (I : Type -> Type) (A : Type) := option (I A). -Definition bar (I : Type -> Type) (A : Type) : A -> option (I A) -> Foo(optionize I) A := foo (optionize I) A A. -Record Rec (I : Type -> Type) := { rec : forall A : Type, A -> I A -> Foo I A }. -Definition barRec : Rec (optionize id) := {| rec := bar id |}. -Inductive Empty {T} : T -> Prop := . -Theorem empty (family : Family) (a : fold_right prod unit (map (Foo family) -nil)) (b : unit) : - Empty (a, b) -> False. -Proof. - intro e. - dependent induction e. -Qed. diff --git a/test-suite/bugs/closed/4616.v b/test-suite/bugs/closed/4616.v index c862f8206..a59975dbc 100644 --- a/test-suite/bugs/closed/4616.v +++ b/test-suite/bugs/closed/4616.v @@ -1,3 +1,5 @@ +Require Coq.extraction.Extraction. + Set Primitive Projections. Record Foo' := Foo { foo : Type }. Axiom f : forall t : Foo', foo t. diff --git a/test-suite/bugs/closed/4656.v b/test-suite/bugs/closed/4656.v deleted file mode 100644 index a59eed2c8..000000000 --- a/test-suite/bugs/closed/4656.v +++ /dev/null @@ -1,4 +0,0 @@ -(* -*- coq-prog-args: ("-compat" "8.4") -*- *) -Goal True. - constructor 1. -Qed. diff --git a/test-suite/bugs/closed/4710.v b/test-suite/bugs/closed/4710.v index fdc850109..5d8ca330a 100644 --- a/test-suite/bugs/closed/4710.v +++ b/test-suite/bugs/closed/4710.v @@ -1,3 +1,5 @@ +Require Coq.extraction.Extraction. + Set Primitive Projections. Record Foo' := Foo { foo : nat }. Extraction foo. diff --git a/test-suite/bugs/closed/4727.v b/test-suite/bugs/closed/4727.v deleted file mode 100644 index cfb4548d2..000000000 --- a/test-suite/bugs/closed/4727.v +++ /dev/null @@ -1,10 +0,0 @@ -(* -*- coq-prog-args: ("-compat" "8.4") -*- *) -Goal forall (P : Set) (l : P) (P0 : Set) (w w0 : P0) (T : Type) (a : P * T) (o : P -> option P0), - (forall (l1 l2 : P) (w1 : P0), o l1 = Some w1 -> o l2 = Some w1 -> l1 = l2) -> - o l = Some w -> o (fst a) = Some w0 -> {w = w0} + {w <> w0} -> False. -Proof. - clear; intros ???????? inj H0 H1 H2. - destruct H2; intuition subst. - eapply inj in H1; [ | eauto ]. - progress subst. (* should succeed, used to not succeed *) -Abort. diff --git a/test-suite/bugs/closed/4733.v b/test-suite/bugs/closed/4733.v deleted file mode 100644 index a90abd71c..000000000 --- a/test-suite/bugs/closed/4733.v +++ /dev/null @@ -1,52 +0,0 @@ -(* -*- coq-prog-args: ("-compat" "8.4") -*- *) -(*Suppose a user wants to declare a new list-like notation with support for singletons in both 8.4 and 8.5. If they use*) -Require Import Coq.Lists.List. -Require Import Coq.Vectors.Vector. -Import ListNotations. -Import VectorNotations. -Set Implicit Arguments. -Inductive mylist T := mynil | mycons (_ : T) (_ : mylist T). -Arguments mynil {_}, _. - -Delimit Scope mylist_scope with mylist. -Bind Scope mylist_scope with mylist. -Delimit Scope vector_scope with vector. - -Notation " [ ] " := mynil (format "[ ]") : mylist_scope. -Notation " [ x ] " := (mycons x mynil) : mylist_scope. -Notation " [ x ; .. ; y ] " := (mycons x .. (mycons y mynil) ..) : mylist_scope. - -(** All of these should work fine in -compat 8.4 mode, just as they do in Coq 8.4. There needs to be a way to specify notations above so that all of these [Check]s go through in both 8.4 and 8.5 *) -Check [ ]%mylist : mylist _. -Check [ ]%list : list _. -Check []%vector : Vector.t _ _. -Check [ _ ]%mylist : mylist _. -Check [ _ ]%list : list _. -Check [ _ ]%vector : Vector.t _ _. -Check [ _ ; _ ]%list : list _. -Check [ _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ]%mylist : mylist _. -Check [ _ ; _ ; _ ]%list : list _. -Check [ _ ; _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ; _ ]%mylist : mylist _. -Check [ _ ; _ ; _ ; _ ]%list : list _. -Check [ _ ; _ ; _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ; _ ; _ ]%mylist : mylist _. - -Notation " [ x ; y ; .. ; z ] " := (mycons x (mycons y .. (mycons z mynil) ..)) : mylist_scope. -(* Now these all work, but not so in 8.4. If we get the ability to remove notations, this section can also just be removed. *) -Check [ ]%mylist : mylist _. -Check [ ]%list : list _. -Check []%vector : Vector.t _ _. -Check [ _ ]%mylist : mylist _. -Check [ _ ]%list : list _. -Check [ _ ]%vector : Vector.t _ _. -Check [ _ ; _ ]%list : list _. -Check [ _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ]%mylist : mylist _. -Check [ _ ; _ ; _ ]%list : list _. -Check [ _ ; _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ; _ ]%mylist : mylist _. -Check [ _ ; _ ; _ ; _ ]%list : list _. -Check [ _ ; _ ; _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ; _ ; _ ]%mylist : mylist _. diff --git a/test-suite/bugs/closed/5019.v b/test-suite/bugs/closed/5019.v new file mode 100644 index 000000000..7c973f88b --- /dev/null +++ b/test-suite/bugs/closed/5019.v @@ -0,0 +1,5 @@ +Require Import Coq.ZArith.ZArith. +Goal forall (T0 : Z -> Type) (k : nat) d (P : T0 (Z.of_nat (S k)) -> Prop), P d. + clear; intros. + Timeout 1 zify. (* used to loop forever; should take < 0.01 s *) +Admitted. diff --git a/test-suite/bugs/closed/5255.v b/test-suite/bugs/closed/5255.v new file mode 100644 index 000000000..5daaf9edb --- /dev/null +++ b/test-suite/bugs/closed/5255.v @@ -0,0 +1,24 @@ +Section foo. + Context (x := 1). + Definition foo : x = 1 := eq_refl. +End foo. + +Module Type Foo. + Context (x := 1). + Definition foo : x = 1 := eq_refl. +End Foo. + +Set Universe Polymorphism. + +Inductive unit := tt. +Inductive eq {A} (x y : A) : Type := eq_refl : eq x y. + +Section bar. + Context (x := tt). + Definition bar : eq x tt := eq_refl _ _. +End bar. + +Module Type Bar. + Context (x := tt). + Definition bar : eq x tt := eq_refl _ _. +End Bar. diff --git a/test-suite/bugs/closed/5372.v b/test-suite/bugs/closed/5372.v index 2dc78d4c7..e60244cd1 100644 --- a/test-suite/bugs/closed/5372.v +++ b/test-suite/bugs/closed/5372.v @@ -1,4 +1,5 @@ (* coq bug 5372: https://coq.inria.fr/bugs/show_bug.cgi?id=5372 *) +Require Import FunInd. Function odd (n:nat) := match n with | 0 => false diff --git a/test-suite/bugs/closed/5414.v b/test-suite/bugs/closed/5414.v new file mode 100644 index 000000000..2522a274f --- /dev/null +++ b/test-suite/bugs/closed/5414.v @@ -0,0 +1,12 @@ +(* Use of idents bound to ltac names in a "match" *) + +Definition foo : Type. +Proof. + let x := fresh "a" in + refine (forall k : nat * nat, let '(x, _) := k in (_ : Type)). + exact (a = a). +Defined. +Goal foo. +intros k. elim k. (* elim because elim keeps names *) +intros. +Check a. (* We check that the name is "a" *) diff --git a/test-suite/bugs/closed/5486.v b/test-suite/bugs/closed/5486.v new file mode 100644 index 000000000..390133162 --- /dev/null +++ b/test-suite/bugs/closed/5486.v @@ -0,0 +1,15 @@ +Axiom proof_admitted : False. +Tactic Notation "admit" := abstract case proof_admitted. +Goal forall (T : Type) (p : prod (prod T T) bool) (Fm : Set) (f : Fm) (k : + forall _ : T, Fm), + @eq Fm + (k + match p return T with + | pair p0 swap => fst p0 + end) f. + intros. + (* next statement failed in Bug 5486 *) + match goal with + | [ |- ?f (let (a, b) := ?d in @?e a b) = ?rhs ] + => pose (let (a, b) := d in e a b) as t0 + end. diff --git a/test-suite/bugs/closed/5526.v b/test-suite/bugs/closed/5526.v new file mode 100644 index 000000000..88f219be3 --- /dev/null +++ b/test-suite/bugs/closed/5526.v @@ -0,0 +1,3 @@ +Fail Notation "x === x" := (eq_refl x) (at level 10). +Reserved Notation "x === x" (only printing, at level 10). +Notation "x === x" := (eq_refl x) (only printing). diff --git a/test-suite/bugs/closed/5550.v b/test-suite/bugs/closed/5550.v new file mode 100644 index 000000000..bb1222489 --- /dev/null +++ b/test-suite/bugs/closed/5550.v @@ -0,0 +1,10 @@ +Section foo. + + Variable bar : Prop. + Variable H : bar. + + Goal bar. + typeclasses eauto with foobar. + Qed. + +End foo. diff --git a/test-suite/bugs/opened/4803.v b/test-suite/bugs/opened/4803.v deleted file mode 100644 index 4541f13d0..000000000 --- a/test-suite/bugs/opened/4803.v +++ /dev/null @@ -1,48 +0,0 @@ -(* -*- coq-prog-args: ("-compat" "8.4") -*- *) -(*Suppose a user wants to declare a new list-like notation with support for singletons in both 8.4 and 8.5. If they use*) -Require Import Coq.Lists.List. -Require Import Coq.Vectors.Vector. -Import ListNotations. -Import VectorNotations. -Set Implicit Arguments. -Inductive mylist T := mynil | mycons (_ : T) (_ : mylist T). -Arguments mynil {_}, _. - -Delimit Scope mylist_scope with mylist. -Bind Scope mylist_scope with mylist. -Delimit Scope vector_scope with vector. - -Notation " [ ] " := mynil (format "[ ]") : mylist_scope. -Notation " [ x ] " := (mycons x mynil) : mylist_scope. -Notation " [ x ; .. ; y ] " := (mycons x .. (mycons y mynil) ..) : mylist_scope. - -(** All of these should work fine in -compat 8.4 mode, just as they do in Coq 8.4. There needs to be a way to specify notations above so that all of these [Check]s go through in both 8.4 and 8.5 *) -Check [ ]%mylist : mylist _. -Check [ ]%list : list _. -Check []%vector : Vector.t _ _. -Check [ _ ]%mylist : mylist _. -Check [ _ ]%list : list _. -Check [ _ ]%vector : Vector.t _ _. -Check [ _ ; _ ]%list : list _. -Check [ _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ]%mylist : mylist _. -Check [ _ ; _ ; _ ]%list : list _. -Check [ _ ; _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ; _ ]%mylist : mylist _. -Check [ _ ; _ ; _ ; _ ]%list : list _. -Check [ _ ; _ ; _ ; _ ]%vector : Vector.t _ _. -Check [ _ ; _ ; _ ; _ ]%mylist : mylist _. - -(** Now check that we can add and then remove notations from the parser *) -(* We should be able to stick some vernacular here to remove [] from the parser *) -Fail Remove Notation "[]". -Goal True. - (* This should not be a syntax error; before moving this file to closed, uncomment this line. *) - (* idtac; []. *) - constructor. -Qed. - -Check { _ : _ & _ }. -Reserved Infix "&" (at level 0). -Fail Remove Infix "&". -(* Check { _ : _ & _ }. *) diff --git a/test-suite/coq-makefile/arg/_CoqProject b/test-suite/coq-makefile/arg/_CoqProject index afdb32e7c..53dc96399 100644 --- a/test-suite/coq-makefile/arg/_CoqProject +++ b/test-suite/coq-makefile/arg/_CoqProject @@ -1,7 +1,7 @@ -R theories test -R src test -I src --arg "-compat 8.4" +-arg "-w default" src/test_plugin.mlpack src/test.ml4 diff --git a/test-suite/coq-makefile/coqdoc1/run.sh b/test-suite/coq-makefile/coqdoc1/run.sh index d6bb52bb4..e8291c89d 100755 --- a/test-suite/coq-makefile/coqdoc1/run.sh +++ b/test-suite/coq-makefile/coqdoc1/run.sh @@ -15,9 +15,7 @@ make install-doc DSTROOT="$PWD/tmp" sort -u > desired <<EOT . ./test -./test/test_plugin.cma ./test/test_plugin.cmi -./test/test_plugin.cmo ./test/test_plugin.cmx ./test/test_plugin.cmxs ./test/test.glob diff --git a/test-suite/coq-makefile/coqdoc2/run.sh b/test-suite/coq-makefile/coqdoc2/run.sh index d6bb52bb4..e8291c89d 100755 --- a/test-suite/coq-makefile/coqdoc2/run.sh +++ b/test-suite/coq-makefile/coqdoc2/run.sh @@ -15,9 +15,7 @@ make install-doc DSTROOT="$PWD/tmp" sort -u > desired <<EOT . ./test -./test/test_plugin.cma ./test/test_plugin.cmi -./test/test_plugin.cmo ./test/test_plugin.cmx ./test/test_plugin.cmxs ./test/test.glob diff --git a/test-suite/coq-makefile/mlpack1/run.sh b/test-suite/coq-makefile/mlpack1/run.sh index f6fb3bcb4..10a200dde 100755 --- a/test-suite/coq-makefile/mlpack1/run.sh +++ b/test-suite/coq-makefile/mlpack1/run.sh @@ -15,9 +15,7 @@ sort > desired <<EOT . ./test ./test/test.glob -./test/test_plugin.cma ./test/test_plugin.cmi -./test/test_plugin.cmo ./test/test_plugin.cmx ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/mlpack2/run.sh b/test-suite/coq-makefile/mlpack2/run.sh index f6fb3bcb4..10a200dde 100755 --- a/test-suite/coq-makefile/mlpack2/run.sh +++ b/test-suite/coq-makefile/mlpack2/run.sh @@ -15,9 +15,7 @@ sort > desired <<EOT . ./test ./test/test.glob -./test/test_plugin.cma ./test/test_plugin.cmi -./test/test_plugin.cmo ./test/test_plugin.cmx ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/multiroot/run.sh b/test-suite/coq-makefile/multiroot/run.sh index 863c39f50..3cd1ac305 100755 --- a/test-suite/coq-makefile/multiroot/run.sh +++ b/test-suite/coq-makefile/multiroot/run.sh @@ -19,12 +19,9 @@ sort > desired <<EOT ./test ./test/test.glob ./test/test.cmi -./test/test.cmo ./test/test.cmx ./test/test_aux.cmi -./test/test_aux.cmo ./test/test_aux.cmx -./test/test_plugin.cma ./test/test_plugin.cmxa ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/native1/run.sh b/test-suite/coq-makefile/native1/run.sh index f07966263..9f6295d64 100755 --- a/test-suite/coq-makefile/native1/run.sh +++ b/test-suite/coq-makefile/native1/run.sh @@ -18,9 +18,7 @@ sort > desired <<EOT . ./test ./test/test.glob -./test/test_plugin.cma ./test/test_plugin.cmi -./test/test_plugin.cmo ./test/test_plugin.cmx ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/plugin-reach-outside-API-and-fail/run.sh b/test-suite/coq-makefile/plugin-reach-outside-API-and-fail/run.sh new file mode 100755 index 000000000..6301aa03c --- /dev/null +++ b/test-suite/coq-makefile/plugin-reach-outside-API-and-fail/run.sh @@ -0,0 +1,37 @@ +#!/usr/bin/env bash + +set -e + +git clean -dfx + +cat > _CoqProject <<EOT +-I src/ + +./src/test_plugin.mllib +./src/test.ml4 +./src/test.mli +EOT + +mkdir src + +cat > src/test_plugin.mllib <<EOT +Test +EOT + +touch src/test.mli + +cat > src/test.ml4 <<EOT +DECLARE PLUGIN "test" + +let _ = Pre_env.empty_env +EOT + +${COQBIN}coq_makefile -f _CoqProject -o Makefile + +if make VERBOSE=1; then + # make command should have failed (but didn't) + exit 1 +else + # make command should have failed (and it indeed did) + exit 0 +fi diff --git a/test-suite/coq-makefile/plugin-reach-outside-API-and-succeed-by-bypassing-the-API/run.sh b/test-suite/coq-makefile/plugin-reach-outside-API-and-succeed-by-bypassing-the-API/run.sh new file mode 100755 index 000000000..991fb4a61 --- /dev/null +++ b/test-suite/coq-makefile/plugin-reach-outside-API-and-succeed-by-bypassing-the-API/run.sh @@ -0,0 +1,32 @@ +#!/usr/bin/env bash + +set -e + +git clean -dfx + +cat > _CoqProject <<EOT +-bypass-API +-I src/ + +./src/test_plugin.mllib +./src/test.ml4 +./src/test.mli +EOT + +mkdir src + +cat > src/test_plugin.mllib <<EOT +Test +EOT + +touch src/test.mli + +cat > src/test.ml4 <<EOT +DECLARE PLUGIN "test" + +let _ = Pre_env.empty_env +EOT + +${COQBIN}coq_makefile -f _CoqProject -o Makefile + +make VERBOSE=1 diff --git a/test-suite/coq-makefile/plugin1/run.sh b/test-suite/coq-makefile/plugin1/run.sh index 24ef8c891..c2d47166f 100755 --- a/test-suite/coq-makefile/plugin1/run.sh +++ b/test-suite/coq-makefile/plugin1/run.sh @@ -17,12 +17,9 @@ sort > desired <<EOT ./test ./test/test.glob ./test/test.cmi -./test/test.cmo ./test/test.cmx ./test/test_aux.cmi -./test/test_aux.cmo ./test/test_aux.cmx -./test/test_plugin.cma ./test/test_plugin.cmxa ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/plugin2/run.sh b/test-suite/coq-makefile/plugin2/run.sh index 24ef8c891..c2d47166f 100755 --- a/test-suite/coq-makefile/plugin2/run.sh +++ b/test-suite/coq-makefile/plugin2/run.sh @@ -17,12 +17,9 @@ sort > desired <<EOT ./test ./test/test.glob ./test/test.cmi -./test/test.cmo ./test/test.cmx ./test/test_aux.cmi -./test/test_aux.cmo ./test/test_aux.cmx -./test/test_plugin.cma ./test/test_plugin.cmxa ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/plugin3/run.sh b/test-suite/coq-makefile/plugin3/run.sh index 24ef8c891..c2d47166f 100755 --- a/test-suite/coq-makefile/plugin3/run.sh +++ b/test-suite/coq-makefile/plugin3/run.sh @@ -17,12 +17,9 @@ sort > desired <<EOT ./test ./test/test.glob ./test/test.cmi -./test/test.cmo ./test/test.cmx ./test/test_aux.cmi -./test/test_aux.cmo ./test/test_aux.cmx -./test/test_plugin.cma ./test/test_plugin.cmxa ./test/test_plugin.cmxs ./test/test.v diff --git a/test-suite/coq-makefile/template/src/test.ml4 b/test-suite/coq-makefile/template/src/test.ml4 index 72765abe0..e7d0bfe1f 100644 --- a/test-suite/coq-makefile/template/src/test.ml4 +++ b/test-suite/coq-makefile/template/src/test.ml4 @@ -1,3 +1,4 @@ +open API open Ltac_plugin DECLARE PLUGIN "test_plugin" let () = Mltop.add_known_plugin (fun () -> ()) "test_plugin";; diff --git a/test-suite/coq-makefile/template/src/test_aux.ml b/test-suite/coq-makefile/template/src/test_aux.ml index a01d0865a..e134abd84 100644 --- a/test-suite/coq-makefile/template/src/test_aux.ml +++ b/test-suite/coq-makefile/template/src/test_aux.ml @@ -1 +1 @@ -let tac = Proofview.tclUNIT () +let tac = API.Proofview.tclUNIT () diff --git a/test-suite/coq-makefile/template/src/test_aux.mli b/test-suite/coq-makefile/template/src/test_aux.mli index 10020f27d..2e7ad1529 100644 --- a/test-suite/coq-makefile/template/src/test_aux.mli +++ b/test-suite/coq-makefile/template/src/test_aux.mli @@ -1 +1 @@ -val tac : unit Proofview.tactic +val tac : unit API.Proofview.tactic diff --git a/test-suite/coqchk/univ.v b/test-suite/coqchk/univ.v index 84a4009d7..19eea94b1 100644 --- a/test-suite/coqchk/univ.v +++ b/test-suite/coqchk/univ.v @@ -33,3 +33,16 @@ Inductive finite_of_order T (D : T -> Type) (n : natural) := (rank_injective : injective_in T natural D rank) (rank_onto : forall i, equivalent (less_than i n) (in_image T natural D rank i)). + +(* Constraints *) +Universes i j. +Inductive constraint1 : (Type -> Type) -> Type := mk_constraint1 : constraint1 (fun x : Type@{i} => (x : Type@{j})). +Constraint i < j. +Inductive constraint2 : Type@{j} := mkc2 (_ : Type@{i}). +Universes i' j'. +Constraint i' = j'. +Inductive constraint3 : (Type -> Type) -> Type := mk_constraint3 : constraint3 (fun x : Type@{i'} => (x : Type@{j'})). +Inductive constraint4 : (Type -> Type) -> Type + := mk_constraint4 : let U1 := Type in + let U2 := Type in + constraint4 (fun x : U1 => (x : U2)). 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/ide/blocking-futures.fake b/test-suite/ide/blocking-futures.fake index b63f09bcf..541fb798c 100644 --- a/test-suite/ide/blocking-futures.fake +++ b/test-suite/ide/blocking-futures.fake @@ -4,6 +4,7 @@ # Extraction will force the future computation, assert it is blocking # Example courtesy of Jonathan (jonikelee) # +ADD { Require Coq.extraction.Extraction. } ADD { Require Import List. } ADD { Import ListNotations. } ADD { Definition myrev{A}(l : list A) : {rl : list A | rl = rev l}. } diff --git a/test-suite/output/Cases.out b/test-suite/output/Cases.out index 8ce6f9795..97fa8e254 100644 --- a/test-suite/output/Cases.out +++ b/test-suite/output/Cases.out @@ -2,18 +2,18 @@ t_rect = fun (P : t -> Type) (f : let x := t in forall x0 : x, P x0 -> P (k x0)) => fix F (t : t) : P t := match t as t0 return (P t0) with - | @k _ x0 => f x0 (F x0) + | k _ x0 => f x0 (F x0) end : forall P : t -> Type, (let x := t in forall x0 : x, P x0 -> P (k x0)) -> forall t : t, P t Argument scopes are [function_scope function_scope _] = fun d : TT => match d with - | @CTT _ _ b => b + | {| f3 := b |} => b end : TT -> 0 = 0 = fun d : TT => match d with - | @CTT _ _ b => b + | {| f3 := b |} => b end : TT -> 0 = 0 proj = @@ -72,3 +72,57 @@ e1 : texp t1 e2 : texp t2 The term "0" has type "nat" while it is expected to have type "typeDenote t0". +fun '{{n, m, _}} => n + m + : J -> nat +fun '{{n, m, p}} => n + m + p + : J -> nat +fun '(D n m p q) => n + m + p + q + : J -> nat +The command has indeed failed with message: +The constructor D (in type J) expects 3 arguments. +lem1 = +fun dd : nat * nat => let (bb, cc) as aa return (aa = aa) := dd in eq_refl + : forall k : nat * nat, k = k +lem2 = +fun dd : bool => if dd as aa return (aa = aa) then eq_refl else eq_refl + : forall k : bool, k = k + +Argument scope is [bool_scope] +lem3 = +fun dd : nat * nat => let (bb, cc) as aa return (aa = aa) := dd in eq_refl + : forall k : nat * nat, k = k +1 subgoal + + x : nat + n, n0 := match x + 0 with + | 0 => 0 + | S _ => 0 + end : nat + e, + e0 := match x + 0 as y return (y = y) with + | 0 => eq_refl + | S n => eq_refl + end : x + 0 = x + 0 + n1, n2 := match x with + | 0 => 0 + | S _ => 0 + end : nat + e1, e2 := match x return (x = x) with + | 0 => eq_refl + | S n => eq_refl + end : x = x + ============================ + x + 0 = 0 +1 subgoal + + p : nat + a, + a0 := match eq_refl as y in (_ = e) return (y = y /\ e = e) with + | eq_refl => conj eq_refl eq_refl + end : eq_refl = eq_refl /\ p = p + a1, + a2 := match eq_refl in (_ = e) return (p = p /\ e = e) with + | eq_refl => conj eq_refl eq_refl + end : p = p /\ p = p + ============================ + eq_refl = eq_refl diff --git a/test-suite/output/Cases.v b/test-suite/output/Cases.v index 407489642..17fee3303 100644 --- a/test-suite/output/Cases.v +++ b/test-suite/output/Cases.v @@ -106,3 +106,81 @@ Fail Fixpoint texpDenote t (e:texp t):typeDenote t:= | TBinop t1 t2 _ b e1 e2 => O end. +(* Test notations with local definitions in constructors *) + +Inductive J := D : forall n m, let p := n+m in nat -> J. +Notation "{{ n , m , q }}" := (D n m q). + +Check fun x : J => let '{{n, m, _}} := x in n + m. +Check fun x : J => let '{{n, m, p}} := x in n + m + p. + +(* Cannot use the notation because of the dependency in p *) + +Check fun x => let '(D n m p q) := x in n+m+p+q. + +(* This used to succeed, being interpreted as "let '{{n, m, p}} := ..." *) + +Fail Check fun x : J => let '{{n, m, _}} p := x in n + m + p. + +(* Test use of idents bound to ltac names in a "match" *) + +Lemma lem1 : forall k, k=k :>nat * nat. +let x := fresh "aa" in +let y := fresh "bb" in +let z := fresh "cc" in +let k := fresh "dd" in +refine (fun k : nat * nat => match k as x return x = x with (y,z) => eq_refl end). +Qed. +Print lem1. + +Lemma lem2 : forall k, k=k :> bool. +let x := fresh "aa" in +let y := fresh "bb" in +let z := fresh "cc" in +let k := fresh "dd" in +refine (fun k => if k as x return x = x then eq_refl else eq_refl). +Qed. +Print lem2. + +Lemma lem3 : forall k, k=k :>nat * nat. +let x := fresh "aa" in +let y := fresh "bb" in +let z := fresh "cc" in +let k := fresh "dd" in +refine (fun k : nat * nat => let (y,z) as x return x = x := k in eq_refl). +Qed. +Print lem3. + +Lemma lem4 x : x+0=0. +match goal with |- ?y = _ => pose (match y with 0 => 0 | S n => 0 end) end. +match goal with |- ?y = _ => pose (match y as y with 0 => 0 | S n => 0 end) end. +match goal with |- ?y = _ => pose (match y as y return y=y with 0 => eq_refl | S n => eq_refl end) end. +match goal with |- ?y = _ => pose (match y return y=y with 0 => eq_refl | S n => eq_refl end) end. +match goal with |- ?y + _ = _ => pose (match y with 0 => 0 | S n => 0 end) end. +match goal with |- ?y + _ = _ => pose (match y as y with 0 => 0 | S n => 0 end) end. +match goal with |- ?y + _ = _ => pose (match y as y return y=y with 0 => eq_refl | S n => eq_refl end) end. +match goal with |- ?y + _ = _ => pose (match y return y=y with 0 => eq_refl | S n => eq_refl end) end. +Show. + +Lemma lem5 (p:nat) : eq_refl p = eq_refl p. +let y := fresh "n" in (* Checking that y is hidden *) + let z := fresh "e" in (* Checking that z is hidden *) + match goal with + |- ?y = _ => pose (match y as y in _ = z return y=y /\ z=z with eq_refl => conj eq_refl eq_refl end) + end. +let y := fresh "n" in + let z := fresh "e" in + match goal with + |- ?y = _ => pose (match y in _ = z return y=y /\ z=z with eq_refl => conj eq_refl eq_refl end) + end. +let y := fresh "n" in + let z := fresh "e" in + match goal with + |- eq_refl ?y = _ => pose (match eq_refl y in _ = z return y=y /\ z=z with eq_refl => conj eq_refl eq_refl end) + end. +let p := fresh "p" in + let z := fresh "e" in + match goal with + |- eq_refl ?p = _ => pose (match eq_refl p in _ = z return p=p /\ z=z with eq_refl => conj eq_refl eq_refl end) + end. +Show. diff --git a/test-suite/output/Extraction_matchs_2413.v b/test-suite/output/Extraction_matchs_2413.v index 6c514b16e..1ecd9771e 100644 --- a/test-suite/output/Extraction_matchs_2413.v +++ b/test-suite/output/Extraction_matchs_2413.v @@ -1,5 +1,7 @@ (** Extraction : tests of optimizations of pattern matching *) +Require Coq.extraction.Extraction. + (** First, a few basic tests *) Definition test1 b := 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/Notations3.out b/test-suite/output/Notations3.out index f4ecfd736..ffea0819a 100644 --- a/test-suite/output/Notations3.out +++ b/test-suite/output/Notations3.out @@ -105,3 +105,7 @@ tele (t : Type) '(y, z) (x : t0) := tt ((nat -> nat) * ((nat -> nat) * ((nat -> nat) * (nat -> nat)))))) foo5 x nat x : nat -> nat +fun x : ?A => x === x + : forall x : ?A, x = x +where +?A : [x : ?A |- Type] (x cannot be used) diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v index 71536c68f..250aecafd 100644 --- a/test-suite/output/Notations3.v +++ b/test-suite/output/Notations3.v @@ -148,5 +148,15 @@ Check [ fun x => x+0 ;; fun x => x+1 ;; fun x => x+2 ]. (* Cyprien's part of bug #4765 *) +Section Bug4765. + Notation foo5 x T y := (fun x : T => y). Check foo5 x nat x. + +End Bug4765. + +(**********************************************************************) +(* Test printing of #5526 *) + +Notation "x === x" := (eq_refl x) (only printing, at level 10). +Check (fun x => eq_refl x). 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/output/Record.out b/test-suite/output/Record.out index 36d643a44..d45343fe6 100644 --- a/test-suite/output/Record.out +++ b/test-suite/output/Record.out @@ -14,3 +14,19 @@ build 5 : test_r build_c 5 : test_c +fun '(C _ p) => p + : N -> True +fun '{| T := T |} => T + : N -> Type +fun '(C T p) => (T, p) + : N -> Type * True +fun '{| q := p |} => p + : M -> True +fun '{| U := T |} => T + : M -> Type +fun '{| U := T; q := p |} => (T, p) + : M -> Type * True +fun '{| U := T; a := a; q := p |} => (T, p, a) + : M -> Type * True * nat +fun '{| U := T; a := a; q := p |} => (T, p, a) + : M -> Type * True * nat diff --git a/test-suite/output/Record.v b/test-suite/output/Record.v index 6aa3df983..d9a649fad 100644 --- a/test-suite/output/Record.v +++ b/test-suite/output/Record.v @@ -19,3 +19,15 @@ Check build 5. Check {| field := 5 |}. Check build_r 5. Check build_c 5. + +Record N := C { T : Type; _ : True }. +Check fun x:N => let 'C _ p := x in p. +Check fun x:N => let 'C T _ := x in T. +Check fun x:N => let 'C T p := x in (T,p). + +Record M := D { U : Type; a := 0; q : True }. +Check fun x:M => let 'D T _ p := x in p. +Check fun x:M => let 'D T _ p := x in T. +Check fun x:M => let 'D T p := x in (T,p). +Check fun x:M => let 'D T a p := x in (T,p,a). +Check fun x:M => let '{|U:=T;a:=a;q:=p|} := x in (T,p,a). diff --git a/test-suite/output/ShowMatch.out b/test-suite/output/ShowMatch.out new file mode 100644 index 000000000..e5520b8df --- /dev/null +++ b/test-suite/output/ShowMatch.out @@ -0,0 +1,8 @@ +match # with + | f => + end + +match # with + | A.f => + end + diff --git a/test-suite/output/ShowMatch.v b/test-suite/output/ShowMatch.v new file mode 100644 index 000000000..02b7eada8 --- /dev/null +++ b/test-suite/output/ShowMatch.v @@ -0,0 +1,13 @@ +(* Bug 5546 complained about unqualified constructors in Show Match output, + when qualification is needed to disambiguate them +*) + +Module A. + Inductive foo := f. + Show Match foo. (* no need to disambiguate *) +End A. + +Module B. + Inductive foo := f. + (* local foo shadows A.foo, so constructor "f" needs disambiguation *) + Show Match A.foo. diff --git a/test-suite/save-logs.sh b/test-suite/save-logs.sh index fb8a1c1b0..b61362108 100755 --- a/test-suite/save-logs.sh +++ b/test-suite/save-logs.sh @@ -9,7 +9,7 @@ mkdir "$SAVEDIR" # keep this synced with test-suite/Makefile FAILMARK="==========> FAILURE <==========" -FAILED=$(mktemp) +FAILED=$(mktemp /tmp/coq-check-XXXXX) find . '(' -path ./bugs/opened -prune ')' -o '(' -name '*.log' -exec grep "$FAILMARK" -q '{}' ';' -print0 ')' > "$FAILED" rsync -a --from0 --files-from="$FAILED" . "$SAVEDIR" diff --git a/test-suite/success/Case19.v b/test-suite/success/Case19.v index e59828def..ce98879a5 100644 --- a/test-suite/success/Case19.v +++ b/test-suite/success/Case19.v @@ -17,3 +17,22 @@ Fail exists (fun x => with | eq_refl => eq_refl end). +Abort. + +(* Some tests with ltac matching on building "if" and "let" *) + +Goal forall b c d, (if negb b then c else d) = 0. +intros. +match goal with +|- (if ?b then ?c else ?d) = 0 => transitivity (if b then d else c) +end. +Abort. + +Definition swap {A} {B} '((x,y):A*B) := (y,x). + +Goal forall p, (let '(x,y) := swap p in x + y) = 0. +intros. +match goal with +|- (let '(x,y) := ?p in x + y) = 0 => transitivity (let (x,y) := p in x+y) +end. +Abort. diff --git a/test-suite/success/Compat84.v b/test-suite/success/Compat84.v deleted file mode 100644 index 732a024fc..000000000 --- a/test-suite/success/Compat84.v +++ /dev/null @@ -1,19 +0,0 @@ -(* -*- coq-prog-args: ("-compat" "8.4") -*- *) - -Goal True. - solve [ constructor 1 ]. Undo. - solve [ econstructor 1 ]. Undo. - solve [ constructor ]. Undo. - solve [ econstructor ]. Undo. - solve [ constructor (fail) ]. Undo. - solve [ econstructor (fail) ]. Undo. - split. -Qed. - -Goal False \/ True. - solve [ constructor (constructor) ]. Undo. - solve [ econstructor (econstructor) ]. Undo. - solve [ constructor 2; constructor ]. Undo. - solve [ econstructor 2; econstructor ]. Undo. - right; esplit. -Qed. diff --git a/test-suite/success/Funind.v b/test-suite/success/Funind.v index 3bf97c131..f87f2e2a9 100644 --- a/test-suite/success/Funind.v +++ b/test-suite/success/Funind.v @@ -1,4 +1,6 @@ +Require Import Coq.funind.FunInd. + Definition iszero (n : nat) : bool := match n with | O => true diff --git a/test-suite/success/InversionSigma.v b/test-suite/success/InversionSigma.v new file mode 100644 index 000000000..51f33c7ce --- /dev/null +++ b/test-suite/success/InversionSigma.v @@ -0,0 +1,40 @@ +Section inversion_sigma. + Local Unset Implicit Arguments. + Context A (B : A -> Prop) (C C' : forall a, B a -> Prop) + (D : forall a b, C a b -> Prop) (E : forall a b c, D a b c -> Prop). + + (* Require that, after destructing sigma types and inverting + equalities, we can subst equalities of variables only, and reduce + down to [eq_refl = eq_refl]. *) + Local Ltac test_inversion_sigma := + intros; + repeat match goal with + | [ H : sig _ |- _ ] => destruct H + | [ H : sigT _ |- _ ] => destruct H + | [ H : sig2 _ _ |- _ ] => destruct H + | [ H : sigT2 _ _ |- _ ] => destruct H + end; simpl in *; + inversion_sigma; + repeat match goal with + | [ H : ?x = ?y |- _ ] => is_var x; is_var y; subst x; simpl in * + end; + match goal with + | [ |- eq_refl = eq_refl ] => reflexivity + end. + + Goal forall (x y : { a : A & { b : { b : B a & C a b } & { d : D a (projT1 b) (projT2 b) & E _ _ _ d } } }) + (p : x = y), p = p. + Proof. test_inversion_sigma. Qed. + + Goal forall (x y : { a : A | { b : { b : B a | C a b } | { d : D a (proj1_sig b) (proj2_sig b) | E _ _ _ d } } }) + (p : x = y), p = p. + Proof. test_inversion_sigma. Qed. + + Goal forall (x y : { a : { a : A & B a } & C _ (projT2 a) & C' _ (projT2 a) }) + (p : x = y), p = p. + Proof. test_inversion_sigma. Qed. + + Goal forall (x y : { a : { a : A & B a } | C _ (projT2 a) & C' _ (projT2 a) }) + (p : x = y), p = p. + Proof. test_inversion_sigma. Qed. +End inversion_sigma. 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/RecTutorial.v b/test-suite/success/RecTutorial.v index d8f804246..841940492 100644 --- a/test-suite/success/RecTutorial.v +++ b/test-suite/success/RecTutorial.v @@ -147,6 +147,7 @@ Proof. intros; absurd (p < p); eauto with arith. Qed. +Require Coq.extraction.Extraction. Extraction max. 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/test-suite/success/cbn.v b/test-suite/success/cbn.v new file mode 100644 index 000000000..6aeb05f54 --- /dev/null +++ b/test-suite/success/cbn.v @@ -0,0 +1,18 @@ +(* cbn is able to refold mutual recursive calls *) + +Fixpoint foo (n : nat) := + match n with + | 0 => true + | S n => g n + end +with g (n : nat) : bool := + match n with + | 0 => true + | S n => foo n + end. +Goal forall n, foo (S n) = g n. + intros. cbn. + match goal with + |- g _ = g _ => reflexivity + end. +Qed.
\ No newline at end of file diff --git a/test-suite/success/evars.v b/test-suite/success/evars.v index 82f726fa7..c36313ec1 100644 --- a/test-suite/success/evars.v +++ b/test-suite/success/evars.v @@ -414,4 +414,10 @@ Axiom test : forall P1 P2, P1 = P2 -> P1 -> P2. Import EqNotations. Definition test2 {A B:Type} {H:A=B} (a:A) : B := rew H in a. +(* Check that pre-existing evars are not counted as newly undefined in "set" *) +(* Reported by Théo *) +Goal exists n : nat, n = n -> True. +eexists. +set (H := _ = _). +Abort. diff --git a/test-suite/success/extraction.v b/test-suite/success/extraction.v index 0086e090b..89be14415 100644 --- a/test-suite/success/extraction.v +++ b/test-suite/success/extraction.v @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +Require Coq.extraction.Extraction. Require Import Arith. Require Import List. diff --git a/test-suite/success/extraction_dep.v b/test-suite/success/extraction_dep.v index 11bb25fda..e770cf779 100644 --- a/test-suite/success/extraction_dep.v +++ b/test-suite/success/extraction_dep.v @@ -1,6 +1,8 @@ (** Examples of code elimination inside modules during extraction *) +Require Coq.extraction.Extraction. + (** NB: we should someday check the produced code instead of simply running the commands. *) diff --git a/test-suite/success/extraction_impl.v b/test-suite/success/extraction_impl.v index dfdeff82f..5bf807b1c 100644 --- a/test-suite/success/extraction_impl.v +++ b/test-suite/success/extraction_impl.v @@ -4,6 +4,8 @@ (** NB: we should someday check the produced code instead of simply running the commands. *) +Require Coq.extraction.Extraction. + (** Bug #4243, part 1 *) Inductive dnat : nat -> Type := diff --git a/test-suite/success/extraction_polyprop.v b/test-suite/success/extraction_polyprop.v index 7215bd990..936d838c5 100644 --- a/test-suite/success/extraction_polyprop.v +++ b/test-suite/success/extraction_polyprop.v @@ -3,6 +3,8 @@ code that segfaults. See Table.error_singleton_become_prop or S. Glondu's thesis for more details. *) +Require Coq.extraction.Extraction. + Definition f {X} (p : (nat -> X) * True) : X * nat := (fst p 0, 0). diff --git a/test-suite/success/primitiveproj.v b/test-suite/success/primitiveproj.v index 2fa770494..789854b2d 100644 --- a/test-suite/success/primitiveproj.v +++ b/test-suite/success/primitiveproj.v @@ -181,6 +181,8 @@ Record wrap (A : Type) := { unwrap : A; unwrap2 : A }. Definition term (x : wrap nat) := x.(unwrap). Definition term' (x : wrap nat) := let f := (@unwrap2 nat) in f x. + +Require Coq.extraction.Extraction. Recursive Extraction term term'. (*Unset Printing Primitive Projection Parameters.*) diff --git a/theories/Classes/CRelationClasses.v b/theories/Classes/CRelationClasses.v index 3d7ef01fb..cfe0e08ed 100644 --- a/theories/Classes/CRelationClasses.v +++ b/theories/Classes/CRelationClasses.v @@ -305,9 +305,7 @@ Section Binary. fun x y => sum (R x y) (R' x y). (** Relation equivalence is an equivalence, and subrelation defines a partial order. *) - - Set Automatic Introduction. - + Global Instance relation_equivalence_equivalence : Equivalence relation_equivalence. Proof. split; red; unfold relation_equivalence, iffT. firstorder. diff --git a/theories/Classes/RelationClasses.v b/theories/Classes/RelationClasses.v index 11c204dae..57728d305 100644 --- a/theories/Classes/RelationClasses.v +++ b/theories/Classes/RelationClasses.v @@ -433,9 +433,7 @@ Section Binary. @predicate_union (A::A::Tnil) R R'. (** Relation equivalence is an equivalence, and subrelation defines a partial order. *) - - Set Automatic Introduction. - + Global Instance relation_equivalence_equivalence : Equivalence relation_equivalence. Proof. exact (@predicate_equivalence_equivalence (A::A::Tnil)). Qed. diff --git a/theories/Compat/Coq84.v b/theories/Compat/Coq84.v deleted file mode 100644 index a3e23f91c..000000000 --- a/theories/Compat/Coq84.v +++ /dev/null @@ -1,79 +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 *) -(************************************************************************) - -(** Compatibility file for making Coq act similar to Coq v8.4 *) - -(** Any compatibility changes to make future versions of Coq behave like Coq 8.5 are likely needed to make them behave like Coq 8.4. *) -Require Export Coq.Compat.Coq85. - -(** See https://coq.inria.fr/bugs/show_bug.cgi?id=4319 for updates *) -(** This is required in Coq 8.5 to use the [omega] tactic; in Coq 8.4, it's automatically available. But ZArith_base puts infix ~ at level 7, and we don't want that, so we don't [Import] it. *) -Require Coq.omega.Omega. -Ltac omega := Coq.omega.Omega.omega. - -(** The number of arguments given in [match] statements has changed from 8.4 to 8.5. *) -Global Set Asymmetric Patterns. - -(** The automatic elimination schemes for records were dropped by default in 8.5. This restores the default behavior of Coq 8.4. *) -Global Set Nonrecursive Elimination Schemes. - -(** See bug 3545 *) -Global Set Universal Lemma Under Conjunction. - -(** Feature introduced in 8.5, disabled by default and configurable since 8.6. *) -Global Unset Refolding Reduction. - -(** In 8.4, [constructor (tac)] allowed backtracking across the use of [constructor]; it has been subsumed by [constructor; tac]. *) -Ltac constructor_84_n n := constructor n. -Ltac constructor_84_tac tac := once (constructor; tac). - -Tactic Notation "constructor" := Coq.Init.Notations.constructor. -Tactic Notation "constructor" int_or_var(n) := constructor_84_n n. -Tactic Notation "constructor" "(" tactic(tac) ")" := constructor_84_tac tac. - -(** In 8.4, [econstructor (tac)] allowed backtracking across the use of [econstructor]; it has been subsumed by [econstructor; tac]. *) -Ltac econstructor_84_n n := econstructor n. -Ltac econstructor_84_tac tac := once (econstructor; tac). - -Tactic Notation "econstructor" := Coq.Init.Notations.econstructor. -Tactic Notation "econstructor" int_or_var(n) := econstructor_84_n n. -Tactic Notation "econstructor" "(" tactic(tac) ")" := econstructor_84_tac tac. - -(** Some tactic notations do not factor well with tactics; we add global parsing entries for some tactics that would otherwise be overwritten by custom variants. See https://coq.inria.fr/bugs/show_bug.cgi?id=4392. *) -Tactic Notation "reflexivity" := reflexivity. -Tactic Notation "assumption" := assumption. -Tactic Notation "etransitivity" := etransitivity. -Tactic Notation "cut" constr(c) := cut c. -Tactic Notation "exact_no_check" constr(c) := exact_no_check c. -Tactic Notation "vm_cast_no_check" constr(c) := vm_cast_no_check c. -Tactic Notation "casetype" constr(c) := casetype c. -Tactic Notation "elimtype" constr(c) := elimtype c. -Tactic Notation "lapply" constr(c) := lapply c. -Tactic Notation "transitivity" constr(c) := transitivity c. -Tactic Notation "left" := left. -Tactic Notation "eleft" := eleft. -Tactic Notation "right" := right. -Tactic Notation "eright" := eright. -Tactic Notation "symmetry" := symmetry. -Tactic Notation "split" := split. -Tactic Notation "esplit" := esplit. - -(** Many things now import [PeanoNat] rather than [NPeano], so we require it so that the old absolute names in [NPeano.Nat] are available. *) -Require Coq.Numbers.Natural.Peano.NPeano. - -(** The following coercions were declared by default in Specif.v. *) -Coercion sig_of_sig2 : sig2 >-> sig. -Coercion sigT_of_sigT2 : sigT2 >-> sigT. -Coercion sigT_of_sig : sig >-> sigT. -Coercion sig_of_sigT : sigT >-> sig. -Coercion sigT2_of_sig2 : sig2 >-> sigT2. -Coercion sig2_of_sigT2 : sigT2 >-> sig2. - -(** In 8.4, the statement of admitted lemmas did not depend on the section - variables. *) -Unset Keep Admitted Variables. diff --git a/theories/Compat/Coq85.v b/theories/Compat/Coq85.v index 64ba6b1e3..b30ad1af8 100644 --- a/theories/Compat/Coq85.v +++ b/theories/Compat/Coq85.v @@ -34,3 +34,6 @@ Global Unset Typeclasses Filtered Unification. (** Allow silently letting unification constraints float after a "." *) Global Unset Solve Unification Constraints. + +Require Export Coq.extraction.Extraction. +Require Export Coq.funind.FunInd. diff --git a/theories/FSets/FMapAVL.v b/theories/FSets/FMapAVL.v index c9e5b8dd2..4a790296b 100644 --- a/theories/FSets/FMapAVL.v +++ b/theories/FSets/FMapAVL.v @@ -16,7 +16,7 @@ See the comments at the beginning of FSetAVL for more details. *) -Require Import FMapInterface FMapList ZArith Int. +Require Import FunInd FMapInterface FMapList ZArith Int. Set Implicit Arguments. Unset Strict Implicit. diff --git a/theories/FSets/FMapFullAVL.v b/theories/FSets/FMapFullAVL.v index a7be32328..b8e362f15 100644 --- a/theories/FSets/FMapFullAVL.v +++ b/theories/FSets/FMapFullAVL.v @@ -25,7 +25,7 @@ *) -Require Import Recdef FMapInterface FMapList ZArith Int FMapAVL ROmega. +Require Import FunInd Recdef FMapInterface FMapList ZArith Int FMapAVL ROmega. Set Implicit Arguments. Unset Strict Implicit. diff --git a/theories/FSets/FMapList.v b/theories/FSets/FMapList.v index 5acdb7eb7..aadef476d 100644 --- a/theories/FSets/FMapList.v +++ b/theories/FSets/FMapList.v @@ -12,7 +12,7 @@ [FMapInterface.S] using lists of pairs ordered (increasing) with respect to left projection. *) -Require Import FMapInterface. +Require Import FunInd FMapInterface. Set Implicit Arguments. Unset Strict Implicit. diff --git a/theories/FSets/FMapWeakList.v b/theories/FSets/FMapWeakList.v index 130cbee87..812409702 100644 --- a/theories/FSets/FMapWeakList.v +++ b/theories/FSets/FMapWeakList.v @@ -11,7 +11,7 @@ (** This file proposes an implementation of the non-dependent interface [FMapInterface.WS] using lists of pairs, unordered but without redundancy. *) -Require Import FMapInterface. +Require Import FunInd FMapInterface. Set Implicit Arguments. Unset Strict Implicit. diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v index 3eefe9a84..4db11ae77 100644 --- a/theories/Init/Logic.v +++ b/theories/Init/Logic.v @@ -313,8 +313,8 @@ Arguments eq_ind [A] x P _ y _. Arguments eq_rec [A] x P _ y _. Arguments eq_rect [A] x P _ y _. -Hint Resolve I conj or_introl or_intror : core. -Hint Resolve eq_refl: core. +Hint Resolve I conj or_introl or_intror : core. +Hint Resolve eq_refl: core. Hint Resolve ex_intro ex_intro2: core. Section Logic_lemmas. @@ -504,6 +504,11 @@ Proof. reflexivity. Defined. +Lemma eq_refl_map_distr : forall A B x (f:A->B), f_equal f (eq_refl x) = eq_refl (f x). +Proof. + reflexivity. +Qed. + Lemma eq_trans_map_distr : forall A B x y z (f:A->B) (e:x=y) (e':y=z), f_equal f (eq_trans e e') = eq_trans (f_equal f e) (f_equal f e'). Proof. destruct e'. @@ -522,6 +527,19 @@ destruct e, e'. reflexivity. Defined. +Lemma eq_trans_rew_distr : forall A (P:A -> Type) (x y z:A) (e:x=y) (e':y=z) (k:P x), + rew (eq_trans e e') in k = rew e' in rew e in k. +Proof. + destruct e, e'; reflexivity. +Qed. + +Lemma rew_const : forall A P (x y:A) (e:x=y) (k:P), + rew [fun _ => P] e in k = k. +Proof. + destruct e; reflexivity. +Qed. + + (* Aliases *) Notation sym_eq := eq_sym (compat "8.3"). @@ -575,7 +593,7 @@ Proof. assert (H : x0 = x1) by (transitivity x; [symmetry|]; auto). destruct H. assumption. -Qed. +Qed. Lemma forall_exists_coincide_unique_domain : forall A (P:A->Prop), @@ -587,7 +605,7 @@ Proof. exists x. split; [trivial|]. destruct H with (Q:=fun x'=>x=x') as (_,Huniq). apply Huniq. exists x; auto. -Qed. +Qed. (** * Being inhabited *) @@ -631,3 +649,97 @@ Qed. Declare Left Step iff_stepl. Declare Right Step iff_trans. + +Local Notation "'rew' 'dependent' H 'in' H'" + := (match H with + | eq_refl => H' + end) + (at level 10, H' at level 10, + format "'[' 'rew' 'dependent' '/ ' H in '/' H' ']'"). + +(** Equality for [ex] *) +Section ex. + Local Unset Implicit Arguments. + Definition eq_ex_uncurried {A : Type} (P : A -> Prop) {u1 v1 : A} {u2 : P u1} {v2 : P v1} + (pq : exists p : u1 = v1, rew p in u2 = v2) + : ex_intro P u1 u2 = ex_intro P v1 v2. + Proof. + destruct pq as [p q]. + destruct q; simpl in *. + destruct p; reflexivity. + Qed. + + Definition eq_ex {A : Type} {P : A -> Prop} (u1 v1 : A) (u2 : P u1) (v2 : P v1) + (p : u1 = v1) (q : rew p in u2 = v2) + : ex_intro P u1 u2 = ex_intro P v1 v2 + := eq_ex_uncurried P (ex_intro _ p q). + + Definition eq_ex_hprop {A} {P : A -> Prop} (P_hprop : forall (x : A) (p q : P x), p = q) + (u1 v1 : A) (u2 : P u1) (v2 : P v1) + (p : u1 = v1) + : ex_intro P u1 u2 = ex_intro P v1 v2 + := eq_ex u1 v1 u2 v2 p (P_hprop _ _ _). + + Lemma rew_ex {A x} {P : A -> Type} (Q : forall a, P a -> Prop) (u : exists p, Q x p) {y} (H : x = y) + : rew [fun a => exists p, Q a p] H in u + = match u with + | ex_intro _ u1 u2 + => ex_intro + (Q y) + (rew H in u1) + (rew dependent H in u2) + end. + Proof. + destruct H, u; reflexivity. + Qed. +End ex. + +(** Equality for [ex2] *) +Section ex2. + Local Unset Implicit Arguments. + + Definition eq_ex2_uncurried {A : Type} (P Q : A -> Prop) {u1 v1 : A} + {u2 : P u1} {v2 : P v1} + {u3 : Q u1} {v3 : Q v1} + (pq : exists2 p : u1 = v1, rew p in u2 = v2 & rew p in u3 = v3) + : ex_intro2 P Q u1 u2 u3 = ex_intro2 P Q v1 v2 v3. + Proof. + destruct pq as [p q r]. + destruct r, q, p; simpl in *. + reflexivity. + Qed. + + Definition eq_ex2 {A : Type} {P Q : A -> Prop} + (u1 v1 : A) + (u2 : P u1) (v2 : P v1) + (u3 : Q u1) (v3 : Q v1) + (p : u1 = v1) (q : rew p in u2 = v2) (r : rew p in u3 = v3) + : ex_intro2 P Q u1 u2 u3 = ex_intro2 P Q v1 v2 v3 + := eq_ex2_uncurried P Q (ex_intro2 _ _ p q r). + + Definition eq_ex2_hprop {A} {P Q : A -> Prop} + (P_hprop : forall (x : A) (p q : P x), p = q) + (Q_hprop : forall (x : A) (p q : Q x), p = q) + (u1 v1 : A) (u2 : P u1) (v2 : P v1) (u3 : Q u1) (v3 : Q v1) + (p : u1 = v1) + : ex_intro2 P Q u1 u2 u3 = ex_intro2 P Q v1 v2 v3 + := eq_ex2 u1 v1 u2 v2 u3 v3 p (P_hprop _ _ _) (Q_hprop _ _ _). + + Lemma rew_ex2 {A x} {P : A -> Type} + (Q : forall a, P a -> Prop) + (R : forall a, P a -> Prop) + (u : exists2 p, Q x p & R x p) {y} (H : x = y) + : rew [fun a => exists2 p, Q a p & R a p] H in u + = match u with + | ex_intro2 _ _ u1 u2 u3 + => ex_intro2 + (Q y) + (R y) + (rew H in u1) + (rew dependent H in u2) + (rew dependent H in u3) + end. + Proof. + destruct H, u; reflexivity. + Qed. +End ex2. diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v index e71a8774e..28049e9ee 100644 --- a/theories/Init/Prelude.v +++ b/theories/Init/Prelude.v @@ -18,9 +18,7 @@ Require Export Coq.Init.Tactics. Require Export Coq.Init.Tauto. (* Initially available plugins (+ nat_syntax_plugin loaded in Datatypes) *) -Declare ML Module "extraction_plugin". Declare ML Module "cc_plugin". Declare ML Module "ground_plugin". -Declare ML Module "recdef_plugin". (* Default substrings not considered by queries like SearchAbout *) Add Search Blacklist "_subproof" "_subterm" "Private_". diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v index 43a441fc5..95734991d 100644 --- a/theories/Init/Specif.v +++ b/theories/Init/Specif.v @@ -218,6 +218,407 @@ Proof. intros [[x y]];exists x;exact y. Qed. +(** Equality of sigma types *) +Import EqNotations. +Local Notation "'rew' 'dependent' H 'in' H'" + := (match H with + | eq_refl => H' + end) + (at level 10, H' at level 10, + format "'[' 'rew' 'dependent' '/ ' H in '/' H' ']'"). + +(** Equality for [sigT] *) +Section sigT. + Local Unset Implicit Arguments. + (** Projecting an equality of a pair to equality of the first components *) + Definition projT1_eq {A} {P : A -> Type} {u v : { a : A & P a }} (p : u = v) + : projT1 u = projT1 v + := f_equal (@projT1 _ _) p. + + (** Projecting an equality of a pair to equality of the second components *) + Definition projT2_eq {A} {P : A -> Type} {u v : { a : A & P a }} (p : u = v) + : rew projT1_eq p in projT2 u = projT2 v + := rew dependent p in eq_refl. + + (** Equality of [sigT] is itself a [sigT] (forwards-reasoning version) *) + Definition eq_existT_uncurried {A : Type} {P : A -> Type} {u1 v1 : A} {u2 : P u1} {v2 : P v1} + (pq : { p : u1 = v1 & rew p in u2 = v2 }) + : existT _ u1 u2 = existT _ v1 v2. + Proof. + destruct pq as [p q]. + destruct q; simpl in *. + destruct p; reflexivity. + Defined. + + (** Equality of [sigT] is itself a [sigT] (backwards-reasoning version) *) + Definition eq_sigT_uncurried {A : Type} {P : A -> Type} (u v : { a : A & P a }) + (pq : { p : projT1 u = projT1 v & rew p in projT2 u = projT2 v }) + : u = v. + Proof. + destruct u as [u1 u2], v as [v1 v2]; simpl in *. + apply eq_existT_uncurried; exact pq. + Defined. + + (** Curried version of proving equality of sigma types *) + Definition eq_sigT {A : Type} {P : A -> Type} (u v : { a : A & P a }) + (p : projT1 u = projT1 v) (q : rew p in projT2 u = projT2 v) + : u = v + := eq_sigT_uncurried u v (existT _ p q). + + (** Equality of [sigT] when the property is an hProp *) + Definition eq_sigT_hprop {A P} (P_hprop : forall (x : A) (p q : P x), p = q) + (u v : { a : A & P a }) + (p : projT1 u = projT1 v) + : u = v + := eq_sigT u v p (P_hprop _ _ _). + + (** Equivalence of equality of [sigT] with a [sigT] of equality *) + (** We could actually prove an isomorphism here, and not just [<->], + but for simplicity, we don't. *) + Definition eq_sigT_uncurried_iff {A P} + (u v : { a : A & P a }) + : u = v <-> { p : projT1 u = projT1 v & rew p in projT2 u = projT2 v }. + Proof. + split; [ intro; subst; exists eq_refl; reflexivity | apply eq_sigT_uncurried ]. + Defined. + + (** Induction principle for [@eq (sigT _)] *) + Definition eq_sigT_rect {A P} {u v : { a : A & P a }} (Q : u = v -> Type) + (f : forall p q, Q (eq_sigT u v p q)) + : forall p, Q p. + Proof. intro p; specialize (f (projT1_eq p) (projT2_eq p)); destruct u, p; exact f. Defined. + Definition eq_sigT_rec {A P u v} (Q : u = v :> { a : A & P a } -> Set) := eq_sigT_rect Q. + Definition eq_sigT_ind {A P u v} (Q : u = v :> { a : A & P a } -> Prop) := eq_sigT_rec Q. + + (** Equivalence of equality of [sigT] involving hProps with equality of the first components *) + Definition eq_sigT_hprop_iff {A P} (P_hprop : forall (x : A) (p q : P x), p = q) + (u v : { a : A & P a }) + : u = v <-> (projT1 u = projT1 v) + := conj (fun p => f_equal (@projT1 _ _) p) (eq_sigT_hprop P_hprop u v). + + (** Non-dependent classification of equality of [sigT] *) + Definition eq_sigT_nondep {A B : Type} (u v : { a : A & B }) + (p : projT1 u = projT1 v) (q : projT2 u = projT2 v) + : u = v + := @eq_sigT _ _ u v p (eq_trans (rew_const _ _) q). + + (** Classification of transporting across an equality of [sigT]s *) + Lemma rew_sigT {A x} {P : A -> Type} (Q : forall a, P a -> Prop) (u : { p : P x & Q x p }) {y} (H : x = y) + : rew [fun a => { p : P a & Q a p }] H in u + = existT + (Q y) + (rew H in projT1 u) + (rew dependent H in (projT2 u)). + Proof. + destruct H, u; reflexivity. + Defined. +End sigT. + +(** Equality for [sig] *) +Section sig. + Local Unset Implicit Arguments. + (** Projecting an equality of a pair to equality of the first components *) + Definition proj1_sig_eq {A} {P : A -> Prop} {u v : { a : A | P a }} (p : u = v) + : proj1_sig u = proj1_sig v + := f_equal (@proj1_sig _ _) p. + + (** Projecting an equality of a pair to equality of the second components *) + Definition proj2_sig_eq {A} {P : A -> Prop} {u v : { a : A | P a }} (p : u = v) + : rew proj1_sig_eq p in proj2_sig u = proj2_sig v + := rew dependent p in eq_refl. + + (** Equality of [sig] is itself a [sig] (forwards-reasoning version) *) + Definition eq_exist_uncurried {A : Type} {P : A -> Prop} {u1 v1 : A} {u2 : P u1} {v2 : P v1} + (pq : { p : u1 = v1 | rew p in u2 = v2 }) + : exist _ u1 u2 = exist _ v1 v2. + Proof. + destruct pq as [p q]. + destruct q; simpl in *. + destruct p; reflexivity. + Defined. + + (** Equality of [sig] is itself a [sig] (backwards-reasoning version) *) + Definition eq_sig_uncurried {A : Type} {P : A -> Prop} (u v : { a : A | P a }) + (pq : { p : proj1_sig u = proj1_sig v | rew p in proj2_sig u = proj2_sig v }) + : u = v. + Proof. + destruct u as [u1 u2], v as [v1 v2]; simpl in *. + apply eq_exist_uncurried; exact pq. + Defined. + + (** Curried version of proving equality of sigma types *) + Definition eq_sig {A : Type} {P : A -> Prop} (u v : { a : A | P a }) + (p : proj1_sig u = proj1_sig v) (q : rew p in proj2_sig u = proj2_sig v) + : u = v + := eq_sig_uncurried u v (exist _ p q). + + (** Induction principle for [@eq (sig _)] *) + Definition eq_sig_rect {A P} {u v : { a : A | P a }} (Q : u = v -> Type) + (f : forall p q, Q (eq_sig u v p q)) + : forall p, Q p. + Proof. intro p; specialize (f (proj1_sig_eq p) (proj2_sig_eq p)); destruct u, p; exact f. Defined. + Definition eq_sig_rec {A P u v} (Q : u = v :> { a : A | P a } -> Set) := eq_sig_rect Q. + Definition eq_sig_ind {A P u v} (Q : u = v :> { a : A | P a } -> Prop) := eq_sig_rec Q. + + (** Equality of [sig] when the property is an hProp *) + Definition eq_sig_hprop {A} {P : A -> Prop} (P_hprop : forall (x : A) (p q : P x), p = q) + (u v : { a : A | P a }) + (p : proj1_sig u = proj1_sig v) + : u = v + := eq_sig u v p (P_hprop _ _ _). + + (** Equivalence of equality of [sig] with a [sig] of equality *) + (** We could actually prove an isomorphism here, and not just [<->], + but for simplicity, we don't. *) + Definition eq_sig_uncurried_iff {A} {P : A -> Prop} + (u v : { a : A | P a }) + : u = v <-> { p : proj1_sig u = proj1_sig v | rew p in proj2_sig u = proj2_sig v }. + Proof. + split; [ intro; subst; exists eq_refl; reflexivity | apply eq_sig_uncurried ]. + Defined. + + (** Equivalence of equality of [sig] involving hProps with equality of the first components *) + Definition eq_sig_hprop_iff {A} {P : A -> Prop} (P_hprop : forall (x : A) (p q : P x), p = q) + (u v : { a : A | P a }) + : u = v <-> (proj1_sig u = proj1_sig v) + := conj (fun p => f_equal (@proj1_sig _ _) p) (eq_sig_hprop P_hprop u v). + + Lemma rew_sig {A x} {P : A -> Type} (Q : forall a, P a -> Prop) (u : { p : P x | Q x p }) {y} (H : x = y) + : rew [fun a => { p : P a | Q a p }] H in u + = exist + (Q y) + (rew H in proj1_sig u) + (rew dependent H in proj2_sig u). + Proof. + destruct H, u; reflexivity. + Defined. +End sig. + +(** Equality for [sigT] *) +Section sigT2. + (* We make [sigT_of_sigT2] a coercion so we can use [projT1], [projT2] on [sigT2] *) + Local Coercion sigT_of_sigT2 : sigT2 >-> sigT. + Local Unset Implicit Arguments. + (** Projecting an equality of a pair to equality of the first components *) + Definition sigT_of_sigT2_eq {A} {P Q : A -> Type} {u v : { a : A & P a & Q a }} (p : u = v) + : u = v :> { a : A & P a } + := f_equal _ p. + Definition projT1_of_sigT2_eq {A} {P Q : A -> Type} {u v : { a : A & P a & Q a }} (p : u = v) + : projT1 u = projT1 v + := projT1_eq (sigT_of_sigT2_eq p). + + (** Projecting an equality of a pair to equality of the second components *) + Definition projT2_of_sigT2_eq {A} {P Q : A -> Type} {u v : { a : A & P a & Q a }} (p : u = v) + : rew projT1_of_sigT2_eq p in projT2 u = projT2 v + := rew dependent p in eq_refl. + + (** Projecting an equality of a pair to equality of the third components *) + Definition projT3_eq {A} {P Q : A -> Type} {u v : { a : A & P a & Q a }} (p : u = v) + : rew projT1_of_sigT2_eq p in projT3 u = projT3 v + := rew dependent p in eq_refl. + + (** Equality of [sigT2] is itself a [sigT2] (forwards-reasoning version) *) + Definition eq_existT2_uncurried {A : Type} {P Q : A -> Type} + {u1 v1 : A} {u2 : P u1} {v2 : P v1} {u3 : Q u1} {v3 : Q v1} + (pqr : { p : u1 = v1 + & rew p in u2 = v2 & rew p in u3 = v3 }) + : existT2 _ _ u1 u2 u3 = existT2 _ _ v1 v2 v3. + Proof. + destruct pqr as [p q r]. + destruct r, q, p; simpl. + reflexivity. + Defined. + + (** Equality of [sigT2] is itself a [sigT2] (backwards-reasoning version) *) + Definition eq_sigT2_uncurried {A : Type} {P Q : A -> Type} (u v : { a : A & P a & Q a }) + (pqr : { p : projT1 u = projT1 v + & rew p in projT2 u = projT2 v & rew p in projT3 u = projT3 v }) + : u = v. + Proof. + destruct u as [u1 u2 u3], v as [v1 v2 v3]; simpl in *. + apply eq_existT2_uncurried; exact pqr. + Defined. + + (** Curried version of proving equality of sigma types *) + Definition eq_sigT2 {A : Type} {P Q : A -> Type} (u v : { a : A & P a & Q a }) + (p : projT1 u = projT1 v) + (q : rew p in projT2 u = projT2 v) + (r : rew p in projT3 u = projT3 v) + : u = v + := eq_sigT2_uncurried u v (existT2 _ _ p q r). + + (** Equality of [sigT2] when the second property is an hProp *) + Definition eq_sigT2_hprop {A P Q} (Q_hprop : forall (x : A) (p q : Q x), p = q) + (u v : { a : A & P a & Q a }) + (p : u = v :> { a : A & P a }) + : u = v + := eq_sigT2 u v (projT1_eq p) (projT2_eq p) (Q_hprop _ _ _). + + (** Equivalence of equality of [sigT2] with a [sigT2] of equality *) + (** We could actually prove an isomorphism here, and not just [<->], + but for simplicity, we don't. *) + Definition eq_sigT2_uncurried_iff {A P Q} + (u v : { a : A & P a & Q a }) + : u = v + <-> { p : projT1 u = projT1 v + & rew p in projT2 u = projT2 v & rew p in projT3 u = projT3 v }. + Proof. + split; [ intro; subst; exists eq_refl; reflexivity | apply eq_sigT2_uncurried ]. + Defined. + + (** Induction principle for [@eq (sigT2 _ _)] *) + Definition eq_sigT2_rect {A P Q} {u v : { a : A & P a & Q a }} (R : u = v -> Type) + (f : forall p q r, R (eq_sigT2 u v p q r)) + : forall p, R p. + Proof. + intro p. + specialize (f (projT1_of_sigT2_eq p) (projT2_of_sigT2_eq p) (projT3_eq p)). + destruct u, p; exact f. + Defined. + Definition eq_sigT2_rec {A P Q u v} (R : u = v :> { a : A & P a & Q a } -> Set) := eq_sigT2_rect R. + Definition eq_sigT2_ind {A P Q u v} (R : u = v :> { a : A & P a & Q a } -> Prop) := eq_sigT2_rec R. + + (** Equivalence of equality of [sigT2] involving hProps with equality of the first components *) + Definition eq_sigT2_hprop_iff {A P Q} (Q_hprop : forall (x : A) (p q : Q x), p = q) + (u v : { a : A & P a & Q a }) + : u = v <-> (u = v :> { a : A & P a }) + := conj (fun p => f_equal (@sigT_of_sigT2 _ _ _) p) (eq_sigT2_hprop Q_hprop u v). + + (** Non-dependent classification of equality of [sigT] *) + Definition eq_sigT2_nondep {A B C : Type} (u v : { a : A & B & C }) + (p : projT1 u = projT1 v) (q : projT2 u = projT2 v) (r : projT3 u = projT3 v) + : u = v + := @eq_sigT2 _ _ _ u v p (eq_trans (rew_const _ _) q) (eq_trans (rew_const _ _) r). + + (** Classification of transporting across an equality of [sigT2]s *) + Lemma rew_sigT2 {A x} {P : A -> Type} (Q R : forall a, P a -> Prop) + (u : { p : P x & Q x p & R x p }) + {y} (H : x = y) + : rew [fun a => { p : P a & Q a p & R a p }] H in u + = existT2 + (Q y) + (R y) + (rew H in projT1 u) + (rew dependent H in projT2 u) + (rew dependent H in projT3 u). + Proof. + destruct H, u; reflexivity. + Defined. +End sigT2. + +(** Equality for [sig2] *) +Section sig2. + (* We make [sig_of_sig2] a coercion so we can use [proj1], [proj2] on [sig2] *) + Local Coercion sig_of_sig2 : sig2 >-> sig. + Local Unset Implicit Arguments. + (** Projecting an equality of a pair to equality of the first components *) + Definition sig_of_sig2_eq {A} {P Q : A -> Prop} {u v : { a : A | P a & Q a }} (p : u = v) + : u = v :> { a : A | P a } + := f_equal _ p. + Definition proj1_sig_of_sig2_eq {A} {P Q : A -> Prop} {u v : { a : A | P a & Q a }} (p : u = v) + : proj1_sig u = proj1_sig v + := proj1_sig_eq (sig_of_sig2_eq p). + + (** Projecting an equality of a pair to equality of the second components *) + Definition proj2_sig_of_sig2_eq {A} {P Q : A -> Prop} {u v : { a : A | P a & Q a }} (p : u = v) + : rew proj1_sig_of_sig2_eq p in proj2_sig u = proj2_sig v + := rew dependent p in eq_refl. + + (** Projecting an equality of a pair to equality of the third components *) + Definition proj3_sig_eq {A} {P Q : A -> Prop} {u v : { a : A | P a & Q a }} (p : u = v) + : rew proj1_sig_of_sig2_eq p in proj3_sig u = proj3_sig v + := rew dependent p in eq_refl. + + (** Equality of [sig2] is itself a [sig2] (fowards-reasoning version) *) + Definition eq_exist2_uncurried {A} {P Q : A -> Prop} + {u1 v1 : A} {u2 : P u1} {v2 : P v1} {u3 : Q u1} {v3 : Q v1} + (pqr : { p : u1 = v1 + | rew p in u2 = v2 & rew p in u3 = v3 }) + : exist2 _ _ u1 u2 u3 = exist2 _ _ v1 v2 v3. + Proof. + destruct pqr as [p q r]. + destruct r, q, p; simpl. + reflexivity. + Defined. + + (** Equality of [sig2] is itself a [sig2] (backwards-reasoning version) *) + Definition eq_sig2_uncurried {A} {P Q : A -> Prop} (u v : { a : A | P a & Q a }) + (pqr : { p : proj1_sig u = proj1_sig v + | rew p in proj2_sig u = proj2_sig v & rew p in proj3_sig u = proj3_sig v }) + : u = v. + Proof. + destruct u as [u1 u2 u3], v as [v1 v2 v3]; simpl in *. + apply eq_exist2_uncurried; exact pqr. + Defined. + + (** Curried version of proving equality of sigma types *) + Definition eq_sig2 {A} {P Q : A -> Prop} (u v : { a : A | P a & Q a }) + (p : proj1_sig u = proj1_sig v) + (q : rew p in proj2_sig u = proj2_sig v) + (r : rew p in proj3_sig u = proj3_sig v) + : u = v + := eq_sig2_uncurried u v (exist2 _ _ p q r). + + (** Equality of [sig2] when the second property is an hProp *) + Definition eq_sig2_hprop {A} {P Q : A -> Prop} (Q_hprop : forall (x : A) (p q : Q x), p = q) + (u v : { a : A | P a & Q a }) + (p : u = v :> { a : A | P a }) + : u = v + := eq_sig2 u v (proj1_sig_eq p) (proj2_sig_eq p) (Q_hprop _ _ _). + + (** Equivalence of equality of [sig2] with a [sig2] of equality *) + (** We could actually prove an isomorphism here, and not just [<->], + but for simplicity, we don't. *) + Definition eq_sig2_uncurried_iff {A P Q} + (u v : { a : A | P a & Q a }) + : u = v + <-> { p : proj1_sig u = proj1_sig v + | rew p in proj2_sig u = proj2_sig v & rew p in proj3_sig u = proj3_sig v }. + Proof. + split; [ intro; subst; exists eq_refl; reflexivity | apply eq_sig2_uncurried ]. + Defined. + + (** Induction principle for [@eq (sig2 _ _)] *) + Definition eq_sig2_rect {A P Q} {u v : { a : A | P a & Q a }} (R : u = v -> Type) + (f : forall p q r, R (eq_sig2 u v p q r)) + : forall p, R p. + Proof. + intro p. + specialize (f (proj1_sig_of_sig2_eq p) (proj2_sig_of_sig2_eq p) (proj3_sig_eq p)). + destruct u, p; exact f. + Defined. + Definition eq_sig2_rec {A P Q u v} (R : u = v :> { a : A | P a & Q a } -> Set) := eq_sig2_rect R. + Definition eq_sig2_ind {A P Q u v} (R : u = v :> { a : A | P a & Q a } -> Prop) := eq_sig2_rec R. + + (** Equivalence of equality of [sig2] involving hProps with equality of the first components *) + Definition eq_sig2_hprop_iff {A} {P Q : A -> Prop} (Q_hprop : forall (x : A) (p q : Q x), p = q) + (u v : { a : A | P a & Q a }) + : u = v <-> (u = v :> { a : A | P a }) + := conj (fun p => f_equal (@sig_of_sig2 _ _ _) p) (eq_sig2_hprop Q_hprop u v). + + (** Non-dependent classification of equality of [sig] *) + Definition eq_sig2_nondep {A} {B C : Prop} (u v : @sig2 A (fun _ => B) (fun _ => C)) + (p : proj1_sig u = proj1_sig v) (q : proj2_sig u = proj2_sig v) (r : proj3_sig u = proj3_sig v) + : u = v + := @eq_sig2 _ _ _ u v p (eq_trans (rew_const _ _) q) (eq_trans (rew_const _ _) r). + + (** Classification of transporting across an equality of [sig2]s *) + Lemma rew_sig2 {A x} {P : A -> Type} (Q R : forall a, P a -> Prop) + (u : { p : P x | Q x p & R x p }) + {y} (H : x = y) + : rew [fun a => { p : P a | Q a p & R a p }] H in u + = exist2 + (Q y) + (R y) + (rew H in proj1_sig u) + (rew dependent H in proj2_sig u) + (rew dependent H in proj3_sig u). + Proof. + destruct H, u; reflexivity. + Defined. +End sig2. + + (** [sumbool] is a boolean type equipped with the justification of their value *) diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v index 7a846cd1b..aab385ef7 100644 --- a/theories/Init/Tactics.v +++ b/theories/Init/Tactics.v @@ -243,3 +243,66 @@ with the actual [dependent induction] tactic. *) Tactic Notation "dependent" "induction" ident(H) := fail "To use dependent induction, first [Require Import Coq.Program.Equality.]". + +(** *** [inversion_sigma] *) +(** The built-in [inversion] will frequently leave equalities of + dependent pairs. When the first type in the pair is an hProp or + otherwise simplifies, [inversion_sigma] is useful; it will replace + the equality of pairs with a pair of equalities, one involving a + term casted along the other. This might also prove useful for + writing a version of [inversion] / [dependent destruction] which + does not lose information, i.e., does not turn a goal which is + provable into one which requires axiom K / UIP. *) +Ltac simpl_proj_exist_in H := + repeat match type of H with + | context G[proj1_sig (exist _ ?x ?p)] + => let G' := context G[x] in change G' in H + | context G[proj2_sig (exist _ ?x ?p)] + => let G' := context G[p] in change G' in H + | context G[projT1 (existT _ ?x ?p)] + => let G' := context G[x] in change G' in H + | context G[projT2 (existT _ ?x ?p)] + => let G' := context G[p] in change G' in H + | context G[proj3_sig (exist2 _ _ ?x ?p ?q)] + => let G' := context G[q] in change G' in H + | context G[projT3 (existT2 _ _ ?x ?p ?q)] + => let G' := context G[q] in change G' in H + | context G[sig_of_sig2 (@exist2 ?A ?P ?Q ?x ?p ?q)] + => let G' := context G[@exist A P x p] in change G' in H + | context G[sigT_of_sigT2 (@existT2 ?A ?P ?Q ?x ?p ?q)] + => let G' := context G[@existT A P x p] in change G' in H + end. +Ltac induction_sigma_in_using H rect := + let H0 := fresh H in + let H1 := fresh H in + induction H as [H0 H1] using (rect _ _ _ _); + simpl_proj_exist_in H0; + simpl_proj_exist_in H1. +Ltac induction_sigma2_in_using H rect := + let H0 := fresh H in + let H1 := fresh H in + let H2 := fresh H in + induction H as [H0 H1 H2] using (rect _ _ _ _ _); + simpl_proj_exist_in H0; + simpl_proj_exist_in H1; + simpl_proj_exist_in H2. +Ltac inversion_sigma_step := + match goal with + | [ H : _ = exist _ _ _ |- _ ] + => induction_sigma_in_using H @eq_sig_rect + | [ H : _ = existT _ _ _ |- _ ] + => induction_sigma_in_using H @eq_sigT_rect + | [ H : exist _ _ _ = _ |- _ ] + => induction_sigma_in_using H @eq_sig_rect + | [ H : existT _ _ _ = _ |- _ ] + => induction_sigma_in_using H @eq_sigT_rect + | [ H : _ = exist2 _ _ _ _ _ |- _ ] + => induction_sigma2_in_using H @eq_sig2_rect + | [ H : _ = existT2 _ _ _ _ _ |- _ ] + => induction_sigma2_in_using H @eq_sigT2_rect + | [ H : exist2 _ _ _ _ _ = _ |- _ ] + => induction_sigma_in_using H @eq_sig2_rect + | [ H : existT2 _ _ _ _ _ = _ |- _ ] + => induction_sigma_in_using H @eq_sigT2_rect + end. +Ltac inversion_sigma := repeat inversion_sigma_step. diff --git a/theories/Logic/vo.itarget b/theories/Logic/vo.itarget deleted file mode 100644 index 5eba0b623..000000000 --- a/theories/Logic/vo.itarget +++ /dev/null @@ -1,35 +0,0 @@ -Berardi.vo -PropExtensionalityFacts.vo -ChoiceFacts.vo -ClassicalChoice.vo -ClassicalDescription.vo -ClassicalEpsilon.vo -ClassicalFacts.vo -Classical_Pred_Type.vo -Classical_Prop.vo -ClassicalUniqueChoice.vo -Classical.vo -ConstructiveEpsilon.vo -Decidable.vo -Description.vo -Diaconescu.vo -Epsilon.vo -Eqdep_dec.vo -EqdepFacts.vo -Eqdep.vo -WeakFan.vo -WKL.vo -FunctionalExtensionality.vo -ExtensionalityFacts.vo -ExtensionalFunctionRepresentative.vo -Hurkens.vo -IndefiniteDescription.vo -JMeq.vo -ProofIrrelevanceFacts.vo -ProofIrrelevance.vo -PropFacts.vo -PropExtensionality.vo -RelationalChoice.vo -SetIsType.vo -SetoidChoice.vo -FinFun.vo diff --git a/theories/MSets/MSetAVL.v b/theories/MSets/MSetAVL.v index a3c265a21..b30cb6b56 100644 --- a/theories/MSets/MSetAVL.v +++ b/theories/MSets/MSetAVL.v @@ -31,7 +31,7 @@ code after extraction. *) -Require Import MSetInterface MSetGenTree BinInt Int. +Require Import FunInd MSetInterface MSetGenTree BinInt Int. Set Implicit Arguments. Unset Strict Implicit. diff --git a/theories/MSets/MSetGenTree.v b/theories/MSets/MSetGenTree.v index 154c2384c..036ff1aa4 100644 --- a/theories/MSets/MSetGenTree.v +++ b/theories/MSets/MSetGenTree.v @@ -27,7 +27,7 @@ - min_elt max_elt choose *) -Require Import Orders OrdersFacts MSetInterface PeanoNat. +Require Import FunInd Orders OrdersFacts MSetInterface PeanoNat. Local Open Scope list_scope. Local Open Scope lazy_bool_scope. 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. diff --git a/theories/Program/Wf.v b/theories/Program/Wf.v index c490ea516..6e51f6187 100644 --- a/theories/Program/Wf.v +++ b/theories/Program/Wf.v @@ -69,6 +69,7 @@ Section Well_founded. End Well_founded. +Require Coq.extraction.Extraction. Extraction Inline Fix_F_sub Fix_sub. Set Implicit Arguments. diff --git a/theories/QArith/Qcabs.v b/theories/QArith/Qcabs.v index c0ababfff..e433ecffa 100644 --- a/theories/QArith/Qcabs.v +++ b/theories/QArith/Qcabs.v @@ -22,7 +22,7 @@ Lemma Qcabs_canon (x : Q) : Qred x = x -> Qred (Qabs x) = Qabs x. Proof. intros H; now rewrite (Qred_abs x), H. Qed. Definition Qcabs (x:Qc) : Qc := {| canon := Qcabs_canon x (canon x) |}. -Notation "[ q ]" := (Qcabs q) (q at next level, format "[ q ]") : Qc_scope. +Notation "[ q ]" := (Qcabs q) : Qc_scope. Ltac Qc_unfolds := unfold Qcabs, Qcminus, Qcopp, Qcplus, Qcmult, Qcle, Q2Qc, this. diff --git a/theories/Reals/SeqProp.v b/theories/Reals/SeqProp.v index 3697999f7..6a5233b64 100644 --- a/theories/Reals/SeqProp.v +++ b/theories/Reals/SeqProp.v @@ -150,7 +150,7 @@ Definition sequence_lb (Un:nat -> R) (pr:has_lb Un) (* Compatibility *) Notation sequence_majorant := sequence_ub (only parsing). Notation sequence_minorant := sequence_lb (only parsing). -Unset Standard Proposition Elimination Names. + Lemma Wn_decreasing : forall (Un:nat -> R) (pr:has_ub Un), Un_decreasing (sequence_ub Un pr). Proof. diff --git a/tools/CoqMakefile.in b/tools/CoqMakefile.in index c25ad1f37..5e223a0b4 100644 --- a/tools/CoqMakefile.in +++ b/tools/CoqMakefile.in @@ -115,6 +115,17 @@ TIMER=$(if $(TIMED), $(STDTIME), $(TIMECMD)) OPT?= +# The DYNOBJ and DYNLIB variables are used by "coqdep -dyndep var" in .v.d +ifeq '$(OPT)' '-byte' +USEBYTE:=true +DYNOBJ:=.cmo +DYNLIB:=.cma +else +USEBYTE:= +DYNOBJ:=.cmxs +DYNLIB:=.cmxs +endif + COQFLAGS?=-q $(OPT) $(COQLIBS) $(OTHERFLAGS) COQCHKFLAGS?=-silent -o $(COQLIBS) COQDOCFLAGS?=-interpolate -utf8 $(COQLIBS_NOML) @@ -213,7 +224,6 @@ CMOFILESTOINSTALL = $(filter-out $(addsuffix .cmo,$(PACKEDFILES)),$(CMOFILES)) OBJFILES = $(call vo_to_obj,$(VOFILES)) ALLNATIVEFILES = \ $(OBJFILES:.o=.cmi) \ - $(OBJFILES:.o=.cmo) \ $(OBJFILES:.o=.cmx) \ $(OBJFILES:.o=.cmxs) # trick: wildcard filters out non-existing files @@ -223,8 +233,9 @@ FILESTOINSTALL = \ $(VFILES) \ $(GLOBFILES) \ $(NATIVEFILESTOINSTALL) \ + $(CMIFILESTOINSTALL) +BYTEFILESTOINSTALL = \ $(CMOFILESTOINSTALL) \ - $(CMIFILESTOINSTALL) \ $(CMAFILES) ifeq '$(HASNATDYNLINK)' 'true' DO_NATDYNLINK = yes @@ -256,9 +267,15 @@ post-all:: @# Extension point .PHONY: post-all -real-all: $(VOFILES) $(CMOFILES) $(CMAFILES) $(if $(DO_NATDYNLINK),$(CMXSFILES)) +real-all: $(VOFILES) $(if $(USEBYTE),bytefiles,optfiles) .PHONY: real-all +bytefiles: $(CMOFILES) $(CMAFILES) +.PHONY: bytefiles + +optfiles: $(if $(DO_NATDYNLINK),$(CMXSFILES)) +.PHONY: optfiles + # FIXME, see Ralph's bugreport quick: $(VOFILES:.vo=.vio) .PHONY: quick @@ -350,6 +367,18 @@ install-extra:: @# Extension point .PHONY: install install-extra +install-byte: + $(HIDE)for f in $(BYTEFILESTOINSTALL); do\ + df="`$(COQMF_MAKEFILE) -destination-of "$$f" $(COQLIBS)`";\ + if [ -z "$$df" ]; then\ + echo SKIP "$$f" since it has no logical path;\ + else\ + install -d "$(DESTDIR)$(COQLIBINSTALL)/$$df"; \ + install -m 0644 "$$f" "$(DESTDIR)$(COQLIBINSTALL)/$$df"; \ + echo INSTALL "$$f" "$(DESTDIR)$(COQLIBINSTALL)/$$df";\ + fi;\ + done + install-doc:: html mlihtml @# Extension point $(HIDE)install -d "$(DESTDIR)$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" @@ -561,7 +590,7 @@ $(addsuffix .d,$(MLPACKFILES)): %.mlpack.d: %.mlpack $(addsuffix .d,$(VFILES)): %.v.d: %.v $(SHOW)'COQDEP $<' - $(HIDE)$(COQDEP) $(COQLIBS) -c "$<" $(redir_if_ok) + $(HIDE)$(COQDEP) $(COQLIBS) -dyndep var -c "$<" $(redir_if_ok) # Misc ######################################################################## diff --git a/tools/coq_makefile.ml b/tools/coq_makefile.ml index 8e2f75fc9..e4f135977 100644 --- a/tools/coq_makefile.ml +++ b/tools/coq_makefile.ml @@ -27,11 +27,6 @@ let rec print_prefix_list sep = function | x :: l -> print sep; print x; print_prefix_list sep l | [] -> () -(* These are the Coq library directories that are used for - * plugin development - *) -let lib_dirs = Envars.coq_src_subdirs - let usage () = output_string stderr "Usage summary:\ \n\ @@ -73,6 +68,7 @@ let usage () = \n[-f file]: take the contents of file as arguments\ \n[-o file]: output should go in file file\ \n Output file outside the current directory is forbidden.\ +\n[-bypass-API]: when compiling plugins, bypass Coq API\ \n[-h]: print this usage summary\ \n[--help]: equivalent to [-h]\n"; exit 1 @@ -197,9 +193,12 @@ let generate_conf_includes oc { ml_includes; r_includes; q_includes } = (S.concat " " (map (fun ({ path },l) -> dash2 "R" path l) r_includes)) ;; -let generate_conf_coq_config oc args = +let generate_conf_coq_config oc args bypass_API = section oc "Coq configuration."; - Envars.print_config ~prefix_var_name:"COQMF_" oc; + let src_dirs = if bypass_API + then Coq_config.all_src_dirs + else Coq_config.api_dirs @ Coq_config.plugins_dirs in + Envars.print_config ~prefix_var_name:"COQMF_" oc src_dirs; fprintf oc "COQMF_MAKEFILE=%s\n" (quote (List.hd args)); ;; @@ -258,7 +257,7 @@ let generate_conf oc project args = fprintf oc "# %s\n\n" (String.concat " " (List.map quote args)); generate_conf_files oc project; generate_conf_includes oc project; - generate_conf_coq_config oc args; + generate_conf_coq_config oc args project.bypass_API; generate_conf_defs oc project; generate_conf_doc oc project; generate_conf_extra_target oc project.extra_targets; diff --git a/tools/coqc.ml b/tools/coqc.ml index 240531f12..c1f0182d9 100644 --- a/tools/coqc.ml +++ b/tools/coqc.ml @@ -83,7 +83,7 @@ let parse_args () = | ("-config" | "--config") :: _ -> Envars.set_coqlib ~fail:(fun x -> x); - Envars.print_config stdout; + Envars.print_config stdout Coq_config.all_src_dirs; exit 0 |"--print-version" :: _ -> diff --git a/tools/coqdep.ml b/tools/coqdep.ml index 044399544..cba9c3eb0 100644 --- a/tools/coqdep.ml +++ b/tools/coqdep.ml @@ -320,19 +320,25 @@ let treat_coq_file chan = List.fold_left (fun accu v -> mark_v_done from accu v) acc strl | Declare sl -> let declare suff dir s = - let base = file_name s dir in - let opt = if !option_natdynlk then " " ^ base ^ ".cmxs" else "" in - (escape base, suff ^ opt) + let base = escape (file_name s dir) in + match !option_dynlink with + | No -> [] + | Byte -> [base,suff] + | Opt -> [base,".cmxs"] + | Both -> [base,suff; base,".cmxs"] + | Variable -> + if suff=".cmo" then [base,"$(DYNOBJ)"] + else [base,"$(DYNLIB)"] in let decl acc str = let s = basename_noext str in if not (StrSet.mem s !deja_vu_ml) then let () = deja_vu_ml := StrSet.add s !deja_vu_ml in match search_mllib_known s with - | Some mldir -> (declare ".cma" mldir s) :: acc + | Some mldir -> (declare ".cma" mldir s) @ acc | None -> match search_ml_known s with - | Some mldir -> (declare ".cmo" mldir s) :: acc + | Some mldir -> (declare ".cmo" mldir s) @ acc | None -> acc else acc in @@ -449,6 +455,7 @@ let usage () = eprintf " -coqlib dir : set the coq standard library directory\n"; eprintf " -suffix s : \n"; eprintf " -slash : deprecated, no effect\n"; + eprintf " -dyndep (opt|byte|both|no|var) : set how dependencies over ML modules are printed"; exit 1 let split_period = Str.split (Str.regexp (Str.quote ".")) @@ -476,17 +483,22 @@ let rec parse = function | "-slash" :: ll -> Printf.eprintf "warning: option -slash has no effect and is deprecated.\n"; parse ll + | "-dyndep" :: "no" :: ll -> option_dynlink := No; parse ll + | "-dyndep" :: "opt" :: ll -> option_dynlink := Opt; parse ll + | "-dyndep" :: "byte" :: ll -> option_dynlink := Byte; parse ll + | "-dyndep" :: "both" :: ll -> option_dynlink := Both; parse ll + | "-dyndep" :: "var" :: ll -> option_dynlink := Variable; parse ll | ("-h"|"--help"|"-help") :: _ -> usage () | f :: ll -> treat_file None f; parse ll | [] -> () let coqdep () = if Array.length Sys.argv < 2 then usage (); + if not Coq_config.has_natdynlink then option_dynlink := No; parse (List.tl (Array.to_list Sys.argv)); (* Add current dir with empty logical path if not set by options above. *) (try ignore (Coqdep_common.find_dir_logpath (Sys.getcwd())) with Not_found -> add_norec_dir_import add_known "." []); - if not Coq_config.has_natdynlink then option_natdynlk := false; (* NOTE: These directories are searched from last to first *) if !option_boot then begin add_rec_dir_import add_known "theories" ["Coq"]; diff --git a/tools/coqdep_boot.ml b/tools/coqdep_boot.ml index 6fc826833..25f62d2be 100644 --- a/tools/coqdep_boot.ml +++ b/tools/coqdep_boot.ml @@ -16,7 +16,11 @@ open Coqdep_common *) let rec parse = function - | "-natdynlink" :: "no" :: ll -> option_natdynlk := false; parse ll + | "-dyndep" :: "no" :: ll -> option_dynlink := No; parse ll + | "-dyndep" :: "opt" :: ll -> option_dynlink := Opt; parse ll + | "-dyndep" :: "byte" :: ll -> option_dynlink := Byte; parse ll + | "-dyndep" :: "both" :: ll -> option_dynlink := Both; parse ll + | "-dyndep" :: "var" :: ll -> option_dynlink := Variable; parse ll | "-c" :: ll -> option_c := true; parse ll | "-boot" :: ll -> parse ll (* We're already in boot mode by default *) | "-mldep" :: ocamldep :: ll -> diff --git a/tools/coqdep_common.ml b/tools/coqdep_common.ml index f5e93527c..bf8bcd0c4 100644 --- a/tools/coqdep_common.ml +++ b/tools/coqdep_common.ml @@ -15,7 +15,7 @@ open Minisys behavior is the one of [coqdep -boot]. Its only dependencies are [Coqdep_lexer], [Unix] and [Minisys], and it should stay so. If it need someday some additional information, pass it via - options (see for instance [option_natdynlk] below). + options (see for instance [option_dynlink] below). *) module StrSet = Set.Make(String) @@ -26,9 +26,11 @@ module StrListMap = Map.Make(StrList) let stderr = Pervasives.stderr let stdout = Pervasives.stdout +type dynlink = Opt | Byte | Both | No | Variable + let option_c = ref false let option_noglob = ref false -let option_natdynlk = ref true +let option_dynlink = ref Both let option_boot = ref false let option_mldep = ref None @@ -383,10 +385,16 @@ let rec traite_fichier_Coq suffixe verbose f = end) strl | Declare sl -> let declare suff dir s = - let base = file_name s dir in - let opt = if !option_natdynlk then " "^base^".cmxs" else "" in - printf " %s%s%s" (escape base) suff opt - in + let base = escape (file_name s dir) in + match !option_dynlink with + | No -> () + | Byte -> printf " %s%s" base suff + | Opt -> printf " %s.cmxs" base + | Both -> printf " %s%s %s.cmxs" base suff base + | Variable -> + printf " %s%s" base + (if suff=".cmo" then "$(DYNOBJ)" else "$(DYNLIB)") + in let decl str = let s = basename_noext str in if not (StrSet.mem s !deja_vu_ml) then begin diff --git a/tools/coqdep_common.mli b/tools/coqdep_common.mli index 10da0240d..8c1787d31 100644 --- a/tools/coqdep_common.mli +++ b/tools/coqdep_common.mli @@ -19,7 +19,10 @@ val find_dir_logpath: string -> string list val option_c : bool ref val option_noglob : bool ref val option_boot : bool ref -val option_natdynlk : bool ref + +type dynlink = Opt | Byte | Both | No | Variable + +val option_dynlink : dynlink ref val option_mldep : string option ref val norec_dirs : StrSet.t ref val suffixe : string ref diff --git a/tools/coqdoc/cdglobals.mli b/tools/coqdoc/cdglobals.mli new file mode 100644 index 000000000..2c9b3fb8e --- /dev/null +++ b/tools/coqdoc/cdglobals.mli @@ -0,0 +1,49 @@ +type target_language = LaTeX | HTML | TeXmacs | Raw +val target_language : target_language ref +type output_t = StdOut | MultFiles | File of string +val output_dir : string ref +val out_to : output_t ref +val out_channel : out_channel ref +val ( / ) : string -> string -> string +val coqdoc_out : string -> string +val open_out_file : string -> unit +val close_out_file : unit -> unit +type glob_source_t = NoGlob | DotGlob | GlobFile of string +val glob_source : glob_source_t ref +val normalize_path : string -> string +val normalize_filename : string -> string * string +val guess_coqlib : unit -> string +val header_trailer : bool ref +val header_file : string ref +val header_file_spec : bool ref +val footer_file : string ref +val footer_file_spec : bool ref +val quiet : bool ref +val light : bool ref +val gallina : bool ref +val short : bool ref +val index : bool ref +val multi_index : bool ref +val index_name : string ref +val toc : bool ref +val page_title : string ref +val title : string ref +val externals : bool ref +val coqlib : string ref +val coqlib_path : string ref +val raw_comments : bool ref +val parse_comments : bool ref +val plain_comments : bool ref +val toc_depth : int option ref +val lib_name : string ref +val lib_subtitles : bool ref +val interpolate : bool ref +val inline_notmono : bool ref +val charset : string ref +val inputenc : string ref +val latin1 : bool ref +val utf8 : bool ref +val set_latin1 : unit -> unit +val set_utf8 : unit -> unit +type coq_module = string +type file = Vernac_file of string * coq_module | Latex_file of string diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml index 8fca30268..f36d0c348 100644 --- a/toplevel/coqinit.ml +++ b/toplevel/coqinit.ml @@ -124,16 +124,14 @@ let init_ocaml_path () = Mltop.add_ml_dir (List.fold_left (/) Envars.coqroot [dl]) in Mltop.add_ml_dir (Envars.coqlib ()); - List.iter add_subdir Envars.coq_src_subdirs + List.iter add_subdir Coq_config.all_src_dirs -let get_compat_version = function +let get_compat_version ?(allow_old = true) = function | "8.7" -> Flags.Current | "8.6" -> Flags.V8_6 | "8.5" -> Flags.V8_5 - | "8.4" -> Flags.V8_4 - | "8.3" -> Flags.V8_3 - | "8.2" -> Flags.V8_2 - | ("8.1" | "8.0") as s -> + | ("8.4" | "8.3" | "8.2" | "8.1" | "8.0") as s -> + if allow_old then Flags.VOld else CErrors.user_err ~hdr:"get_compat_version" (str "Compatibility with version " ++ str s ++ str " not supported.") | s -> CErrors.user_err ~hdr:"get_compat_version" diff --git a/toplevel/coqinit.mli b/toplevel/coqinit.mli index 3b42289ee..787dfb61a 100644 --- a/toplevel/coqinit.mli +++ b/toplevel/coqinit.mli @@ -25,4 +25,4 @@ val init_library_roots : unit -> unit val init_ocaml_path : unit -> unit -val get_compat_version : string -> Flags.compat_version +val get_compat_version : ?allow_old:bool -> string -> Flags.compat_version diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml index 908786565..0b0ef6717 100644 --- a/toplevel/coqloop.ml +++ b/toplevel/coqloop.ml @@ -187,7 +187,7 @@ end from cycling. *) let make_prompt () = try - (Names.Id.to_string (Pfedit.get_current_proof_name ())) ^ " < " + (Names.Id.to_string (Proof_global.get_current_proof_name ())) ^ " < " with Proof_global.NoCurrentProof -> "Coq < " diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index 26ee413fb..5f0716fd9 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -205,9 +205,9 @@ let require () = let add_compat_require v = match v with - | Flags.V8_4 -> add_require "Coq.Compat.Coq84" | Flags.V8_5 -> add_require "Coq.Compat.Coq85" - | _ -> () + | Flags.V8_6 -> add_require "Coq.Compat.Coq86" + | Flags.VOld | Flags.Current -> () let compile_list = ref ([] : (bool * string) list) @@ -514,7 +514,9 @@ let parse_args arglist = |"-async-proofs-delegation-threshold" -> Flags.async_proofs_delegation_threshold:= get_float opt (next ()) |"-worker-id" -> set_worker_id opt (next ()) - |"-compat" -> let v = get_compat_version (next ()) in Flags.compat_version := v; add_compat_require v + |"-compat" -> + let v = get_compat_version ~allow_old:false (next ()) in + Flags.compat_version := v; add_compat_require v |"-compile" -> add_compile false (next ()) |"-compile-verbose" -> add_compile true (next ()) |"-dump-glob" -> Dumpglob.dump_into_file (next ()); glob_opt := true @@ -621,7 +623,7 @@ let init_toplevel arglist = Spawned.init_channels (); Envars.set_coqlib ~fail:(fun msg -> CErrors.user_err Pp.(str msg)); if !print_where then (print_endline(Envars.coqlib ()); exit(exitcode ())); - if !print_config then (Envars.print_config stdout; exit (exitcode ())); + if !print_config then (Envars.print_config stdout Coq_config.all_src_dirs; exit (exitcode ())); if !print_tags then (print_style_tags (); exit (exitcode ())); if !filter_opts then (print_string (String.concat "\n" extras); exit 0); init_load_path (); diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index a61ade784..92730c14d 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -111,7 +111,7 @@ let pr_open_cur_subgoals () = with Proof_global.NoCurrentProof -> Pp.str "" let vernac_error msg = - Format.fprintf !Topfmt.err_ft "@[%a@]%!" Pp.pp_with msg; + Topfmt.std_logger Feedback.Error msg; flush_all (); exit 1 @@ -285,7 +285,7 @@ let ensure_exists f = (* Compile a vernac file *) let compile verbosely f = let check_pending_proofs () = - let pfs = Pfedit.get_all_proof_names () in + let pfs = Proof_global.get_all_proof_names () in if not (List.is_empty pfs) then vernac_error (str "There are pending proofs") in match !Flags.compilation_mode with diff --git a/vernac/classes.ml b/vernac/classes.ml index dc5ce1a53..aba61146c 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -341,7 +341,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p if not (Option.is_empty term) then let init_refine = Tacticals.New.tclTHENLIST [ - Refine.refine (fun evm -> (evm,EConstr.of_constr (Option.get term))); + Refine.refine ~typecheck:false (fun evm -> (evm,EConstr.of_constr (Option.get term))); Proofview.Unsafe.tclNEWGOALS gls; Tactics.New.reduce_after_refine; ] @@ -386,7 +386,13 @@ let context poly l = let ctx = Univ.ContextSet.to_context !uctx in (* Declare the universe context once *) let () = uctx := Univ.ContextSet.empty in - let decl = (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) in + let decl = match b with + | None -> + (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) + | Some b -> + let entry = Declare.definition_entry ~poly ~univs:ctx ~types:t b in + (DefinitionEntry entry, IsAssumption Logical) + in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in match class_of_constr !evars (EConstr.of_constr t) with | Some (rels, ((tc,_), args) as _cl) -> @@ -402,9 +408,17 @@ let context poly l = in let impl = List.exists test impls in let decl = (Discharge, poly, Definitional) in - let nstatus = + let nstatus = match b with + | None -> pi3 (Command.declare_assumption false decl (t, !uctx) [] [] impl Vernacexpr.NoInline (Loc.tag id)) + | Some b -> + let ctx = Univ.ContextSet.to_context !uctx in + let decl = (Discharge, poly, Definition) in + let entry = Declare.definition_entry ~poly ~univs:ctx ~types:t b in + let hook = Lemmas.mk_hook (fun _ gr -> gr) in + let _ = Command.declare_definition id decl entry [] [] hook in + Lib.sections_are_opened () || Lib.is_modtype_strict () in let () = uctx := Univ.ContextSet.empty in status && nstatus diff --git a/vernac/command.ml b/vernac/command.ml index b1425d703..998e7803e 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -187,7 +187,7 @@ let declare_definition ident (local, p, k) ce pl imps hook = let () = definition_message ident in let gr = VarRef ident in let () = maybe_declare_manual_implicits false gr imps in - let () = if Pfedit.refining () then + let () = if Proof_global.there_are_pending_proofs () then warn_definition_not_visible ident in gr @@ -233,7 +233,7 @@ match local with let _ = declare_variable ident decl in let () = assumption_message ident in let () = - if not !Flags.quiet && Pfedit.refining () then + if not !Flags.quiet && Proof_global.there_are_pending_proofs () then Feedback.msg_info (str"Variable" ++ spc () ++ pr_id ident ++ strbrk " is not visible from current goals") in diff --git a/vernac/command.mli b/vernac/command.mli index 9bbc2fdac..2a52d9bcb 100644 --- a/vernac/command.mli +++ b/vernac/command.mli @@ -15,7 +15,6 @@ open Vernacexpr open Constrexpr open Decl_kinds open Redexpr -open Pfedit (** This file is about the interpretation of raw commands into typed ones and top-level declaration of the main Gallina objects *) @@ -151,7 +150,7 @@ val declare_fixpoint : locality -> polymorphic -> recursive_preentry * lident list option * Evd.evar_universe_context * (Context.Rel.t * Impargs.manual_implicits * int option) list -> - lemma_possible_guards -> decl_notation list -> unit + Proof_global.lemma_possible_guards -> decl_notation list -> unit val declare_cofixpoint : locality -> polymorphic -> recursive_preentry * lident list option * Evd.evar_universe_context * diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml index c2c27eb78..44d6f37cc 100644 --- a/vernac/indschemes.ml +++ b/vernac/indschemes.ml @@ -84,15 +84,8 @@ let _ = optkey = ["Boolean";"Equality";"Schemes"]; optread = (fun () -> !eq_flag) ; optwrite = (fun b -> eq_flag := b) } -let _ = (* compatibility *) - declare_bool_option - { optdepr = true; - optname = "automatic declaration of boolean equality"; - optkey = ["Equality";"Scheme"]; - optread = (fun () -> !eq_flag) ; - optwrite = (fun b -> eq_flag := b) } -let is_eq_flag () = !eq_flag && Flags.version_strictly_greater Flags.V8_2 +let is_eq_flag () = !eq_flag let eq_dec_flag = ref false let _ = diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 77e356eb2..5bf419caf 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -209,7 +209,7 @@ let compute_proof_name locality = function user_err ?loc (pr_id id ++ str " already exists."); id, pl | None -> - next_global_ident_away default_thm_id (Pfedit.get_all_proof_names ()), None + next_global_ident_away default_thm_id (Proof_global.get_all_proof_names ()), None let save_remaining_recthms (locality,p,kind) norm ctx body opaq i ((id,pl),(t_i,(_,imps))) = let t_i = norm t_i in @@ -487,7 +487,7 @@ let save_proof ?proof = function let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in Admitted(id, k, (sec_vars, pi2 k, (typ, ctx), None), universes) | None -> - let pftree = Pfedit.get_pftreestate () in + let pftree = Proof_global.give_me_the_proof () in let id, k, typ = Pfedit.current_proof_statement () in let typ = EConstr.Unsafe.to_constr typ in let universes = Proof.initial_euctx pftree in @@ -496,7 +496,7 @@ let save_proof ?proof = function Proof_global.return_proof ~allow_partial:true () in let sec_vars = if not !keep_admitted_vars then None - else match Pfedit.get_used_variables(), pproofs with + else match Proof_global.get_used_variables(), pproofs with | Some _ as x, _ -> x | None, (pproof, _) :: _ -> let env = Global.env () in @@ -504,7 +504,7 @@ let save_proof ?proof = function let ids_def = Environ.global_vars_set env pproof in Some (Environ.keep_hyps env (Idset.union ids_typ ids_def)) | _ -> None in - let names = Pfedit.get_universe_binders () in + let names = Proof_global.get_universe_binders () in let evd = Evd.from_ctx universes in let binders, ctx = Evd.universe_context ?names evd in Admitted(id,k,(sec_vars, pi2 k, (typ, ctx), None), @@ -519,7 +519,7 @@ let save_proof ?proof = function | Some proof -> proof in (* if the proof is given explicitly, nothing has to be deleted *) - if Option.is_empty proof then Pfedit.delete_current_proof (); + if Option.is_empty proof then Proof_global.discard_current (); Proof_global.(apply_terminator terminator (Proved (is_opaque,idopt,proof_obj))) (* Miscellaneous *) diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli index d06b8fd14..a9c0d99f3 100644 --- a/vernac/lemmas.mli +++ b/vernac/lemmas.mli @@ -9,7 +9,6 @@ open Names open Term open Decl_kinds -open Pfedit type 'a declaration_hook val mk_hook : @@ -21,16 +20,16 @@ val call_hook : (** A hook start_proof calls on the type of the definition being started *) val set_start_hook : (EConstr.types -> unit) -> unit -val start_proof : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> - ?terminator:(lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) -> +val start_proof : Id.t -> ?pl:Proof_global.universe_binders -> goal_kind -> Evd.evar_map -> + ?terminator:(Proof_global.lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) -> ?sign:Environ.named_context_val -> EConstr.types -> - ?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards -> + ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards -> unit declaration_hook -> unit -val start_proof_univs : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> - ?terminator:(lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> Proof_global.proof_terminator) -> +val start_proof_univs : Id.t -> ?pl:Proof_global.universe_binders -> goal_kind -> Evd.evar_map -> + ?terminator:(Proof_global.lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> Proof_global.proof_terminator) -> ?sign:Environ.named_context_val -> EConstr.types -> - ?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards -> + ?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> unit val start_proof_com : @@ -40,8 +39,8 @@ val start_proof_com : val start_proof_with_initialization : goal_kind -> Evd.evar_map -> - (bool * lemma_possible_guards * unit Proofview.tactic list option) option -> - ((Id.t (* name of thm *) * universe_binders option) * + (bool * Proof_global.lemma_possible_guards * unit Proofview.tactic list option) option -> + ((Id.t (* name of thm *) * Proof_global.universe_binders option) * (types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list -> int list option -> unit declaration_hook -> unit diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index 34b9b97d8..a114553cd 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -301,22 +301,22 @@ let is_numeral symbs = | _ -> false -let rec get_notation_vars = function +let rec get_notation_vars onlyprint = function | [] -> [] | NonTerminal id :: sl -> - let vars = get_notation_vars sl in + let vars = get_notation_vars onlyprint sl in if Id.equal id ldots_var then vars else - if Id.List.mem id vars then + (* don't check for nonlinearity if printing only, see Bug 5526 *) + if not onlyprint && Id.List.mem id vars then user_err ~hdr:"Metasyntax.get_notation_vars" (str "Variable " ++ pr_id id ++ str " occurs more than once.") - else - id::vars - | (Terminal _ | Break _) :: sl -> get_notation_vars sl + else id::vars + | (Terminal _ | Break _) :: sl -> get_notation_vars onlyprint sl | SProdList _ :: _ -> assert false -let analyze_notation_tokens l = +let analyze_notation_tokens ~onlyprint l = let l = raw_analyze_notation_tokens l in - let vars = get_notation_vars l in + let vars = get_notation_vars onlyprint l in let recvars,l = interp_list_parser [] l in recvars, List.subtract Id.equal vars (List.map snd recvars), l @@ -1084,12 +1084,12 @@ let compute_syntax_data df modifiers = if onlyprint && onlyparse then user_err (str "A notation cannot be both 'only printing' and 'only parsing'."); let assoc = Option.append mods.assoc (Some NonA) in let toks = split_notation_string df in - let recvars,mainvars,symbols = analyze_notation_tokens toks in + let (recvars,mainvars,symbols) = analyze_notation_tokens ~onlyprint toks in let _ = check_useless_entry_types recvars mainvars mods.etyps in let _ = check_binder_type recvars mods.etyps in (* Notations for interp and grammar *) -let ntn_for_interp = make_notation_key symbols in + let ntn_for_interp = make_notation_key symbols in let symbols' = remove_curly_brackets symbols in let ntn_for_grammar = make_notation_key symbols' in if not onlyprint then check_rule_productivity symbols'; @@ -1333,7 +1333,7 @@ let add_notation_in_scope local df c mods scope = let add_notation_interpretation_core local df ?(impls=empty_internalization_env) c scope onlyparse onlyprint compat = let dfs = split_notation_string df in - let recvars,mainvars,symbs = analyze_notation_tokens dfs in + let (recvars,mainvars,symbs) = analyze_notation_tokens ~onlyprint dfs in (* Recover types of variables and pa/pp rules; redeclare them if needed *) let i_typs, onlyprint = if not (is_numeral symbs) then begin let i_typs,sy_rules,onlyprint' = recover_notation_syntax (make_notation_key symbs) in @@ -1410,7 +1410,7 @@ let add_notation local c ((loc,df),modifiers) sc = let add_notation_extra_printing_rule df k v = let notk = let dfs = split_notation_string df in - let _,_, symbs = analyze_notation_tokens dfs in + let _,_, symbs = analyze_notation_tokens ~onlyprint:true dfs in make_notation_key symbs in Notation.add_notation_extra_printing_rule notk k v diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 6dee95bc5..e03e9b803 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -947,7 +947,7 @@ let rec solve_obligation prg num tac = let hook ctx = Lemmas.mk_hook (obligation_hook prg obl num auto ctx) in let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd (EConstr.of_constr obl.obl_type) ~terminator hook in let _ = Pfedit.by !default_tactic in - Option.iter (fun tac -> Pfedit.set_end_tac tac) tac + Option.iter (fun tac -> Proof_global.set_endline_tactic tac) tac and obligation (user_num, name, typ) tac = let num = pred user_num in diff --git a/vernac/search.ml b/vernac/search.ml index 916015800..0ff78f439 100644 --- a/vernac/search.ml +++ b/vernac/search.ml @@ -142,7 +142,7 @@ module ConstrPriority = struct -(3*(num_symbols t) + size t) let compare (_,_,_,p1) (_,_,_,p2) = - compare p1 p2 + Pervasives.compare p1 p2 end module PriorityQueue = Heap.Functional(ConstrPriority) diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 69492759b..d0f9c7de7 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -15,7 +15,6 @@ open Flags open Names open Nameops open Term -open Pfedit open Tacmach open Constrintern open Prettyp @@ -61,35 +60,25 @@ let show_proof () = let pprf = Proof.partial_proof p in Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_econstr pprf) -let show_node () = - (* spiwack: I'm have little clue what this function used to do. I deactivated it, - could, possibly, be cleaned away. (Feb. 2010) *) - () - -let show_thesis () = CErrors.anomaly (Pp.str "Show Thesis: TODO.") - let show_top_evars () = (* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *) - let pfts = get_pftreestate () in + let pfts = Proof_global.give_me_the_proof () in let gls = Proof.V82.subgoals pfts in let sigma = gls.Evd.sigma in Feedback.msg_notice (pr_evars_int sigma 1 (Evarutil.non_instantiated sigma)) let show_universes () = - let pfts = get_pftreestate () in + let pfts = Proof_global.give_me_the_proof () in let gls = Proof.V82.subgoals pfts in let sigma = gls.Evd.sigma in let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in Feedback.msg_notice (Termops.pr_evar_universe_context (Evd.evar_universe_context sigma)); Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx) -(* Spiwack: proof tree is currently not working *) -let show_prooftree () = () - (* Simulate the Intro(s) tactic *) let show_intro all = let open EConstr in - let pf = get_pftreestate() in + let pf = Proof_global.give_me_the_proof() in let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in if not (List.is_empty gls) then begin let gl = {Evd.it=List.hd gls ; sigma = sigma; } in @@ -108,14 +97,29 @@ let show_intro all = [Not_found] is raised if the given string isn't the qualid of a known inductive type. *) +(* + + HH notes in PR #679: + + The Show Match could also be made more robust, for instance in the + presence of let in the branch of a constructor. A + decompose_prod_assum would probably suffice for that, but then, it + is a Context.Rel.Declaration.t which needs to be matched and not + just a pair (name,type). + + Otherwise, this is OK. After all, the API on inductive types is not + so canonical in general, and in this simple case, working at the + low-level of mind_nf_lc seems reasonable (compared to working at the + higher-level of Inductiveops). + +*) + let make_cases_aux glob_ref = match glob_ref with - | Globnames.IndRef i -> - let {Declarations.mind_nparams = np} - , {Declarations.mind_consnames = carr ; Declarations.mind_nf_lc = tarr } - = Global.lookup_inductive i in - Util.Array.fold_right2 - (fun consname typ l -> + | Globnames.IndRef ind -> + let {Declarations.mind_nparams = np} , {Declarations.mind_nf_lc = tarr} = Global.lookup_inductive ind in + Util.Array.fold_right_i + (fun i typ l -> let al = List.rev (fst (decompose_prod typ)) in let al = Util.List.skipn np al in let rec rename avoid = function @@ -124,8 +128,9 @@ let make_cases_aux glob_ref = let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n avoid in Id.to_string n' :: rename (n'::avoid) l in let al' = rename [] al in - (Id.to_string consname :: al') :: l) - carr tarr [] + let consref = ConstructRef (ith_constructor_of_inductive ind (i + 1)) in + (Libnames.string_of_qualid (Nametab.shortest_qualid_of_global Id.Set.empty consref) :: al') :: l) + tarr [] | _ -> raise Not_found let make_cases s = @@ -492,7 +497,7 @@ let vernac_start_proof locality p kind l lettop = match id with | Some (lid,_) -> Dumpglob.dump_definition lid false "prf" | None -> ()) l; - if not(refining ()) then + if not(Proof_global.there_are_pending_proofs ()) then if lettop then user_err ~hdr:"Vernacentries.StartProof" (str "Let declarations can only be used in proof editing mode."); @@ -505,7 +510,7 @@ let vernac_end_proof ?proof = function let vernac_exact_proof c = (* spiwack: for simplicity I do not enforce that "Proof proof_term" is called only at the begining of a proof. *) - let status = by (Tactics.exact_proof c) in + let status = Pfedit.by (Tactics.exact_proof c) in save_proof (Vernacexpr.(Proved(Opaque None,None))); if not status then Feedback.feedback Feedback.AddedAxiom @@ -623,8 +628,7 @@ let vernac_constraint loc poly l = (* Modules *) let vernac_import export refl = - Library.import_module export (List.map qualid_of_reference refl); - Lib.add_frozen_state () + Library.import_module export (List.map qualid_of_reference refl) let vernac_declare_module export (loc, id) binders_ast mty_ast = (* We check the state of the system (in section, in module type) @@ -651,7 +655,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = user_err Pp.(str "Modules and Module Types are not allowed inside sections."); match mexpr_ast_l with | [] -> - check_no_pending_proofs (); + Proof_global.check_no_pending_proof (); let binders_ast,argsexport = List.fold_right (fun (export,idl,ty) (args,argsexport) -> @@ -697,7 +701,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l = match mty_ast_l with | [] -> - check_no_pending_proofs (); + Proof_global.check_no_pending_proof (); let binders_ast,argsexport = List.fold_right (fun (export,idl,ty) (args,argsexport) -> @@ -745,7 +749,7 @@ let vernac_include l = (* Sections *) let vernac_begin_section (_, id as lid) = - check_no_pending_proofs (); + Proof_global.check_no_pending_proof (); Dumpglob.dump_definition lid true "sec"; Lib.open_section id @@ -759,7 +763,7 @@ let vernac_name_sec_hyp (_,id) set = Proof_using.name_set id set (* Dispatcher of the "End" command *) let vernac_end_segment (_,id as lid) = - check_no_pending_proofs (); + Proof_global.check_no_pending_proof (); match Lib.find_opening_node id with | Lib.OpenedModule (false,export,_,_) -> vernac_end_module export lid | Lib.OpenedModule (true,_,_,_) -> vernac_end_modtype lid @@ -839,14 +843,14 @@ let focus_command_cond = Proof.no_cond command_focus there are no more goals to solve. It cannot be a tactic since all tactics fail if there are no further goals to prove. *) -let vernac_solve_existential = instantiate_nth_evar_com +let vernac_solve_existential = Pfedit.instantiate_nth_evar_com let vernac_set_end_tac tac = let env = Genintern.empty_glob_sign (Global.env ()) in let _, tac = Genintern.generic_intern env tac in - if not (refining ()) then + if not (Proof_global.there_are_pending_proofs ()) then user_err Pp.(str "Unknown command of the non proof-editing mode."); - set_end_tac tac + Proof_global.set_endline_tactic tac (* TO DO verifier s'il faut pas mettre exist s | TacId s ici*) let vernac_set_used_variables e = @@ -861,13 +865,13 @@ let vernac_set_used_variables e = user_err ~hdr:"vernac_set_used_variables" (str "Unknown variable: " ++ pr_id id)) l; - let _, to_clear = set_used_variables l in + let _, to_clear = Proof_global.set_used_variables l in let to_clear = List.map snd to_clear in Proof_global.with_current_proof begin fun _ p -> if List.is_empty to_clear then (p, ()) else let tac = Tactics.clear to_clear in - fst (solve SelectAll None tac p), () + fst (Pfedit.solve SelectAll None tac p), () end (*****************************) @@ -911,12 +915,12 @@ let vernac_chdir = function (* State management *) let vernac_write_state file = - Pfedit.delete_all_proofs (); + Proof_global.discard_all (); let file = CUnix.make_suffix file ".coq" in States.extern_state file let vernac_restore_state file = - Pfedit.delete_all_proofs (); + Proof_global.discard_all (); let file = Loadpath.locate_file (CUnix.make_suffix file ".coq") in States.intern_state file @@ -1282,7 +1286,7 @@ let _ = let _ = declare_bool_option - { optdepr = false; + { optdepr = true; (* remove in 8.8 *) optname = "automatic introduction of variables"; optkey = ["Automatic";"Introduction"]; optread = Flags.is_auto_intros; @@ -1378,17 +1382,6 @@ let _ = optread = (fun () -> !CClosure.share); optwrite = (fun b -> CClosure.share := b) } -(* No more undo limit in the new proof engine. - The command still exists for compatibility (e.g. with ProofGeneral) *) - -let _ = - declare_int_option - { optdepr = true; - optname = "the undo limit (OBSOLETE)"; - optkey = ["Undo"]; - optread = (fun _ -> None); - optwrite = (fun _ -> ()) } - let _ = declare_bool_option { optdepr = false; @@ -1510,7 +1503,7 @@ let vernac_print_option key = with Not_found -> error_undeclared_key key let get_current_context_of_args = function - | Some n -> get_goal_context n + | Some n -> Pfedit.get_goal_context n | None -> get_current_context () let query_command_selector ?loc = function @@ -1572,7 +1565,7 @@ let vernac_global_check c = let get_nth_goal n = - let pf = get_pftreestate() in + let pf = Proof_global.give_me_the_proof() in let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in let gl = {Evd.it=List.nth gls (n-1) ; sigma = sigma; } in gl @@ -1761,7 +1754,7 @@ let vernac_locate = let open Feedback in function | LocateFile f -> msg_notice (locate_file f) let vernac_register id r = - if Pfedit.refining () then + if Proof_global.there_are_pending_proofs () then user_err Pp.(str "Cannot register a primitive while in proof editing mode."); let kn = Constrintern.global_reference (snd id) in if not (isConstRef kn) then @@ -1828,24 +1821,16 @@ let vernac_show = let open Feedback in function | GoalUid id -> pr_goal_by_uid id in msg_notice info - | ShowGoalImplicitly None -> - Constrextern.with_implicits msg_notice (pr_open_subgoals ()) - | ShowGoalImplicitly (Some n) -> - Constrextern.with_implicits msg_notice (pr_nth_open_subgoal n) | ShowProof -> show_proof () - | ShowNode -> show_node () | ShowExistentials -> show_top_evars () | ShowUniverses -> show_universes () - | ShowTree -> show_prooftree () | ShowProofNames -> - msg_notice (pr_sequence pr_id (Pfedit.get_all_proof_names())) + msg_notice (pr_sequence pr_id (Proof_global.get_all_proof_names())) | ShowIntros all -> show_intro all | ShowMatch id -> show_match id - | ShowThesis -> show_thesis () - let vernac_check_guard () = - let pts = get_pftreestate () in + let pts = Proof_global.give_me_the_proof () in let pfterm = List.hd (Proof.partial_proof pts) in let message = try |