aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Maxime Dénès <mail@maximedenes.fr>2017-05-21 14:50:25 +0200
committerGravatar Maxime Dénès <mail@maximedenes.fr>2017-06-06 17:41:37 +0200
commit014749917e5de9fe1885a1b1edc52b01cefa6f3f (patch)
tree5b6952c2d67b1c18ce5f6e22cc75cefc005ad460
parent0a577d3c979af094ca00be3e7e323109c7e1f6ab (diff)
Merge the ssr plugin.
-rw-r--r--.merlin5
-rw-r--r--Makefile.common5
-rw-r--r--plugins/ssr/ssrast.mli149
-rw-r--r--plugins/ssr/ssrbool.v1871
-rw-r--r--plugins/ssr/ssrbwd.ml126
-rw-r--r--plugins/ssr/ssrbwd.mli20
-rw-r--r--plugins/ssr/ssrcommon.ml1297
-rw-r--r--plugins/ssr/ssrcommon.mli410
-rw-r--r--plugins/ssr/ssreflect.v451
-rw-r--r--plugins/ssr/ssreflect_plugin.mlpack13
-rw-r--r--plugins/ssr/ssrelim.ml441
-rw-r--r--plugins/ssr/ssrelim.mli53
-rw-r--r--plugins/ssr/ssrequality.ml663
-rw-r--r--plugins/ssr/ssrequality.mli62
-rw-r--r--plugins/ssr/ssrfun.v791
-rw-r--r--plugins/ssr/ssrfwd.ml409
-rw-r--r--plugins/ssr/ssrfwd.mli65
-rw-r--r--plugins/ssr/ssripats.ml400
-rw-r--r--plugins/ssr/ssripats.mli82
-rw-r--r--plugins/ssr/ssrparser.ml42349
-rw-r--r--plugins/ssr/ssrparser.mli20
-rw-r--r--plugins/ssr/ssrprinters.ml85
-rw-r--r--plugins/ssr/ssrprinters.mli45
-rw-r--r--plugins/ssr/ssrtacticals.ml160
-rw-r--r--plugins/ssr/ssrtacticals.mli44
-rw-r--r--plugins/ssr/ssrvernac.ml4600
-rw-r--r--plugins/ssr/ssrvernac.mli9
-rw-r--r--plugins/ssr/ssrview.ml125
-rw-r--r--plugins/ssr/ssrview.mli36
-rw-r--r--plugins/ssr/vo.itarget3
30 files changed, 10785 insertions, 4 deletions
diff --git a/.merlin b/.merlin
index b78f24551..c8d7d322f 100644
--- a/.merlin
+++ b/.merlin
@@ -1,7 +1,5 @@
FLG -rectypes -thread -safe-string -w +a-4-9-27-41-42-44-45-48-50
-S ltac
-B ltac
S config
B config
S ide
@@ -46,4 +44,7 @@ B tools/coqdoc
S dev
B dev
+S plugins/**
+B plugins/**
+
PKG threads.posix camlp5
diff --git a/Makefile.common b/Makefile.common
index b936eb4c7..4545fad05 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -62,7 +62,7 @@ PLUGINDIRS:=\
setoid_ring extraction fourier \
cc funind firstorder derive \
rtauto nsatz syntax btauto \
- ssrmatching ltac
+ ssrmatching ltac ssr
SRCDIRS:=\
$(CORESRCDIRS) \
@@ -120,13 +120,14 @@ OTHERSYNTAXCMO:=$(addprefix plugins/syntax/, \
DERIVECMO:=plugins/derive/derive_plugin.cmo
LTACCMO:=plugins/ltac/ltac_plugin.cmo
SSRMATCHINGCMO:=plugins/ssrmatching/ssrmatching_plugin.cmo
+SSRCMO:=plugins/ssr/ssreflect_plugin.cmo
PLUGINSCMO:=$(LTACCMO) $(OMEGACMO) $(ROMEGACMO) $(MICROMEGACMO) \
$(QUOTECMO) $(RINGCMO) \
$(FOURIERCMO) $(EXTRACTIONCMO) \
$(CCCMO) $(FOCMO) $(RTAUTOCMO) $(BTAUTOCMO) \
$(FUNINDCMO) $(NSATZCMO) $(NATSYNTAXCMO) $(OTHERSYNTAXCMO) \
- $(DERIVECMO) $(SSRMATCHINGCMO)
+ $(DERIVECMO) $(SSRMATCHINGCMO) $(SSRCMO)
ifeq ($(HASNATDYNLINK)-$(BEST),false-opt)
STATICPLUGINS:=$(PLUGINSCMO)
diff --git a/plugins/ssr/ssrast.mli b/plugins/ssr/ssrast.mli
new file mode 100644
index 000000000..69202ae2d
--- /dev/null
+++ b/plugins/ssr/ssrast.mli
@@ -0,0 +1,149 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Ltac_plugin
+
+(* Names of variables to be cleared (automatic check: not a section var) *)
+type ssrhyp = SsrHyp of Id.t Loc.located
+(* Variant of the above *)
+type ssrhyp_or_id = Hyp of ssrhyp | Id of ssrhyp
+
+(* Variant of the above *)
+type ssrhyps = ssrhyp list
+
+(* Direction to be used for rewriting as in -> or rewrite flag *)
+type ssrdir = Ssrmatching_plugin.Ssrmatching.ssrdir = L2R | R2L
+
+(* simpl: "/=", cut: "//", simplcut: "//=" nop: commodity placeholder *)
+type ssrsimpl = Simpl of int | Cut of int | SimplCut of int * int | Nop
+
+(* modality for rewrite and do: ! ? *)
+type ssrmmod = May | Must | Once
+
+(* modality with a bound for rewrite and do: !n ?n *)
+type ssrmult = int * ssrmmod
+
+(** Occurrence switch {1 2}, all is Some(false,[]) *)
+type ssrocc = (bool * int list) option
+
+(* index MAYBE REMOVE ONLY INTERNAL stuff between {} *)
+type ssrindex = int Misctypes.or_var
+
+(* clear switch {H G} *)
+type ssrclear = ssrhyps
+
+(* Discharge occ switch (combined occurrence / clear switch) *)
+type ssrdocc = ssrclear option * ssrocc
+
+(* FIXME, make algebraic *)
+type ssrtermkind = char
+
+type ssrterm = ssrtermkind * Tacexpr.glob_constr_and_expr
+
+type ssrview = ssrterm list
+
+(* TODO
+type id_mod = Hat | HatTilde | Sharp
+ *)
+
+(* Only [One] forces an introduction, possibly reducing the goal. *)
+type anon_iter =
+ | One
+ | Drop
+ | All
+
+(* TODO
+ | Dependent (* fast mode *)
+ | UntilMark
+ | Temporary (* "+" *)
+ *)
+
+type ssripat =
+ | IPatNoop
+ | IPatId of (*TODO id_mod option * *) Id.t
+ | IPatAnon of anon_iter (* inaccessible name *)
+(* TODO | IPatClearMark *)
+(* TODO | IPatDispatch of ssripatss (* /[..|..] *) *)
+ | IPatCase of (* ipats_mod option * *) ssripatss (* this is not equivalent to /case /[..|..] if there are already multiple goals *)
+ | IPatInj of ssripatss
+ | IPatRewrite of (*occurrence option * rewrite_pattern **) ssrocc * ssrdir
+ | IPatView of ssrterm list (* /view *)
+ | IPatClear of ssrclear (* {H1 H2} *)
+ | IPatSimpl of ssrsimpl
+ | IPatNewHidden of identifier list
+(* | IPatVarsForAbstract of Id.t list *)
+
+and ssripats = ssripat list
+and ssripatss = ssripats list
+type ssrhpats = ((ssrclear * ssripats) * ssripats) * ssripats
+type ssrhpats_wtransp = bool * ssrhpats
+
+(* tac => inpats *)
+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
+ | 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 *)
+(* base term from a term, for the two first levels of representation of *)
+(* of constr terms. *)
+type ssrbindfmt =
+ | BFvar
+ | BFdecl of int (* #xs *)
+ | BFcast (* final cast *)
+ | BFdef (* has cast? *)
+ | BFrec of bool * bool (* has struct? * has cast? *)
+type 'term ssrbindval = 'term ssrbind list * 'term
+
+(** Forward chaining argument *)
+(* There are three kinds of forward definitions: *)
+(* - Hint: type only, cast to Type, may have proof hint. *)
+(* - Have: type option + value, no space before type *)
+(* - Pose: binders + value, space before binders. *)
+type ssrfwdkind = FwdHint of string * bool | FwdHave | FwdPose
+type ssrfwdfmt = ssrfwdkind * ssrbindfmt list
+
+(* in *)
+type ssrclseq = InGoal | InHyps
+ | InHypsGoal | InHypsSeqGoal | InSeqGoal | InHypsSeq | InAll | InAllHyps
+
+type 'tac ssrhint = bool * 'tac option list
+
+type 'tac fwdbinders =
+ bool * (ssrhpats * ((ssrfwdfmt * ssrterm) * 'tac ssrhint))
+
+type clause =
+ (ssrclear * ((ssrhyp_or_id * string) *
+ Ssrmatching_plugin.Ssrmatching.cpattern option) option)
+type clauses = clause list * ssrclseq
+
+type wgen =
+ (ssrclear *
+ ((ssrhyp_or_id * string) *
+ Ssrmatching_plugin.Ssrmatching.cpattern option)
+ option)
+
+type 'a ssrdoarg = ((ssrindex * ssrmmod) * 'a ssrhint) * clauses
+type 'a ssrseqarg = ssrindex * ('a ssrhint * 'a option)
+
+(* OOP : these are general shortcuts *)
+type gist = Tacintern.glob_sign
+type ist = Tacinterp.interp_sign
+type goal = Proof_type.goal
+type 'a sigma = 'a Evd.sigma
+type v82tac = Proof_type.tactic
diff --git a/plugins/ssr/ssrbool.v b/plugins/ssr/ssrbool.v
new file mode 100644
index 000000000..63bf0116c
--- /dev/null
+++ b/plugins/ssr/ssrbool.v
@@ -0,0 +1,1871 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+Require Bool.
+Require Import ssreflect ssrfun.
+
+(******************************************************************************)
+(* A theory of boolean predicates and operators. A large part of this file is *)
+(* concerned with boolean reflection. *)
+(* Definitions and notations: *)
+(* is_true b == the coercion of b : bool to Prop (:= b = true). *)
+(* This is just input and displayed as `b''. *)
+(* reflect P b == the reflection inductive predicate, asserting *)
+(* that the logical proposition P : prop with the *)
+(* formula b : bool. Lemmas asserting reflect P b *)
+(* are often referred to as "views". *)
+(* iffP, appP, sameP, rwP :: lemmas for direct manipulation of reflection *)
+(* views: iffP is used to prove reflection from *)
+(* logical equivalence, appP to compose views, and *)
+(* sameP and rwP to perform boolean and setoid *)
+(* rewriting. *)
+(* elimT :: coercion reflect >-> Funclass, which allows the *)
+(* direct application of `reflect' views to *)
+(* boolean assertions. *)
+(* decidable P <-> P is effectively decidable (:= {P} + {~ P}. *)
+(* contra, contraL, ... :: contraposition lemmas. *)
+(* altP my_viewP :: natural alternative for reflection; given *)
+(* lemma myviewP: reflect my_Prop my_formula, *)
+(* have [myP | not_myP] := altP my_viewP. *)
+(* generates two subgoals, in which my_formula has *)
+(* been replaced by true and false, resp., with *)
+(* new assumptions myP : my_Prop and *)
+(* not_myP: ~~ my_formula. *)
+(* Caveat: my_formula must be an APPLICATION, not *)
+(* a variable, constant, let-in, etc. (due to the *)
+(* poor behaviour of dependent index matching). *)
+(* boolP my_formula :: boolean disjunction, equivalent to *)
+(* altP (idP my_formula) but circumventing the *)
+(* dependent index capture issue; destructing *)
+(* boolP my_formula generates two subgoals with *)
+(* assumtions my_formula and ~~ myformula. As *)
+(* with altP, my_formula must be an application. *)
+(* \unless C, P <-> we can assume property P when a something that *)
+(* holds under condition C (such as C itself). *)
+(* := forall G : Prop, (C -> G) -> (P -> G) -> G. *)
+(* This is just C \/ P or rather its impredicative *)
+(* encoding, whose usage better fits the above *)
+(* description: given a lemma UCP whose conclusion *)
+(* is \unless C, P we can assume P by writing: *)
+(* wlog hP: / P by apply/UCP; (prove C -> goal). *)
+(* or even apply: UCP id _ => hP if the goal is C. *)
+(* classically P <-> we can assume P when proving is_true b. *)
+(* := forall b : bool, (P -> b) -> b. *)
+(* This is equivalent to ~ (~ P) when P : Prop. *)
+(* implies P Q == wrapper coinductive type that coerces to P -> Q *)
+(* and can be used as a P -> Q view unambigously. *)
+(* Useful to avoid spurious insertion of <-> views *)
+(* when Q is a conjunction of foralls, as in Lemma *)
+(* all_and2 below; conversely, avoids confusion in *)
+(* apply views for impredicative properties, such *)
+(* as \unless C, P. Also supports contrapositives. *)
+(* a && b == the boolean conjunction of a and b. *)
+(* a || b == the boolean disjunction of a and b. *)
+(* a ==> b == the boolean implication of b by a. *)
+(* ~~ a == the boolean negation of a. *)
+(* a (+) b == the boolean exclusive or (or sum) of a and b. *)
+(* [ /\ P1 , P2 & P3 ] == multiway logical conjunction, up to 5 terms. *)
+(* [ \/ P1 , P2 | P3 ] == multiway logical disjunction, up to 4 terms. *)
+(* [&& a, b, c & d] == iterated, right associative boolean conjunction *)
+(* with arbitrary arity. *)
+(* [|| a, b, c | d] == iterated, right associative boolean disjunction *)
+(* with arbitrary arity. *)
+(* [==> a, b, c => d] == iterated, right associative boolean implication *)
+(* with arbitrary arity. *)
+(* and3P, ... == specific reflection lemmas for iterated *)
+(* connectives. *)
+(* andTb, orbAC, ... == systematic names for boolean connective *)
+(* properties (see suffix conventions below). *)
+(* prop_congr == a tactic to move a boolean equality from *)
+(* its coerced form in Prop to the equality *)
+(* in bool. *)
+(* bool_congr == resolution tactic for blindly weeding out *)
+(* like terms from boolean equalities (can fail). *)
+(* This file provides a theory of boolean predicates and relations: *)
+(* pred T == the type of bool predicates (:= T -> bool). *)
+(* simpl_pred T == the type of simplifying bool predicates, using *)
+(* the simpl_fun from ssrfun.v. *)
+(* rel T == the type of bool relations. *)
+(* := T -> pred T or T -> T -> bool. *)
+(* simpl_rel T == type of simplifying relations. *)
+(* predType == the generic predicate interface, supported for *)
+(* for lists and sets. *)
+(* pred_class == a coercion class for the predType projection to *)
+(* pred; declaring a coercion to pred_class is an *)
+(* alternative way of equipping a type with a *)
+(* predType structure, which interoperates better *)
+(* with coercion subtyping. This is used, e.g., *)
+(* for finite sets, so that finite groups inherit *)
+(* the membership operation by coercing to sets. *)
+(* If P is a predicate the proposition "x satisfies P" can be written *)
+(* applicatively as (P x), or using an explicit connective as (x \in P); in *)
+(* the latter case we say that P is a "collective" predicate. We use A, B *)
+(* rather than P, Q for collective predicates: *)
+(* x \in A == x satisfies the (collective) predicate A. *)
+(* x \notin A == x doesn't satisfy the (collective) predicate A. *)
+(* The pred T type can be used as a generic predicate type for either kind, *)
+(* but the two kinds of predicates should not be confused. When a "generic" *)
+(* pred T value of one type needs to be passed as the other the following *)
+(* conversions should be used explicitly: *)
+(* SimplPred P == a (simplifying) applicative equivalent of P. *)
+(* mem A == an applicative equivalent of A: *)
+(* mem A x simplifies to x \in A. *)
+(* Alternatively one can use the syntax for explicit simplifying predicates *)
+(* and relations (in the following x is bound in E): *)
+(* [pred x | E] == simplifying (see ssrfun) predicate x => E. *)
+(* [pred x : T | E] == predicate x => E, with a cast on the argument. *)
+(* [pred : T | P] == constant predicate P on type T. *)
+(* [pred x | E1 & E2] == [pred x | E1 && E2]; an x : T cast is allowed. *)
+(* [pred x in A] == [pred x | x in A]. *)
+(* [pred x in A | E] == [pred x | x in A & E]. *)
+(* [pred x in A | E1 & E2] == [pred x in A | E1 && E2]. *)
+(* [predU A & B] == union of two collective predicates A and B. *)
+(* [predI A & B] == intersection of collective predicates A and B. *)
+(* [predD A & B] == difference of collective predicates A and B. *)
+(* [predC A] == complement of the collective predicate A. *)
+(* [preim f of A] == preimage under f of the collective predicate A. *)
+(* predU P Q, ... == union, etc of applicative predicates. *)
+(* pred0 == the empty predicate. *)
+(* predT == the total (always true) predicate. *)
+(* if T : predArgType, then T coerces to predT. *)
+(* {: T} == T cast to predArgType (e.g., {: bool * nat}) *)
+(* In the following, x and y are bound in E: *)
+(* [rel x y | E] == simplifying relation x, y => E. *)
+(* [rel x y : T | E] == simplifying relation with arguments cast. *)
+(* [rel x y in A & B | E] == [rel x y | [&& x \in A, y \in B & E]]. *)
+(* [rel x y in A & B] == [rel x y | (x \in A) && (y \in B)]. *)
+(* [rel x y in A | E] == [rel x y in A & A | E]. *)
+(* [rel x y in A] == [rel x y in A & A]. *)
+(* relU R S == union of relations R and S. *)
+(* Explicit values of type pred T (i.e., lamdba terms) should always be used *)
+(* applicatively, while values of collection types implementing the predType *)
+(* interface, such as sequences or sets should always be used as collective *)
+(* predicates. Defined constants and functions of type pred T or simpl_pred T *)
+(* as well as the explicit simpl_pred T values described below, can generally *)
+(* be used either way. Note however that x \in A will not auto-simplify when *)
+(* A is an explicit simpl_pred T value; the generic simplification rule inE *)
+(* must be used (when A : pred T, the unfold_in rule can be used). Constants *)
+(* of type pred T with an explicit simpl_pred value do not auto-simplify when *)
+(* used applicatively, but can still be expanded with inE. This behavior can *)
+(* be controlled as follows: *)
+(* Let A : collective_pred T := [pred x | ... ]. *)
+(* The collective_pred T type is just an alias for pred T, but this cast *)
+(* stops rewrite inE from expanding the definition of A, thus treating A *)
+(* into an abstract collection (unfold_in or in_collective can be used to *)
+(* expand manually). *)
+(* Let A : applicative_pred T := [pred x | ...]. *)
+(* This cast causes inE to turn x \in A into the applicative A x form; *)
+(* A will then have to unfolded explicitly with the /A rule. This will *)
+(* also apply to any definition that reduces to A (e.g., Let B := A). *)
+(* Canonical A_app_pred := ApplicativePred A. *)
+(* This declaration, given after definition of A, similarly causes inE to *)
+(* turn x \in A into A x, but in addition allows the app_predE rule to *)
+(* turn A x back into x \in A; it can be used for any definition of type *)
+(* pred T, which makes it especially useful for ambivalent predicates *)
+(* as the relational transitive closure connect, that are used in both *)
+(* applicative and collective styles. *)
+(* Purely for aesthetics, we provide a subtype of collective predicates: *)
+(* qualifier q T == a pred T pretty-printing wrapper. An A : qualifier q T *)
+(* coerces to pred_class and thus behaves as a collective *)
+(* predicate, but x \in A and x \notin A are displayed as: *)
+(* x \is A and x \isn't A when q = 0, *)
+(* x \is a A and x \isn't a A when q = 1, *)
+(* x \is an A and x \isn't an A when q = 2, respectively. *)
+(* [qualify x | P] := Qualifier 0 (fun x => P), constructor for the above. *)
+(* [qualify x : T | P], [qualify a x | P], [qualify an X | P], etc. *)
+(* variants of the above with type constraints and different *)
+(* values of q. *)
+(* We provide an internal interface to support attaching properties (such as *)
+(* being multiplicative) to predicates: *)
+(* pred_key p == phantom type that will serve as a support for properties *)
+(* to be attached to p : pred_class; instances should be *)
+(* created with Fact/Qed so as to be opaque. *)
+(* KeyedPred k_p == an instance of the interface structure that attaches *)
+(* (k_p : pred_key P) to P; the structure projection is a *)
+(* coercion to pred_class. *)
+(* KeyedQualifier k_q == an instance of the interface structure that attaches *)
+(* (k_q : pred_key q) to (q : qualifier n T). *)
+(* DefaultPredKey p == a default value for pred_key p; the vernacular command *)
+(* Import DefaultKeying attaches this key to all predicates *)
+(* that are not explicitly keyed. *)
+(* Keys can be used to attach properties to predicates, qualifiers and *)
+(* generic nouns in a way that allows them to be used transparently. The key *)
+(* projection of a predicate property structure such as unsignedPred should *)
+(* be a pred_key, not a pred, and corresponding lemmas will have the form *)
+(* Lemma rpredN R S (oppS : @opprPred R S) (kS : keyed_pred oppS) : *)
+(* {mono -%R: x / x \in kS}. *)
+(* Because x \in kS will be displayed as x \in S (or x \is S, etc), the *)
+(* canonical instance of opprPred will not normally be exposed (it will also *)
+(* be erased by /= simplification). In addition each predicate structure *)
+(* should have a DefaultPredKey Canonical instance that simply issues the *)
+(* property as a proof obligation (which can be caught by the Prop-irrelevant *)
+(* feature of the ssreflect plugin). *)
+(* Some properties of predicates and relations: *)
+(* A =i B <-> A and B are extensionally equivalent. *)
+(* {subset A <= B} <-> A is a (collective) subpredicate of B. *)
+(* subpred P Q <-> P is an (applicative) subpredicate or Q. *)
+(* subrel R S <-> R is a subrelation of S. *)
+(* In the following R is in rel T: *)
+(* reflexive R <-> R is reflexive. *)
+(* irreflexive R <-> R is irreflexive. *)
+(* symmetric R <-> R (in rel T) is symmetric (equation). *)
+(* pre_symmetric R <-> R is symmetric (implication). *)
+(* antisymmetric R <-> R is antisymmetric. *)
+(* total R <-> R is total. *)
+(* transitive R <-> R is transitive. *)
+(* left_transitive R <-> R is a congruence on its left hand side. *)
+(* right_transitive R <-> R is a congruence on its right hand side. *)
+(* equivalence_rel R <-> R is an equivalence relation. *)
+(* Localization of (Prop) predicates; if P1 is convertible to forall x, Qx, *)
+(* P2 to forall x y, Qxy and P3 to forall x y z, Qxyz : *)
+(* {for y, P1} <-> Qx{y / x}. *)
+(* {in A, P1} <-> forall x, x \in A -> Qx. *)
+(* {in A1 & A2, P2} <-> forall x y, x \in A1 -> y \in A2 -> Qxy. *)
+(* {in A &, P2} <-> forall x y, x \in A -> y \in A -> Qxy. *)
+(* {in A1 & A2 & A3, Q3} <-> forall x y z, *)
+(* x \in A1 -> y \in A2 -> z \in A3 -> Qxyz. *)
+(* {in A1 & A2 &, Q3} == {in A1 & A2 & A2, Q3}. *)
+(* {in A1 && A3, Q3} == {in A1 & A1 & A3, Q3}. *)
+(* {in A &&, Q3} == {in A & A & A, Q3}. *)
+(* {in A, bijective f} == f has a right inverse in A. *)
+(* {on C, P1} == forall x, (f x) \in C -> Qx *)
+(* when P1 is also convertible to Pf f. *)
+(* {on C &, P2} == forall x y, f x \in C -> f y \in C -> Qxy *)
+(* when P2 is also convertible to Pf f. *)
+(* {on C, P1' & g} == forall x, (f x) \in cd -> Qx *)
+(* when P1' is convertible to Pf f *)
+(* and P1' g is convertible to forall x, Qx. *)
+(* {on C, bijective f} == f has a right inverse on C. *)
+(* This file extends the lemma name suffix conventions of ssrfun as follows: *)
+(* A -- associativity, as in andbA : associative andb. *)
+(* AC -- right commutativity. *)
+(* ACA -- self-interchange (inner commutativity), e.g., *)
+(* orbACA : (a || b) || (c || d) = (a || c) || (b || d). *)
+(* b -- a boolean argument, as in andbb : idempotent andb. *)
+(* C -- commutativity, as in andbC : commutative andb, *)
+(* or predicate complement, as in predC. *)
+(* CA -- left commutativity. *)
+(* D -- predicate difference, as in predD. *)
+(* E -- elimination, as in negbFE : ~~ b = false -> b. *)
+(* F or f -- boolean false, as in andbF : b && false = false. *)
+(* I -- left/right injectivity, as in addbI : right_injective addb, *)
+(* or predicate intersection, as in predI. *)
+(* l -- a left-hand operation, as andb_orl : left_distributive andb orb. *)
+(* N or n -- boolean negation, as in andbN : a && (~~ a) = false. *)
+(* P -- a characteristic property, often a reflection lemma, as in *)
+(* andP : reflect (a /\ b) (a && b). *)
+(* r -- a right-hand operation, as orb_andr : rightt_distributive orb andb. *)
+(* T or t -- boolean truth, as in andbT: right_id true andb. *)
+(* U -- predicate union, as in predU. *)
+(* W -- weakening, as in in1W : {in D, forall x, P} -> forall x, P. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+Set Warnings "-projection-no-head-constant".
+
+Notation reflect := Bool.reflect.
+Notation ReflectT := Bool.ReflectT.
+Notation ReflectF := Bool.ReflectF.
+
+Reserved Notation "~~ b" (at level 35, right associativity).
+Reserved Notation "b ==> c" (at level 55, right associativity).
+Reserved Notation "b1 (+) b2" (at level 50, left associativity).
+Reserved Notation "x \in A"
+ (at level 70, format "'[hv' x '/ ' \in A ']'", no associativity).
+Reserved Notation "x \notin A"
+ (at level 70, format "'[hv' x '/ ' \notin A ']'", no associativity).
+Reserved Notation "p1 =i p2"
+ (at level 70, format "'[hv' p1 '/ ' =i p2 ']'", no associativity).
+
+(* We introduce a number of n-ary "list-style" notations that share a common *)
+(* format, namely *)
+(* [op arg1, arg2, ... last_separator last_arg] *)
+(* This usually denotes a right-associative applications of op, e.g., *)
+(* [&& a, b, c & d] denotes a && (b && (c && d)) *)
+(* The last_separator must be a non-operator token. Here we use &, | or =>; *)
+(* our default is &, but we try to match the intended meaning of op. The *)
+(* separator is a workaround for limitations of the parsing engine; the same *)
+(* limitations mean the separator cannot be omitted even when last_arg can. *)
+(* The Notation declarations are complicated by the separate treatment for *)
+(* some fixed arities (binary for bool operators, and all arities for Prop *)
+(* operators). *)
+(* We also use the square brackets in comprehension-style notations *)
+(* [type var separator expr] *)
+(* where "type" is the type of the comprehension (e.g., pred) and "separator" *)
+(* is | or => . It is important that in other notations a leading square *)
+(* bracket [ is always followed by an operator symbol or a fixed identifier. *)
+
+Reserved Notation "[ /\ P1 & P2 ]" (at level 0, only parsing).
+Reserved Notation "[ /\ P1 , P2 & P3 ]" (at level 0, format
+ "'[hv' [ /\ '[' P1 , '/' P2 ']' '/ ' & P3 ] ']'").
+Reserved Notation "[ /\ P1 , P2 , P3 & P4 ]" (at level 0, format
+ "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 ']' '/ ' & P4 ] ']'").
+Reserved Notation "[ /\ P1 , P2 , P3 , P4 & P5 ]" (at level 0, format
+ "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' & P5 ] ']'").
+
+Reserved Notation "[ \/ P1 | P2 ]" (at level 0, only parsing).
+Reserved Notation "[ \/ P1 , P2 | P3 ]" (at level 0, format
+ "'[hv' [ \/ '[' P1 , '/' P2 ']' '/ ' | P3 ] ']'").
+Reserved Notation "[ \/ P1 , P2 , P3 | P4 ]" (at level 0, format
+ "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 ']' '/ ' | P4 ] ']'").
+
+Reserved Notation "[ && b1 & c ]" (at level 0, only parsing).
+Reserved Notation "[ && b1 , b2 , .. , bn & c ]" (at level 0, format
+ "'[hv' [ && '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/ ' & c ] ']'").
+
+Reserved Notation "[ || b1 | c ]" (at level 0, only parsing).
+Reserved Notation "[ || b1 , b2 , .. , bn | c ]" (at level 0, format
+ "'[hv' [ || '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/ ' | c ] ']'").
+
+Reserved Notation "[ ==> b1 => c ]" (at level 0, only parsing).
+Reserved Notation "[ ==> b1 , b2 , .. , bn => c ]" (at level 0, format
+ "'[hv' [ ==> '[' b1 , '/' b2 , '/' .. , '/' bn ']' '/' => c ] ']'").
+
+Reserved Notation "[ 'pred' : T => E ]" (at level 0, format
+ "'[hv' [ 'pred' : T => '/ ' E ] ']'").
+Reserved Notation "[ 'pred' x => E ]" (at level 0, x at level 8, format
+ "'[hv' [ 'pred' x => '/ ' E ] ']'").
+Reserved Notation "[ 'pred' x : T => E ]" (at level 0, x at level 8, format
+ "'[hv' [ 'pred' x : T => '/ ' E ] ']'").
+
+Reserved Notation "[ 'rel' x y => E ]" (at level 0, x, y at level 8, format
+ "'[hv' [ 'rel' x y => '/ ' E ] ']'").
+Reserved Notation "[ 'rel' x y : T => E ]" (at level 0, x, y at level 8, format
+ "'[hv' [ 'rel' x y : T => '/ ' E ] ']'").
+
+(* Shorter delimiter *)
+Delimit Scope bool_scope with B.
+Open Scope bool_scope.
+
+(* An alternative to xorb that behaves somewhat better wrt simplification. *)
+Definition addb b := if b then negb else id.
+
+(* Notation for && and || is declared in Init.Datatypes. *)
+Notation "~~ b" := (negb b) : bool_scope.
+Notation "b ==> c" := (implb b c) : bool_scope.
+Notation "b1 (+) b2" := (addb b1 b2) : bool_scope.
+
+(* Constant is_true b := b = true is defined in Init.Datatypes. *)
+Coercion is_true : bool >-> Sortclass. (* Prop *)
+
+Lemma prop_congr : forall b b' : bool, b = b' -> b = b' :> Prop.
+Proof. by move=> b b' ->. Qed.
+
+Ltac prop_congr := apply: prop_congr.
+
+(* Lemmas for trivial. *)
+Lemma is_true_true : true. Proof. by []. Qed.
+Lemma not_false_is_true : ~ false. Proof. by []. Qed.
+Lemma is_true_locked_true : locked true. Proof. by unlock. Qed.
+Hint Resolve is_true_true not_false_is_true is_true_locked_true.
+
+(* Shorter names. *)
+Definition isT := is_true_true.
+Definition notF := not_false_is_true.
+
+(* Negation lemmas. *)
+
+(* We generally take NEGATION as the standard form of a false condition: *)
+(* negative boolean hypotheses should be of the form ~~ b, rather than ~ b or *)
+(* b = false, as much as possible. *)
+
+Lemma negbT b : b = false -> ~~ b. Proof. by case: b. Qed.
+Lemma negbTE b : ~~ b -> b = false. Proof. by case: b. Qed.
+Lemma negbF b : (b : bool) -> ~~ b = false. Proof. by case: b. Qed.
+Lemma negbFE b : ~~ b = false -> b. Proof. by case: b. Qed.
+Lemma negbK : involutive negb. Proof. by case. Qed.
+Lemma negbNE b : ~~ ~~ b -> b. Proof. by case: b. Qed.
+
+Lemma negb_inj : injective negb. Proof. exact: can_inj negbK. Qed.
+Lemma negbLR b c : b = ~~ c -> ~~ b = c. Proof. exact: canLR negbK. Qed.
+Lemma negbRL b c : ~~ b = c -> b = ~~ c. Proof. exact: canRL negbK. Qed.
+
+Lemma contra (c b : bool) : (c -> b) -> ~~ b -> ~~ c.
+Proof. by case: b => //; case: c. Qed.
+Definition contraNN := contra.
+
+Lemma contraL (c b : bool) : (c -> ~~ b) -> b -> ~~ c.
+Proof. by case: b => //; case: c. Qed.
+Definition contraTN := contraL.
+
+Lemma contraR (c b : bool) : (~~ c -> b) -> ~~ b -> c.
+Proof. by case: b => //; case: c. Qed.
+Definition contraNT := contraR.
+
+Lemma contraLR (c b : bool) : (~~ c -> ~~ b) -> b -> c.
+Proof. by case: b => //; case: c. Qed.
+Definition contraTT := contraLR.
+
+Lemma contraT b : (~~ b -> false) -> b. Proof. by case: b => // ->. Qed.
+
+Lemma wlog_neg b : (~~ b -> b) -> b. Proof. by case: b => // ->. Qed.
+
+Lemma contraFT (c b : bool) : (~~ c -> b) -> b = false -> c.
+Proof. by move/contraR=> notb_c /negbT. Qed.
+
+Lemma contraFN (c b : bool) : (c -> b) -> b = false -> ~~ c.
+Proof. by move/contra=> notb_notc /negbT. Qed.
+
+Lemma contraTF (c b : bool) : (c -> ~~ b) -> b -> c = false.
+Proof. by move/contraL=> b_notc /b_notc/negbTE. Qed.
+
+Lemma contraNF (c b : bool) : (c -> b) -> ~~ b -> c = false.
+Proof. by move/contra=> notb_notc /notb_notc/negbTE. Qed.
+
+Lemma contraFF (c b : bool) : (c -> b) -> b = false -> c = false.
+Proof. by move/contraFN=> bF_notc /bF_notc/negbTE. Qed.
+
+(* Coercion of sum-style datatypes into bool, which makes it possible *)
+(* to use ssr's boolean if rather than Coq's "generic" if. *)
+
+Coercion isSome T (u : option T) := if u is Some _ then true else false.
+
+Coercion is_inl A B (u : A + B) := if u is inl _ then true else false.
+
+Coercion is_left A B (u : {A} + {B}) := if u is left _ then true else false.
+
+Coercion is_inleft A B (u : A + {B}) := if u is inleft _ then true else false.
+
+Prenex Implicits isSome is_inl is_left is_inleft.
+
+Definition decidable P := {P} + {~ P}.
+
+(* Lemmas for ifs with large conditions, which allow reasoning about the *)
+(* condition without repeating it inside the proof (the latter IS *)
+(* preferable when the condition is short). *)
+(* Usage : *)
+(* if the goal contains (if cond then ...) = ... *)
+(* case: ifP => Hcond. *)
+(* generates two subgoal, with the assumption Hcond : cond = true/false *)
+(* Rewrite if_same eliminates redundant ifs *)
+(* Rewrite (fun_if f) moves a function f inside an if *)
+(* Rewrite if_arg moves an argument inside a function-valued if *)
+
+Section BoolIf.
+
+Variables (A B : Type) (x : A) (f : A -> B) (b : bool) (vT vF : A).
+
+CoInductive if_spec (not_b : Prop) : bool -> A -> Set :=
+ | IfSpecTrue of b : if_spec not_b true vT
+ | IfSpecFalse of not_b : if_spec not_b false vF.
+
+Lemma ifP : if_spec (b = false) b (if b then vT else vF).
+Proof. by case def_b: b; constructor. Qed.
+
+Lemma ifPn : if_spec (~~ b) b (if b then vT else vF).
+Proof. by case def_b: b; constructor; rewrite ?def_b. Qed.
+
+Lemma ifT : b -> (if b then vT else vF) = vT. Proof. by move->. Qed.
+Lemma ifF : b = false -> (if b then vT else vF) = vF. Proof. by move->. Qed.
+Lemma ifN : ~~ b -> (if b then vT else vF) = vF. Proof. by move/negbTE->. Qed.
+
+Lemma if_same : (if b then vT else vT) = vT.
+Proof. by case b. Qed.
+
+Lemma if_neg : (if ~~ b then vT else vF) = if b then vF else vT.
+Proof. by case b. Qed.
+
+Lemma fun_if : f (if b then vT else vF) = if b then f vT else f vF.
+Proof. by case b. Qed.
+
+Lemma if_arg (fT fF : A -> B) :
+ (if b then fT else fF) x = if b then fT x else fF x.
+Proof. by case b. Qed.
+
+(* Turning a boolean "if" form into an application. *)
+Definition if_expr := if b then vT else vF.
+Lemma ifE : (if b then vT else vF) = if_expr. Proof. by []. Qed.
+
+End BoolIf.
+
+(* Core (internal) reflection lemmas, used for the three kinds of views. *)
+
+Section ReflectCore.
+
+Variables (P Q : Prop) (b c : bool).
+
+Hypothesis Hb : reflect P b.
+
+Lemma introNTF : (if c then ~ P else P) -> ~~ b = c.
+Proof. by case c; case Hb. Qed.
+
+Lemma introTF : (if c then P else ~ P) -> b = c.
+Proof. by case c; case Hb. Qed.
+
+Lemma elimNTF : ~~ b = c -> if c then ~ P else P.
+Proof. by move <-; case Hb. Qed.
+
+Lemma elimTF : b = c -> if c then P else ~ P.
+Proof. by move <-; case Hb. Qed.
+
+Lemma equivPif : (Q -> P) -> (P -> Q) -> if b then Q else ~ Q.
+Proof. by case Hb; auto. Qed.
+
+Lemma xorPif : Q \/ P -> ~ (Q /\ P) -> if b then ~ Q else Q.
+Proof. by case Hb => [? _ H ? | ? H _]; case: H. Qed.
+
+End ReflectCore.
+
+(* Internal negated reflection lemmas *)
+Section ReflectNegCore.
+
+Variables (P Q : Prop) (b c : bool).
+Hypothesis Hb : reflect P (~~ b).
+
+Lemma introTFn : (if c then ~ P else P) -> b = c.
+Proof. by move/(introNTF Hb) <-; case b. Qed.
+
+Lemma elimTFn : b = c -> if c then ~ P else P.
+Proof. by move <-; apply: (elimNTF Hb); case b. Qed.
+
+Lemma equivPifn : (Q -> P) -> (P -> Q) -> if b then ~ Q else Q.
+Proof. by rewrite -if_neg; apply: equivPif. Qed.
+
+Lemma xorPifn : Q \/ P -> ~ (Q /\ P) -> if b then Q else ~ Q.
+Proof. by rewrite -if_neg; apply: xorPif. Qed.
+
+End ReflectNegCore.
+
+(* User-oriented reflection lemmas *)
+Section Reflect.
+
+Variables (P Q : Prop) (b b' c : bool).
+Hypotheses (Pb : reflect P b) (Pb' : reflect P (~~ b')).
+
+Lemma introT : P -> b. Proof. exact: introTF true _. Qed.
+Lemma introF : ~ P -> b = false. Proof. exact: introTF false _. Qed.
+Lemma introN : ~ P -> ~~ b. Proof. exact: introNTF true _. Qed.
+Lemma introNf : P -> ~~ b = false. Proof. exact: introNTF false _. Qed.
+Lemma introTn : ~ P -> b'. Proof. exact: introTFn true _. Qed.
+Lemma introFn : P -> b' = false. Proof. exact: introTFn false _. Qed.
+
+Lemma elimT : b -> P. Proof. exact: elimTF true _. Qed.
+Lemma elimF : b = false -> ~ P. Proof. exact: elimTF false _. Qed.
+Lemma elimN : ~~ b -> ~P. Proof. exact: elimNTF true _. Qed.
+Lemma elimNf : ~~ b = false -> P. Proof. exact: elimNTF false _. Qed.
+Lemma elimTn : b' -> ~ P. Proof. exact: elimTFn true _. Qed.
+Lemma elimFn : b' = false -> P. Proof. exact: elimTFn false _. Qed.
+
+Lemma introP : (b -> Q) -> (~~ b -> ~ Q) -> reflect Q b.
+Proof. by case b; constructor; auto. Qed.
+
+Lemma iffP : (P -> Q) -> (Q -> P) -> reflect Q b.
+Proof. by case: Pb; constructor; auto. Qed.
+
+Lemma equivP : (P <-> Q) -> reflect Q b.
+Proof. by case; apply: iffP. Qed.
+
+Lemma sumboolP (decQ : decidable Q) : reflect Q decQ.
+Proof. by case: decQ; constructor. Qed.
+
+Lemma appP : reflect Q b -> P -> Q.
+Proof. by move=> Qb; move/introT; case: Qb. Qed.
+
+Lemma sameP : reflect P c -> b = c.
+Proof. by case; [apply: introT | apply: introF]. Qed.
+
+Lemma decPcases : if b then P else ~ P. Proof. by case Pb. Qed.
+
+Definition decP : decidable P. by case: b decPcases; [left | right]. Defined.
+
+Lemma rwP : P <-> b. Proof. by split; [apply: introT | apply: elimT]. Qed.
+
+Lemma rwP2 : reflect Q b -> (P <-> Q).
+Proof. by move=> Qb; split=> ?; [apply: appP | apply: elimT; case: Qb]. Qed.
+
+(* Predicate family to reflect excluded middle in bool. *)
+CoInductive alt_spec : bool -> Type :=
+ | AltTrue of P : alt_spec true
+ | AltFalse of ~~ b : alt_spec false.
+
+Lemma altP : alt_spec b.
+Proof. by case def_b: b / Pb; constructor; rewrite ?def_b. Qed.
+
+End Reflect.
+
+Hint View for move/ elimTF|3 elimNTF|3 elimTFn|3 introT|2 introTn|2 introN|2.
+
+Hint View for apply/ introTF|3 introNTF|3 introTFn|3 elimT|2 elimTn|2 elimN|2.
+
+Hint View for apply// equivPif|3 xorPif|3 equivPifn|3 xorPifn|3.
+
+(* Allow the direct application of a reflection lemma to a boolean assertion. *)
+Coercion elimT : reflect >-> Funclass.
+
+CoInductive implies P Q := Implies of P -> Q.
+Lemma impliesP P Q : implies P Q -> P -> Q. Proof. by case. Qed.
+Lemma impliesPn (P Q : Prop) : implies P Q -> ~ Q -> ~ P.
+Proof. by case=> iP ? /iP. Qed.
+Coercion impliesP : implies >-> Funclass.
+Hint View for move/ impliesPn|2 impliesP|2.
+Hint View for apply/ impliesPn|2 impliesP|2.
+
+(* Impredicative or, which can emulate a classical not-implies. *)
+Definition unless condition property : Prop :=
+ forall goal : Prop, (condition -> goal) -> (property -> goal) -> goal.
+
+Notation "\unless C , P" := (unless C P)
+ (at level 200, C at level 100,
+ format "'[' \unless C , '/ ' P ']'") : type_scope.
+
+Lemma unlessL C P : implies C (\unless C, P).
+Proof. by split=> hC G /(_ hC). Qed.
+
+Lemma unlessR C P : implies P (\unless C, P).
+Proof. by split=> hP G _ /(_ hP). Qed.
+
+Lemma unless_sym C P : implies (\unless C, P) (\unless P, C).
+Proof. by split; apply; [apply/unlessR | apply/unlessL]. Qed.
+
+Lemma unlessP (C P : Prop) : (\unless C, P) <-> C \/ P.
+Proof. by split=> [|[/unlessL | /unlessR]]; apply; [left | right]. Qed.
+
+Lemma bind_unless C P {Q} : implies (\unless C, P) (\unless (\unless C, Q), P).
+Proof. by split; apply=> [hC|hP]; [apply/unlessL/unlessL | apply/unlessR]. Qed.
+
+Lemma unless_contra b C : implies (~~ b -> C) (\unless C, b).
+Proof. by split; case: b => [_ | hC]; [apply/unlessR | apply/unlessL/hC]. Qed.
+
+(* Classical reasoning becomes directly accessible for any bool subgoal. *)
+(* Note that we cannot use "unless" here for lack of universe polymorphism. *)
+Definition classically P : Prop := forall b : bool, (P -> b) -> b.
+
+Lemma classicP (P : Prop) : classically P <-> ~ ~ P.
+Proof.
+split=> [cP nP | nnP [] // nP]; last by case nnP; move/nP.
+by have: P -> false; [move/nP | move/cP].
+Qed.
+
+Lemma classicW P : P -> classically P. Proof. by move=> hP _ ->. Qed.
+
+Lemma classic_bind P Q : (P -> classically Q) -> classically P -> classically Q.
+Proof. by move=> iPQ cP b /iPQ-/cP. Qed.
+
+Lemma classic_EM P : classically (decidable P).
+Proof.
+by case=> // undecP; apply/undecP; right=> notP; apply/notF/undecP; left.
+Qed.
+
+Lemma classic_pick T P : classically ({x : T | P x} + (forall x, ~ P x)).
+Proof.
+case=> // undecP; apply/undecP; right=> x Px.
+by apply/notF/undecP; left; exists x.
+Qed.
+
+Lemma classic_imply P Q : (P -> classically Q) -> classically (P -> Q).
+Proof.
+move=> iPQ []// notPQ; apply/notPQ=> /iPQ-cQ.
+by case: notF; apply: cQ => hQ; apply: notPQ.
+Qed.
+
+(* List notations for wider connectives; the Prop connectives have a fixed *)
+(* width so as to avoid iterated destruction (we go up to width 5 for /\, and *)
+(* width 4 for or). The bool connectives have arbitrary widths, but denote *)
+(* expressions that associate to the RIGHT. This is consistent with the right *)
+(* associativity of list expressions and thus more convenient in most proofs. *)
+
+Inductive and3 (P1 P2 P3 : Prop) : Prop := And3 of P1 & P2 & P3.
+
+Inductive and4 (P1 P2 P3 P4 : Prop) : Prop := And4 of P1 & P2 & P3 & P4.
+
+Inductive and5 (P1 P2 P3 P4 P5 : Prop) : Prop :=
+ And5 of P1 & P2 & P3 & P4 & P5.
+
+Inductive or3 (P1 P2 P3 : Prop) : Prop := Or31 of P1 | Or32 of P2 | Or33 of P3.
+
+Inductive or4 (P1 P2 P3 P4 : Prop) : Prop :=
+ Or41 of P1 | Or42 of P2 | Or43 of P3 | Or44 of P4.
+
+Notation "[ /\ P1 & P2 ]" := (and P1 P2) (only parsing) : type_scope.
+Notation "[ /\ P1 , P2 & P3 ]" := (and3 P1 P2 P3) : type_scope.
+Notation "[ /\ P1 , P2 , P3 & P4 ]" := (and4 P1 P2 P3 P4) : type_scope.
+Notation "[ /\ P1 , P2 , P3 , P4 & P5 ]" := (and5 P1 P2 P3 P4 P5) : type_scope.
+
+Notation "[ \/ P1 | P2 ]" := (or P1 P2) (only parsing) : type_scope.
+Notation "[ \/ P1 , P2 | P3 ]" := (or3 P1 P2 P3) : type_scope.
+Notation "[ \/ P1 , P2 , P3 | P4 ]" := (or4 P1 P2 P3 P4) : type_scope.
+
+Notation "[ && b1 & c ]" := (b1 && c) (only parsing) : bool_scope.
+Notation "[ && b1 , b2 , .. , bn & c ]" := (b1 && (b2 && .. (bn && c) .. ))
+ : bool_scope.
+
+Notation "[ || b1 | c ]" := (b1 || c) (only parsing) : bool_scope.
+Notation "[ || b1 , b2 , .. , bn | c ]" := (b1 || (b2 || .. (bn || c) .. ))
+ : bool_scope.
+
+Notation "[ ==> b1 , b2 , .. , bn => c ]" :=
+ (b1 ==> (b2 ==> .. (bn ==> c) .. )) : bool_scope.
+Notation "[ ==> b1 => c ]" := (b1 ==> c) (only parsing) : bool_scope.
+
+Section AllAnd.
+
+Variables (T : Type) (P1 P2 P3 P4 P5 : T -> Prop).
+Local Notation a P := (forall x, P x).
+
+Lemma all_and2 : implies (forall x, [/\ P1 x & P2 x]) [/\ a P1 & a P2].
+Proof. by split=> haveP; split=> x; case: (haveP x). Qed.
+
+Lemma all_and3 : implies (forall x, [/\ P1 x, P2 x & P3 x])
+ [/\ a P1, a P2 & a P3].
+Proof. by split=> haveP; split=> x; case: (haveP x). Qed.
+
+Lemma all_and4 : implies (forall x, [/\ P1 x, P2 x, P3 x & P4 x])
+ [/\ a P1, a P2, a P3 & a P4].
+Proof. by split=> haveP; split=> x; case: (haveP x). Qed.
+
+Lemma all_and5 : implies (forall x, [/\ P1 x, P2 x, P3 x, P4 x & P5 x])
+ [/\ a P1, a P2, a P3, a P4 & a P5].
+Proof. by split=> haveP; split=> x; case: (haveP x). Qed.
+
+End AllAnd.
+
+Arguments all_and2 {T P1 P2}.
+Arguments all_and3 {T P1 P2 P3}.
+Arguments all_and4 {T P1 P2 P3 P4}.
+Arguments all_and5 {T P1 P2 P3 P4 P5}.
+
+Lemma pair_andP P Q : P /\ Q <-> P * Q. Proof. by split; case. Qed.
+
+Section ReflectConnectives.
+
+Variable b1 b2 b3 b4 b5 : bool.
+
+Lemma idP : reflect b1 b1.
+Proof. by case b1; constructor. Qed.
+
+Lemma boolP : alt_spec b1 b1 b1.
+Proof. exact: (altP idP). Qed.
+
+Lemma idPn : reflect (~~ b1) (~~ b1).
+Proof. by case b1; constructor. Qed.
+
+Lemma negP : reflect (~ b1) (~~ b1).
+Proof. by case b1; constructor; auto. Qed.
+
+Lemma negPn : reflect b1 (~~ ~~ b1).
+Proof. by case b1; constructor. Qed.
+
+Lemma negPf : reflect (b1 = false) (~~ b1).
+Proof. by case b1; constructor. Qed.
+
+Lemma andP : reflect (b1 /\ b2) (b1 && b2).
+Proof. by case b1; case b2; constructor=> //; case. Qed.
+
+Lemma and3P : reflect [/\ b1, b2 & b3] [&& b1, b2 & b3].
+Proof. by case b1; case b2; case b3; constructor; try by case. Qed.
+
+Lemma and4P : reflect [/\ b1, b2, b3 & b4] [&& b1, b2, b3 & b4].
+Proof. by case b1; case b2; case b3; case b4; constructor; try by case. Qed.
+
+Lemma and5P : reflect [/\ b1, b2, b3, b4 & b5] [&& b1, b2, b3, b4 & b5].
+Proof.
+by case b1; case b2; case b3; case b4; case b5; constructor; try by case.
+Qed.
+
+Lemma orP : reflect (b1 \/ b2) (b1 || b2).
+Proof. by case b1; case b2; constructor; auto; case. Qed.
+
+Lemma or3P : reflect [\/ b1, b2 | b3] [|| b1, b2 | b3].
+Proof.
+case b1; first by constructor; constructor 1.
+case b2; first by constructor; constructor 2.
+case b3; first by constructor; constructor 3.
+by constructor; case.
+Qed.
+
+Lemma or4P : reflect [\/ b1, b2, b3 | b4] [|| b1, b2, b3 | b4].
+Proof.
+case b1; first by constructor; constructor 1.
+case b2; first by constructor; constructor 2.
+case b3; first by constructor; constructor 3.
+case b4; first by constructor; constructor 4.
+by constructor; case.
+Qed.
+
+Lemma nandP : reflect (~~ b1 \/ ~~ b2) (~~ (b1 && b2)).
+Proof. by case b1; case b2; constructor; auto; case; auto. Qed.
+
+Lemma norP : reflect (~~ b1 /\ ~~ b2) (~~ (b1 || b2)).
+Proof. by case b1; case b2; constructor; auto; case; auto. Qed.
+
+Lemma implyP : reflect (b1 -> b2) (b1 ==> b2).
+Proof. by case b1; case b2; constructor; auto. Qed.
+
+End ReflectConnectives.
+
+Arguments idP [b1].
+Arguments idPn [b1].
+Arguments negP [b1].
+Arguments negPn [b1].
+Arguments negPf [b1].
+Arguments andP [b1 b2].
+Arguments and3P [b1 b2 b3].
+Arguments and4P [b1 b2 b3 b4].
+Arguments and5P [b1 b2 b3 b4 b5].
+Arguments orP [b1 b2].
+Arguments or3P [b1 b2 b3].
+Arguments or4P [b1 b2 b3 b4].
+Arguments nandP [b1 b2].
+Arguments norP [b1 b2].
+Arguments implyP [b1 b2].
+Prenex Implicits idP idPn negP negPn negPf.
+Prenex Implicits andP and3P and4P and5P orP or3P or4P nandP norP implyP.
+
+(* Shorter, more systematic names for the boolean connectives laws. *)
+
+Lemma andTb : left_id true andb. Proof. by []. Qed.
+Lemma andFb : left_zero false andb. Proof. by []. Qed.
+Lemma andbT : right_id true andb. Proof. by case. Qed.
+Lemma andbF : right_zero false andb. Proof. by case. Qed.
+Lemma andbb : idempotent andb. Proof. by case. Qed.
+Lemma andbC : commutative andb. Proof. by do 2!case. Qed.
+Lemma andbA : associative andb. Proof. by do 3!case. Qed.
+Lemma andbCA : left_commutative andb. Proof. by do 3!case. Qed.
+Lemma andbAC : right_commutative andb. Proof. by do 3!case. Qed.
+Lemma andbACA : interchange andb andb. Proof. by do 4!case. Qed.
+
+Lemma orTb : forall b, true || b. Proof. by []. Qed.
+Lemma orFb : left_id false orb. Proof. by []. Qed.
+Lemma orbT : forall b, b || true. Proof. by case. Qed.
+Lemma orbF : right_id false orb. Proof. by case. Qed.
+Lemma orbb : idempotent orb. Proof. by case. Qed.
+Lemma orbC : commutative orb. Proof. by do 2!case. Qed.
+Lemma orbA : associative orb. Proof. by do 3!case. Qed.
+Lemma orbCA : left_commutative orb. Proof. by do 3!case. Qed.
+Lemma orbAC : right_commutative orb. Proof. by do 3!case. Qed.
+Lemma orbACA : interchange orb orb. Proof. by do 4!case. Qed.
+
+Lemma andbN b : b && ~~ b = false. Proof. by case: b. Qed.
+Lemma andNb b : ~~ b && b = false. Proof. by case: b. Qed.
+Lemma orbN b : b || ~~ b = true. Proof. by case: b. Qed.
+Lemma orNb b : ~~ b || b = true. Proof. by case: b. Qed.
+
+Lemma andb_orl : left_distributive andb orb. Proof. by do 3!case. Qed.
+Lemma andb_orr : right_distributive andb orb. Proof. by do 3!case. Qed.
+Lemma orb_andl : left_distributive orb andb. Proof. by do 3!case. Qed.
+Lemma orb_andr : right_distributive orb andb. Proof. by do 3!case. Qed.
+
+Lemma andb_idl (a b : bool) : (b -> a) -> a && b = b.
+Proof. by case: a; case: b => // ->. Qed.
+Lemma andb_idr (a b : bool) : (a -> b) -> a && b = a.
+Proof. by case: a; case: b => // ->. Qed.
+Lemma andb_id2l (a b c : bool) : (a -> b = c) -> a && b = a && c.
+Proof. by case: a; case: b; case: c => // ->. Qed.
+Lemma andb_id2r (a b c : bool) : (b -> a = c) -> a && b = c && b.
+Proof. by case: a; case: b; case: c => // ->. Qed.
+
+Lemma orb_idl (a b : bool) : (a -> b) -> a || b = b.
+Proof. by case: a; case: b => // ->. Qed.
+Lemma orb_idr (a b : bool) : (b -> a) -> a || b = a.
+Proof. by case: a; case: b => // ->. Qed.
+Lemma orb_id2l (a b c : bool) : (~~ a -> b = c) -> a || b = a || c.
+Proof. by case: a; case: b; case: c => // ->. Qed.
+Lemma orb_id2r (a b c : bool) : (~~ b -> a = c) -> a || b = c || b.
+Proof. by case: a; case: b; case: c => // ->. Qed.
+
+Lemma negb_and (a b : bool) : ~~ (a && b) = ~~ a || ~~ b.
+Proof. by case: a; case: b. Qed.
+
+Lemma negb_or (a b : bool) : ~~ (a || b) = ~~ a && ~~ b.
+Proof. by case: a; case: b. Qed.
+
+(* Pseudo-cancellation -- i.e, absorbtion *)
+
+Lemma andbK a b : a && b || a = a. Proof. by case: a; case: b. Qed.
+Lemma andKb a b : a || b && a = a. Proof. by case: a; case: b. Qed.
+Lemma orbK a b : (a || b) && a = a. Proof. by case: a; case: b. Qed.
+Lemma orKb a b : a && (b || a) = a. Proof. by case: a; case: b. Qed.
+
+(* Imply *)
+
+Lemma implybT b : b ==> true. Proof. by case: b. Qed.
+Lemma implybF b : (b ==> false) = ~~ b. Proof. by case: b. Qed.
+Lemma implyFb b : false ==> b. Proof. by []. Qed.
+Lemma implyTb b : (true ==> b) = b. Proof. by []. Qed.
+Lemma implybb b : b ==> b. Proof. by case: b. Qed.
+
+Lemma negb_imply a b : ~~ (a ==> b) = a && ~~ b.
+Proof. by case: a; case: b. Qed.
+
+Lemma implybE a b : (a ==> b) = ~~ a || b.
+Proof. by case: a; case: b. Qed.
+
+Lemma implyNb a b : (~~ a ==> b) = a || b.
+Proof. by case: a; case: b. Qed.
+
+Lemma implybN a b : (a ==> ~~ b) = (b ==> ~~ a).
+Proof. by case: a; case: b. Qed.
+
+Lemma implybNN a b : (~~ a ==> ~~ b) = b ==> a.
+Proof. by case: a; case: b. Qed.
+
+Lemma implyb_idl (a b : bool) : (~~ a -> b) -> (a ==> b) = b.
+Proof. by case: a; case: b => // ->. Qed.
+Lemma implyb_idr (a b : bool) : (b -> ~~ a) -> (a ==> b) = ~~ a.
+Proof. by case: a; case: b => // ->. Qed.
+Lemma implyb_id2l (a b c : bool) : (a -> b = c) -> (a ==> b) = (a ==> c).
+Proof. by case: a; case: b; case: c => // ->. Qed.
+
+(* Addition (xor) *)
+
+Lemma addFb : left_id false addb. Proof. by []. Qed.
+Lemma addbF : right_id false addb. Proof. by case. Qed.
+Lemma addbb : self_inverse false addb. Proof. by case. Qed.
+Lemma addbC : commutative addb. Proof. by do 2!case. Qed.
+Lemma addbA : associative addb. Proof. by do 3!case. Qed.
+Lemma addbCA : left_commutative addb. Proof. by do 3!case. Qed.
+Lemma addbAC : right_commutative addb. Proof. by do 3!case. Qed.
+Lemma addbACA : interchange addb addb. Proof. by do 4!case. Qed.
+Lemma andb_addl : left_distributive andb addb. Proof. by do 3!case. Qed.
+Lemma andb_addr : right_distributive andb addb. Proof. by do 3!case. Qed.
+Lemma addKb : left_loop id addb. Proof. by do 2!case. Qed.
+Lemma addbK : right_loop id addb. Proof. by do 2!case. Qed.
+Lemma addIb : left_injective addb. Proof. by do 3!case. Qed.
+Lemma addbI : right_injective addb. Proof. by do 3!case. Qed.
+
+Lemma addTb b : true (+) b = ~~ b. Proof. by []. Qed.
+Lemma addbT b : b (+) true = ~~ b. Proof. by case: b. Qed.
+
+Lemma addbN a b : a (+) ~~ b = ~~ (a (+) b).
+Proof. by case: a; case: b. Qed.
+Lemma addNb a b : ~~ a (+) b = ~~ (a (+) b).
+Proof. by case: a; case: b. Qed.
+
+Lemma addbP a b : reflect (~~ a = b) (a (+) b).
+Proof. by case: a; case: b; constructor. Qed.
+Arguments addbP [a b].
+
+(* Resolution tactic for blindly weeding out common terms from boolean *)
+(* equalities. When faced with a goal of the form (andb/orb/addb b1 b2) = b3 *)
+(* they will try to locate b1 in b3 and remove it. This can fail! *)
+
+Ltac bool_congr :=
+ match goal with
+ | |- (?X1 && ?X2 = ?X3) => first
+ [ symmetry; rewrite -1?(andbC X1) -?(andbCA X1); congr 1 (andb X1); symmetry
+ | case: (X1); [ rewrite ?andTb ?andbT // | by rewrite ?andbF /= ] ]
+ | |- (?X1 || ?X2 = ?X3) => first
+ [ symmetry; rewrite -1?(orbC X1) -?(orbCA X1); congr 1 (orb X1); symmetry
+ | case: (X1); [ by rewrite ?orbT //= | rewrite ?orFb ?orbF ] ]
+ | |- (?X1 (+) ?X2 = ?X3) =>
+ symmetry; rewrite -1?(addbC X1) -?(addbCA X1); congr 1 (addb X1); symmetry
+ | |- (~~ ?X1 = ?X2) => congr 1 negb
+ end.
+
+(******************************************************************************)
+(* Predicates, i.e., packaged functions to bool. *)
+(* - pred T, the basic type for predicates over a type T, is simply an alias *)
+(* for T -> bool. *)
+(* We actually distinguish two kinds of predicates, which we call applicative *)
+(* and collective, based on the syntax used to test them at some x in T: *)
+(* - For an applicative predicate P, one uses prefix syntax: *)
+(* P x *)
+(* Also, most operations on applicative predicates use prefix syntax as *)
+(* well (e.g., predI P Q). *)
+(* - For a collective predicate A, one uses infix syntax: *)
+(* x \in A *)
+(* and all operations on collective predicates use infix syntax as well *)
+(* (e.g., [predI A & B]). *)
+(* There are only two kinds of applicative predicates: *)
+(* - pred T, the alias for T -> bool mentioned above *)
+(* - simpl_pred T, an alias for simpl_fun T bool with a coercion to pred T *)
+(* that auto-simplifies on application (see ssrfun). *)
+(* On the other hand, the set of collective predicate types is open-ended via *)
+(* - predType T, a Structure that can be used to put Canonical collective *)
+(* predicate interpretation on other types, such as lists, tuples, *)
+(* finite sets, etc. *)
+(* Indeed, we define such interpretations for applicative predicate types, *)
+(* which can therefore also be used with the infix syntax, e.g., *)
+(* x \in predI P Q *)
+(* Moreover these infix forms are convertible to their prefix counterpart *)
+(* (e.g., predI P Q x which in turn simplifies to P x && Q x). The converse *)
+(* is not true, however; collective predicate types cannot, in general, be *)
+(* general, be used applicatively, because of the "uniform inheritance" *)
+(* restriction on implicit coercions. *)
+(* However, we do define an explicit generic coercion *)
+(* - mem : forall (pT : predType), pT -> mem_pred T *)
+(* where mem_pred T is a variant of simpl_pred T that preserves the infix *)
+(* syntax, i.e., mem A x auto-simplifies to x \in A. *)
+(* Indeed, the infix "collective" operators are notation for a prefix *)
+(* operator with arguments of type mem_pred T or pred T, applied to coerced *)
+(* collective predicates, e.g., *)
+(* Notation "x \in A" := (in_mem x (mem A)). *)
+(* This prevents the variability in the predicate type from interfering with *)
+(* the application of generic lemmas. Moreover this also makes it much easier *)
+(* to define generic lemmas, because the simplest type -- pred T -- can be *)
+(* used as the type of generic collective predicates, provided one takes care *)
+(* not to use it applicatively; this avoids the burden of having to declare a *)
+(* different predicate type for each predicate parameter of each section or *)
+(* lemma. *)
+(* This trick is made possible by the fact that the constructor of the *)
+(* mem_pred T type aligns the unification process, forcing a generic *)
+(* "collective" predicate A : pred T to unify with the actual collective B, *)
+(* which mem has coerced to pred T via an internal, hidden implicit coercion, *)
+(* supplied by the predType structure for B. Users should take care not to *)
+(* inadvertently "strip" (mem B) down to the coerced B, since this will *)
+(* expose the internal coercion: Coq will display a term B x that cannot be *)
+(* typed as such. The topredE lemma can be used to restore the x \in B *)
+(* syntax in this case. While -topredE can conversely be used to change *)
+(* x \in P into P x, it is safer to use the inE and memE lemmas instead, as *)
+(* they do not run the risk of exposing internal coercions. As a consequence *)
+(* it is better to explicitly cast a generic applicative pred T to simpl_pred *)
+(* using the SimplPred constructor, when it is used as a collective predicate *)
+(* (see, e.g., Lemma eq_big in bigop). *)
+(* We also sometimes "instantiate" the predType structure by defining a *)
+(* coercion to the sort of the predPredType structure. This works better for *)
+(* types such as {set T} that have subtypes that coerce to them, since the *)
+(* same coercion will be inserted by the application of mem. It also lets us *)
+(* turn any Type aT : predArgType into the total predicate over that type, *)
+(* i.e., fun _: aT => true. This allows us to write, e.g., #|'I_n| for the *)
+(* cardinal of the (finite) type of integers less than n. *)
+(* Collective predicates have a specific extensional equality, *)
+(* - A =i B, *)
+(* while applicative predicates use the extensional equality of functions, *)
+(* - P =1 Q *)
+(* The two forms are convertible, however. *)
+(* We lift boolean operations to predicates, defining: *)
+(* - predU (union), predI (intersection), predC (complement), *)
+(* predD (difference), and preim (preimage, i.e., composition) *)
+(* For each operation we define three forms, typically: *)
+(* - predU : pred T -> pred T -> simpl_pred T *)
+(* - [predU A & B], a Notation for predU (mem A) (mem B) *)
+(* - xpredU, a Notation for the lambda-expression inside predU, *)
+(* which is mostly useful as an argument of =1, since it exposes the head *)
+(* head constant of the expression to the ssreflect matching algorithm. *)
+(* The syntax for the preimage of a collective predicate A is *)
+(* - [preim f of A] *)
+(* Finally, the generic syntax for defining a simpl_pred T is *)
+(* - [pred x : T | P(x)], [pred x | P(x)], [pred x in A | P(x)], etc. *)
+(* We also support boolean relations, but only the applicative form, with *)
+(* types *)
+(* - rel T, an alias for T -> pred T *)
+(* - simpl_rel T, an auto-simplifying version, and syntax *)
+(* [rel x y | P(x,y)], [rel x y in A & B | P(x,y)], etc. *)
+(* The notation [rel of fA] can be used to coerce a function returning a *)
+(* collective predicate to one returning pred T. *)
+(* Finally, note that there is specific support for ambivalent predicates *)
+(* that can work in either style, as per this file's head descriptor. *)
+(******************************************************************************)
+
+Definition pred T := T -> bool.
+
+Identity Coercion fun_of_pred : pred >-> Funclass.
+
+Definition rel T := T -> pred T.
+
+Identity Coercion fun_of_rel : rel >-> Funclass.
+
+Notation xpred0 := (fun _ => false).
+Notation xpredT := (fun _ => true).
+Notation xpredI := (fun (p1 p2 : pred _) x => p1 x && p2 x).
+Notation xpredU := (fun (p1 p2 : pred _) x => p1 x || p2 x).
+Notation xpredC := (fun (p : pred _) x => ~~ p x).
+Notation xpredD := (fun (p1 p2 : pred _) x => ~~ p2 x && p1 x).
+Notation xpreim := (fun f (p : pred _) x => p (f x)).
+Notation xrelU := (fun (r1 r2 : rel _) x y => r1 x y || r2 x y).
+
+Section Predicates.
+
+Variables T : Type.
+
+Definition subpred (p1 p2 : pred T) := forall x, p1 x -> p2 x.
+
+Definition subrel (r1 r2 : rel T) := forall x y, r1 x y -> r2 x y.
+
+Definition simpl_pred := simpl_fun T bool.
+Definition applicative_pred := pred T.
+Definition collective_pred := pred T.
+
+Definition SimplPred (p : pred T) : simpl_pred := SimplFun p.
+
+Coercion pred_of_simpl (p : simpl_pred) : pred T := fun_of_simpl p.
+Coercion applicative_pred_of_simpl (p : simpl_pred) : applicative_pred :=
+ fun_of_simpl p.
+Coercion collective_pred_of_simpl (p : simpl_pred) : collective_pred :=
+ fun x => (let: SimplFun f := p in fun _ => f x) x.
+(* Note: applicative_of_simpl is convertible to pred_of_simpl, while *)
+(* collective_of_simpl is not. *)
+
+Definition pred0 := SimplPred xpred0.
+Definition predT := SimplPred xpredT.
+Definition predI p1 p2 := SimplPred (xpredI p1 p2).
+Definition predU p1 p2 := SimplPred (xpredU p1 p2).
+Definition predC p := SimplPred (xpredC p).
+Definition predD p1 p2 := SimplPred (xpredD p1 p2).
+Definition preim rT f (d : pred rT) := SimplPred (xpreim f d).
+
+Definition simpl_rel := simpl_fun T (pred T).
+
+Definition SimplRel (r : rel T) : simpl_rel := [fun x => r x].
+
+Coercion rel_of_simpl_rel (r : simpl_rel) : rel T := fun x y => r x y.
+
+Definition relU r1 r2 := SimplRel (xrelU r1 r2).
+
+Lemma subrelUl r1 r2 : subrel r1 (relU r1 r2).
+Proof. by move=> *; apply/orP; left. Qed.
+
+Lemma subrelUr r1 r2 : subrel r2 (relU r1 r2).
+Proof. by move=> *; apply/orP; right. Qed.
+
+CoInductive mem_pred := Mem of pred T.
+
+Definition isMem pT topred mem := mem = (fun p : pT => Mem [eta topred p]).
+
+Structure predType := PredType {
+ pred_sort :> Type;
+ topred : pred_sort -> pred T;
+ _ : {mem | isMem topred mem}
+}.
+
+Definition mkPredType pT toP := PredType (exist (@isMem pT toP) _ (erefl _)).
+
+Canonical predPredType := Eval hnf in @mkPredType (pred T) id.
+Canonical simplPredType := Eval hnf in mkPredType pred_of_simpl.
+Canonical boolfunPredType := Eval hnf in @mkPredType (T -> bool) id.
+
+Coercion pred_of_mem mp : pred_sort predPredType := let: Mem p := mp in [eta p].
+Canonical memPredType := Eval hnf in mkPredType pred_of_mem.
+
+Definition clone_pred U :=
+ fun pT & pred_sort pT -> U =>
+ fun a mP (pT' := @PredType U a mP) & phant_id pT' pT => pT'.
+
+End Predicates.
+
+Arguments pred0 [T].
+Arguments predT [T].
+Prenex Implicits pred0 predT predI predU predC predD preim relU.
+
+Notation "[ 'pred' : T | E ]" := (SimplPred (fun _ : T => E%B))
+ (at level 0, format "[ 'pred' : T | E ]") : fun_scope.
+Notation "[ 'pred' x | E ]" := (SimplPred (fun x => E%B))
+ (at level 0, x ident, format "[ 'pred' x | E ]") : fun_scope.
+Notation "[ 'pred' x | E1 & E2 ]" := [pred x | E1 && E2 ]
+ (at level 0, x ident, format "[ 'pred' x | E1 & E2 ]") : fun_scope.
+Notation "[ 'pred' x : T | E ]" := (SimplPred (fun x : T => E%B))
+ (at level 0, x ident, only parsing) : fun_scope.
+Notation "[ 'pred' x : T | E1 & E2 ]" := [pred x : T | E1 && E2 ]
+ (at level 0, x ident, only parsing) : fun_scope.
+Notation "[ 'rel' x y | E ]" := (SimplRel (fun x y => E%B))
+ (at level 0, x ident, y ident, format "[ 'rel' x y | E ]") : fun_scope.
+Notation "[ 'rel' x y : T | E ]" := (SimplRel (fun x y : T => E%B))
+ (at level 0, x ident, y ident, only parsing) : fun_scope.
+
+Notation "[ 'predType' 'of' T ]" := (@clone_pred _ T _ id _ _ id)
+ (at level 0, format "[ 'predType' 'of' T ]") : form_scope.
+
+(* This redundant coercion lets us "inherit" the simpl_predType canonical *)
+(* instance by declaring a coercion to simpl_pred. This hack is the only way *)
+(* to put a predType structure on a predArgType. We use simpl_pred rather *)
+(* than pred to ensure that /= removes the identity coercion. Note that the *)
+(* coercion will never be used directly for simpl_pred, since the canonical *)
+(* instance should always be resolved. *)
+
+Notation pred_class := (pred_sort (predPredType _)).
+Coercion sort_of_simpl_pred T (p : simpl_pred T) : pred_class := p : pred T.
+
+(* This lets us use some types as a synonym for their universal predicate. *)
+(* Unfortunately, this won't work for existing types like bool, unless we *)
+(* redefine bool, true, false and all bool ops. *)
+Definition predArgType := Type.
+Bind Scope type_scope with predArgType.
+Identity Coercion sort_of_predArgType : predArgType >-> Sortclass.
+Coercion pred_of_argType (T : predArgType) : simpl_pred T := predT.
+
+Notation "{ : T }" := (T%type : predArgType)
+ (at level 0, format "{ : T }") : type_scope.
+
+(* These must be defined outside a Section because "cooking" kills the *)
+(* nosimpl tag. *)
+
+Definition mem T (pT : predType T) : pT -> mem_pred T :=
+ nosimpl (let: @PredType _ _ _ (exist _ mem _) := pT return pT -> _ in mem).
+Definition in_mem T x mp := nosimpl pred_of_mem T mp x.
+
+Prenex Implicits mem.
+
+Coercion pred_of_mem_pred T mp := [pred x : T | in_mem x mp].
+
+Definition eq_mem T p1 p2 := forall x : T, in_mem x p1 = in_mem x p2.
+Definition sub_mem T p1 p2 := forall x : T, in_mem x p1 -> in_mem x p2.
+
+Typeclasses Opaque eq_mem.
+
+Lemma sub_refl T (p : mem_pred T) : sub_mem p p. Proof. by []. Qed.
+Arguments sub_refl {T p}.
+
+Notation "x \in A" := (in_mem x (mem A)) : bool_scope.
+Notation "x \in A" := (in_mem x (mem A)) : bool_scope.
+Notation "x \notin A" := (~~ (x \in A)) : bool_scope.
+Notation "A =i B" := (eq_mem (mem A) (mem B)) : type_scope.
+Notation "{ 'subset' A <= B }" := (sub_mem (mem A) (mem B))
+ (at level 0, A, B at level 69,
+ format "{ '[hv' 'subset' A '/ ' <= B ']' }") : type_scope.
+Notation "[ 'mem' A ]" := (pred_of_simpl (pred_of_mem_pred (mem A)))
+ (at level 0, only parsing) : fun_scope.
+Notation "[ 'rel' 'of' fA ]" := (fun x => [mem (fA x)])
+ (at level 0, format "[ 'rel' 'of' fA ]") : fun_scope.
+Notation "[ 'predI' A & B ]" := (predI [mem A] [mem B])
+ (at level 0, format "[ 'predI' A & B ]") : fun_scope.
+Notation "[ 'predU' A & B ]" := (predU [mem A] [mem B])
+ (at level 0, format "[ 'predU' A & B ]") : fun_scope.
+Notation "[ 'predD' A & B ]" := (predD [mem A] [mem B])
+ (at level 0, format "[ 'predD' A & B ]") : fun_scope.
+Notation "[ 'predC' A ]" := (predC [mem A])
+ (at level 0, format "[ 'predC' A ]") : fun_scope.
+Notation "[ 'preim' f 'of' A ]" := (preim f [mem A])
+ (at level 0, format "[ 'preim' f 'of' A ]") : fun_scope.
+
+Notation "[ 'pred' x 'in' A ]" := [pred x | x \in A]
+ (at level 0, x ident, format "[ 'pred' x 'in' A ]") : fun_scope.
+Notation "[ 'pred' x 'in' A | E ]" := [pred x | x \in A & E]
+ (at level 0, x ident, format "[ 'pred' x 'in' A | E ]") : fun_scope.
+Notation "[ 'pred' x 'in' A | E1 & E2 ]" := [pred x | x \in A & E1 && E2 ]
+ (at level 0, x ident,
+ format "[ 'pred' x 'in' A | E1 & E2 ]") : fun_scope.
+Notation "[ 'rel' x y 'in' A & B | E ]" :=
+ [rel x y | (x \in A) && (y \in B) && E]
+ (at level 0, x ident, y ident,
+ format "[ 'rel' x y 'in' A & B | E ]") : fun_scope.
+Notation "[ 'rel' x y 'in' A & B ]" := [rel x y | (x \in A) && (y \in B)]
+ (at level 0, x ident, y ident,
+ format "[ 'rel' x y 'in' A & B ]") : fun_scope.
+Notation "[ 'rel' x y 'in' A | E ]" := [rel x y in A & A | E]
+ (at level 0, x ident, y ident,
+ format "[ 'rel' x y 'in' A | E ]") : fun_scope.
+Notation "[ 'rel' x y 'in' A ]" := [rel x y in A & A]
+ (at level 0, x ident, y ident,
+ format "[ 'rel' x y 'in' A ]") : fun_scope.
+
+Section simpl_mem.
+
+Variables (T : Type) (pT : predType T).
+Implicit Types (x : T) (p : pred T) (sp : simpl_pred T) (pp : pT).
+
+(* Bespoke structures that provide fine-grained control over matching the *)
+(* various forms of the \in predicate; note in particular the different forms *)
+(* of hoisting that are used. We had to work around several bugs in the *)
+(* implementation of unification, notably improper expansion of telescope *)
+(* projections and overwriting of a variable assignment by a later *)
+(* unification (probably due to conversion cache cross-talk). *)
+Structure manifest_applicative_pred p := ManifestApplicativePred {
+ manifest_applicative_pred_value :> pred T;
+ _ : manifest_applicative_pred_value = p
+}.
+Definition ApplicativePred p := ManifestApplicativePred (erefl p).
+Canonical applicative_pred_applicative sp :=
+ ApplicativePred (applicative_pred_of_simpl sp).
+
+Structure manifest_simpl_pred p := ManifestSimplPred {
+ manifest_simpl_pred_value :> simpl_pred T;
+ _ : manifest_simpl_pred_value = SimplPred p
+}.
+Canonical expose_simpl_pred p := ManifestSimplPred (erefl (SimplPred p)).
+
+Structure manifest_mem_pred p := ManifestMemPred {
+ manifest_mem_pred_value :> mem_pred T;
+ _ : manifest_mem_pred_value= Mem [eta p]
+}.
+Canonical expose_mem_pred p := @ManifestMemPred p _ (erefl _).
+
+Structure applicative_mem_pred p :=
+ ApplicativeMemPred {applicative_mem_pred_value :> manifest_mem_pred p}.
+Canonical check_applicative_mem_pred p (ap : manifest_applicative_pred p) mp :=
+ @ApplicativeMemPred ap mp.
+
+Lemma mem_topred (pp : pT) : mem (topred pp) = mem pp.
+Proof. by rewrite /mem; case: pT pp => T1 app1 [mem1 /= ->]. Qed.
+
+Lemma topredE x (pp : pT) : topred pp x = (x \in pp).
+Proof. by rewrite -mem_topred. Qed.
+
+Lemma app_predE x p (ap : manifest_applicative_pred p) : ap x = (x \in p).
+Proof. by case: ap => _ /= ->. Qed.
+
+Lemma in_applicative x p (amp : applicative_mem_pred p) : in_mem x amp = p x.
+Proof. by case: amp => [[_ /= ->]]. Qed.
+
+Lemma in_collective x p (msp : manifest_simpl_pred p) :
+ (x \in collective_pred_of_simpl msp) = p x.
+Proof. by case: msp => _ /= ->. Qed.
+
+Lemma in_simpl x p (msp : manifest_simpl_pred p) :
+ in_mem x (Mem [eta fun_of_simpl (msp : simpl_pred T)]) = p x.
+Proof. by case: msp => _ /= ->. Qed.
+
+(* Because of the explicit eta expansion in the left-hand side, this lemma *)
+(* should only be used in a right-to-left direction. The 8.3 hack allowing *)
+(* partial right-to-left use does not work with the improved expansion *)
+(* heuristics in 8.4. *)
+Lemma unfold_in x p : (x \in ([eta p] : pred T)) = p x.
+Proof. by []. Qed.
+
+Lemma simpl_predE p : SimplPred p =1 p.
+Proof. by []. Qed.
+
+Definition inE := (in_applicative, in_simpl, simpl_predE). (* to be extended *)
+
+Lemma mem_simpl sp : mem sp = sp :> pred T.
+Proof. by []. Qed.
+
+Definition memE := mem_simpl. (* could be extended *)
+
+Lemma mem_mem (pp : pT) : (mem (mem pp) = mem pp) * (mem [mem pp] = mem pp).
+Proof. by rewrite -mem_topred. Qed.
+
+End simpl_mem.
+
+(* Qualifiers and keyed predicates. *)
+
+CoInductive qualifier (q : nat) T := Qualifier of predPredType T.
+
+Coercion has_quality n T (q : qualifier n T) : pred_class :=
+ fun x => let: Qualifier _ p := q in p x.
+Arguments has_quality n [T].
+
+Lemma qualifE n T p x : (x \in @Qualifier n T p) = p x. Proof. by []. Qed.
+
+Notation "x \is A" := (x \in has_quality 0 A)
+ (at level 70, no associativity,
+ format "'[hv' x '/ ' \is A ']'") : bool_scope.
+Notation "x \is 'a' A" := (x \in has_quality 1 A)
+ (at level 70, no associativity,
+ format "'[hv' x '/ ' \is 'a' A ']'") : bool_scope.
+Notation "x \is 'an' A" := (x \in has_quality 2 A)
+ (at level 70, no associativity,
+ format "'[hv' x '/ ' \is 'an' A ']'") : bool_scope.
+Notation "x \isn't A" := (x \notin has_quality 0 A)
+ (at level 70, no associativity,
+ format "'[hv' x '/ ' \isn't A ']'") : bool_scope.
+Notation "x \isn't 'a' A" := (x \notin has_quality 1 A)
+ (at level 70, no associativity,
+ format "'[hv' x '/ ' \isn't 'a' A ']'") : bool_scope.
+Notation "x \isn't 'an' A" := (x \notin has_quality 2 A)
+ (at level 70, no associativity,
+ format "'[hv' x '/ ' \isn't 'an' A ']'") : bool_scope.
+Notation "[ 'qualify' x | P ]" := (Qualifier 0 (fun x => P%B))
+ (at level 0, x at level 99,
+ format "'[hv' [ 'qualify' x | '/ ' P ] ']'") : form_scope.
+Notation "[ 'qualify' x : T | P ]" := (Qualifier 0 (fun x : T => P%B))
+ (at level 0, x at level 99, only parsing) : form_scope.
+Notation "[ 'qualify' 'a' x | P ]" := (Qualifier 1 (fun x => P%B))
+ (at level 0, x at level 99,
+ format "'[hv' [ 'qualify' 'a' x | '/ ' P ] ']'") : form_scope.
+Notation "[ 'qualify' 'a' x : T | P ]" := (Qualifier 1 (fun x : T => P%B))
+ (at level 0, x at level 99, only parsing) : form_scope.
+Notation "[ 'qualify' 'an' x | P ]" := (Qualifier 2 (fun x => P%B))
+ (at level 0, x at level 99,
+ format "'[hv' [ 'qualify' 'an' x | '/ ' P ] ']'") : form_scope.
+Notation "[ 'qualify' 'an' x : T | P ]" := (Qualifier 2 (fun x : T => P%B))
+ (at level 0, x at level 99, only parsing) : form_scope.
+
+(* Keyed predicates: support for property-bearing predicate interfaces. *)
+
+Section KeyPred.
+
+Variable T : Type.
+CoInductive pred_key (p : predPredType T) := DefaultPredKey.
+
+Variable p : predPredType T.
+Structure keyed_pred (k : pred_key p) :=
+ PackKeyedPred {unkey_pred :> pred_class; _ : unkey_pred =i p}.
+
+Variable k : pred_key p.
+Definition KeyedPred := @PackKeyedPred k p (frefl _).
+
+Variable k_p : keyed_pred k.
+Lemma keyed_predE : k_p =i p. Proof. by case: k_p. Qed.
+
+(* Instances that strip the mem cast; the first one has "pred_of_mem" as its *)
+(* projection head value, while the second has "pred_of_simpl". The latter *)
+(* has the side benefit of preempting accidental misdeclarations. *)
+(* Note: pred_of_mem is the registered mem >-> pred_class coercion, while *)
+(* simpl_of_mem; pred_of_simpl is the mem >-> pred >=> Funclass coercion. We *)
+(* must write down the coercions explicitly as the Canonical head constant *)
+(* computation does not strip casts !! *)
+Canonical keyed_mem :=
+ @PackKeyedPred k (pred_of_mem (mem k_p)) keyed_predE.
+Canonical keyed_mem_simpl :=
+ @PackKeyedPred k (pred_of_simpl (mem k_p)) keyed_predE.
+
+End KeyPred.
+
+Notation "x \i 'n' S" := (x \in @unkey_pred _ S _ _)
+ (at level 70, format "'[hv' x '/ ' \i 'n' S ']'") : bool_scope.
+
+Section KeyedQualifier.
+
+Variables (T : Type) (n : nat) (q : qualifier n T).
+
+Structure keyed_qualifier (k : pred_key q) :=
+ PackKeyedQualifier {unkey_qualifier; _ : unkey_qualifier = q}.
+Definition KeyedQualifier k := PackKeyedQualifier k (erefl q).
+Variables (k : pred_key q) (k_q : keyed_qualifier k).
+Fact keyed_qualifier_suproof : unkey_qualifier k_q =i q.
+Proof. by case: k_q => /= _ ->. Qed.
+Canonical keyed_qualifier_keyed := PackKeyedPred k keyed_qualifier_suproof.
+
+End KeyedQualifier.
+
+Notation "x \i 's' A" := (x \i n has_quality 0 A)
+ (at level 70, format "'[hv' x '/ ' \i 's' A ']'") : bool_scope.
+Notation "x \i 's' 'a' A" := (x \i n has_quality 1 A)
+ (at level 70, format "'[hv' x '/ ' \i 's' 'a' A ']'") : bool_scope.
+Notation "x \i 's' 'an' A" := (x \i n has_quality 2 A)
+ (at level 70, format "'[hv' x '/ ' \i 's' 'an' A ']'") : bool_scope.
+
+Module DefaultKeying.
+
+Canonical default_keyed_pred T p := KeyedPred (@DefaultPredKey T p).
+Canonical default_keyed_qualifier T n (q : qualifier n T) :=
+ KeyedQualifier (DefaultPredKey q).
+
+End DefaultKeying.
+
+(* Skolemizing with conditions. *)
+
+Lemma all_tag_cond_dep I T (C : pred I) U :
+ (forall x, T x) -> (forall x, C x -> {y : T x & U x y}) ->
+ {f : forall x, T x & forall x, C x -> U x (f x)}.
+Proof.
+move=> f0 fP; apply: all_tag (fun x y => C x -> U x y) _ => x.
+by case Cx: (C x); [case/fP: Cx => y; exists y | exists (f0 x)].
+Qed.
+
+Lemma all_tag_cond I T (C : pred I) U :
+ T -> (forall x, C x -> {y : T & U x y}) ->
+ {f : I -> T & forall x, C x -> U x (f x)}.
+Proof. by move=> y0; apply: all_tag_cond_dep. Qed.
+
+Lemma all_sig_cond_dep I T (C : pred I) P :
+ (forall x, T x) -> (forall x, C x -> {y : T x | P x y}) ->
+ {f : forall x, T x | forall x, C x -> P x (f x)}.
+Proof. by move=> f0 /(all_tag_cond_dep f0)[f]; exists f. Qed.
+
+Lemma all_sig_cond I T (C : pred I) P :
+ T -> (forall x, C x -> {y : T | P x y}) ->
+ {f : I -> T | forall x, C x -> P x (f x)}.
+Proof. by move=> y0; apply: all_sig_cond_dep. Qed.
+
+Section RelationProperties.
+
+(* Caveat: reflexive should not be used to state lemmas, as auto and trivial *)
+(* will not expand the constant. *)
+
+Variable T : Type.
+
+Variable R : rel T.
+
+Definition total := forall x y, R x y || R y x.
+Definition transitive := forall y x z, R x y -> R y z -> R x z.
+
+Definition symmetric := forall x y, R x y = R y x.
+Definition antisymmetric := forall x y, R x y && R y x -> x = y.
+Definition pre_symmetric := forall x y, R x y -> R y x.
+
+Lemma symmetric_from_pre : pre_symmetric -> symmetric.
+Proof. by move=> symR x y; apply/idP/idP; apply: symR. Qed.
+
+Definition reflexive := forall x, R x x.
+Definition irreflexive := forall x, R x x = false.
+
+Definition left_transitive := forall x y, R x y -> R x =1 R y.
+Definition right_transitive := forall x y, R x y -> R^~ x =1 R^~ y.
+
+Section PER.
+
+Hypotheses (symR : symmetric) (trR : transitive).
+
+Lemma sym_left_transitive : left_transitive.
+Proof. by move=> x y Rxy z; apply/idP/idP; apply: trR; rewrite // symR. Qed.
+
+Lemma sym_right_transitive : right_transitive.
+Proof. by move=> x y /sym_left_transitive Rxy z; rewrite !(symR z) Rxy. Qed.
+
+End PER.
+
+(* We define the equivalence property with prenex quantification so that it *)
+(* can be localized using the {in ..., ..} form defined below. *)
+
+Definition equivalence_rel := forall x y z, R z z * (R x y -> R x z = R y z).
+
+Lemma equivalence_relP : equivalence_rel <-> reflexive /\ left_transitive.
+Proof.
+split=> [eqiR | [Rxx trR] x y z]; last by split=> [|/trR->].
+by split=> [x | x y Rxy z]; [rewrite (eqiR x x x) | rewrite (eqiR x y z)].
+Qed.
+
+End RelationProperties.
+
+Lemma rev_trans T (R : rel T) : transitive R -> transitive (fun x y => R y x).
+Proof. by move=> trR x y z Ryx Rzy; apply: trR Rzy Ryx. Qed.
+
+(* Property localization *)
+
+Local Notation "{ 'all1' P }" := (forall x, P x : Prop) (at level 0).
+Local Notation "{ 'all2' P }" := (forall x y, P x y : Prop) (at level 0).
+Local Notation "{ 'all3' P }" := (forall x y z, P x y z: Prop) (at level 0).
+Local Notation ph := (phantom _).
+
+Section LocalProperties.
+
+Variables T1 T2 T3 : Type.
+
+Variables (d1 : mem_pred T1) (d2 : mem_pred T2) (d3 : mem_pred T3).
+Local Notation ph := (phantom Prop).
+
+Definition prop_for (x : T1) P & ph {all1 P} := P x.
+
+Lemma forE x P phP : @prop_for x P phP = P x. Proof. by []. Qed.
+
+Definition prop_in1 P & ph {all1 P} :=
+ forall x, in_mem x d1 -> P x.
+
+Definition prop_in11 P & ph {all2 P} :=
+ forall x y, in_mem x d1 -> in_mem y d2 -> P x y.
+
+Definition prop_in2 P & ph {all2 P} :=
+ forall x y, in_mem x d1 -> in_mem y d1 -> P x y.
+
+Definition prop_in111 P & ph {all3 P} :=
+ forall x y z, in_mem x d1 -> in_mem y d2 -> in_mem z d3 -> P x y z.
+
+Definition prop_in12 P & ph {all3 P} :=
+ forall x y z, in_mem x d1 -> in_mem y d2 -> in_mem z d2 -> P x y z.
+
+Definition prop_in21 P & ph {all3 P} :=
+ forall x y z, in_mem x d1 -> in_mem y d1 -> in_mem z d2 -> P x y z.
+
+Definition prop_in3 P & ph {all3 P} :=
+ forall x y z, in_mem x d1 -> in_mem y d1 -> in_mem z d1 -> P x y z.
+
+Variable f : T1 -> T2.
+
+Definition prop_on1 Pf P & phantom T3 (Pf f) & ph {all1 P} :=
+ forall x, in_mem (f x) d2 -> P x.
+
+Definition prop_on2 Pf P & phantom T3 (Pf f) & ph {all2 P} :=
+ forall x y, in_mem (f x) d2 -> in_mem (f y) d2 -> P x y.
+
+End LocalProperties.
+
+Definition inPhantom := Phantom Prop.
+Definition onPhantom T P (x : T) := Phantom Prop (P x).
+
+Definition bijective_in aT rT (d : mem_pred aT) (f : aT -> rT) :=
+ exists2 g, prop_in1 d (inPhantom (cancel f g))
+ & prop_on1 d (Phantom _ (cancel g)) (onPhantom (cancel g) f).
+
+Definition bijective_on aT rT (cd : mem_pred rT) (f : aT -> rT) :=
+ exists2 g, prop_on1 cd (Phantom _ (cancel f)) (onPhantom (cancel f) g)
+ & prop_in1 cd (inPhantom (cancel g f)).
+
+Notation "{ 'for' x , P }" :=
+ (prop_for x (inPhantom P))
+ (at level 0, format "{ 'for' x , P }") : type_scope.
+
+Notation "{ 'in' d , P }" :=
+ (prop_in1 (mem d) (inPhantom P))
+ (at level 0, format "{ 'in' d , P }") : type_scope.
+
+Notation "{ 'in' d1 & d2 , P }" :=
+ (prop_in11 (mem d1) (mem d2) (inPhantom P))
+ (at level 0, format "{ 'in' d1 & d2 , P }") : type_scope.
+
+Notation "{ 'in' d & , P }" :=
+ (prop_in2 (mem d) (inPhantom P))
+ (at level 0, format "{ 'in' d & , P }") : type_scope.
+
+Notation "{ 'in' d1 & d2 & d3 , P }" :=
+ (prop_in111 (mem d1) (mem d2) (mem d3) (inPhantom P))
+ (at level 0, format "{ 'in' d1 & d2 & d3 , P }") : type_scope.
+
+Notation "{ 'in' d1 & & d3 , P }" :=
+ (prop_in21 (mem d1) (mem d3) (inPhantom P))
+ (at level 0, format "{ 'in' d1 & & d3 , P }") : type_scope.
+
+Notation "{ 'in' d1 & d2 & , P }" :=
+ (prop_in12 (mem d1) (mem d2) (inPhantom P))
+ (at level 0, format "{ 'in' d1 & d2 & , P }") : type_scope.
+
+Notation "{ 'in' d & & , P }" :=
+ (prop_in3 (mem d) (inPhantom P))
+ (at level 0, format "{ 'in' d & & , P }") : type_scope.
+
+Notation "{ 'on' cd , P }" :=
+ (prop_on1 (mem cd) (inPhantom P) (inPhantom P))
+ (at level 0, format "{ 'on' cd , P }") : type_scope.
+
+Notation "{ 'on' cd & , P }" :=
+ (prop_on2 (mem cd) (inPhantom P) (inPhantom P))
+ (at level 0, format "{ 'on' cd & , P }") : type_scope.
+
+Local Arguments onPhantom {_%type_scope} _ _.
+
+Notation "{ 'on' cd , P & g }" :=
+ (prop_on1 (mem cd) (Phantom (_ -> Prop) P) (onPhantom P g))
+ (at level 0, format "{ 'on' cd , P & g }") : type_scope.
+
+Notation "{ 'in' d , 'bijective' f }" := (bijective_in (mem d) f)
+ (at level 0, f at level 8,
+ format "{ 'in' d , 'bijective' f }") : type_scope.
+
+Notation "{ 'on' cd , 'bijective' f }" := (bijective_on (mem cd) f)
+ (at level 0, f at level 8,
+ format "{ 'on' cd , 'bijective' f }") : type_scope.
+
+(* Weakening and monotonicity lemmas for localized predicates. *)
+(* Note that using these lemmas in backward reasoning will force expansion of *)
+(* the predicate definition, as Coq needs to expose the quantifier to apply *)
+(* these lemmas. We define a few specialized variants to avoid this for some *)
+(* of the ssrfun predicates. *)
+
+Section LocalGlobal.
+
+Variables T1 T2 T3 : predArgType.
+Variables (D1 : pred T1) (D2 : pred T2) (D3 : pred T3).
+Variables (d1 d1' : mem_pred T1) (d2 d2' : mem_pred T2) (d3 d3' : mem_pred T3).
+Variables (f f' : T1 -> T2) (g : T2 -> T1) (h : T3).
+Variables (P1 : T1 -> Prop) (P2 : T1 -> T2 -> Prop).
+Variable P3 : T1 -> T2 -> T3 -> Prop.
+Variable Q1 : (T1 -> T2) -> T1 -> Prop.
+Variable Q1l : (T1 -> T2) -> T3 -> T1 -> Prop.
+Variable Q2 : (T1 -> T2) -> T1 -> T1 -> Prop.
+
+Hypothesis sub1 : sub_mem d1 d1'.
+Hypothesis sub2 : sub_mem d2 d2'.
+Hypothesis sub3 : sub_mem d3 d3'.
+
+Lemma in1W : {all1 P1} -> {in D1, {all1 P1}}.
+Proof. by move=> ? ?. Qed.
+Lemma in2W : {all2 P2} -> {in D1 & D2, {all2 P2}}.
+Proof. by move=> ? ?. Qed.
+Lemma in3W : {all3 P3} -> {in D1 & D2 & D3, {all3 P3}}.
+Proof. by move=> ? ?. Qed.
+
+Lemma in1T : {in T1, {all1 P1}} -> {all1 P1}.
+Proof. by move=> ? ?; auto. Qed.
+Lemma in2T : {in T1 & T2, {all2 P2}} -> {all2 P2}.
+Proof. by move=> ? ?; auto. Qed.
+Lemma in3T : {in T1 & T2 & T3, {all3 P3}} -> {all3 P3}.
+Proof. by move=> ? ?; auto. Qed.
+
+Lemma sub_in1 (Ph : ph {all1 P1}) : prop_in1 d1' Ph -> prop_in1 d1 Ph.
+Proof. by move=> allP x /sub1; apply: allP. Qed.
+
+Lemma sub_in11 (Ph : ph {all2 P2}) : prop_in11 d1' d2' Ph -> prop_in11 d1 d2 Ph.
+Proof. by move=> allP x1 x2 /sub1 d1x1 /sub2; apply: allP. Qed.
+
+Lemma sub_in111 (Ph : ph {all3 P3}) :
+ prop_in111 d1' d2' d3' Ph -> prop_in111 d1 d2 d3 Ph.
+Proof. by move=> allP x1 x2 x3 /sub1 d1x1 /sub2 d2x2 /sub3; apply: allP. Qed.
+
+Let allQ1 f'' := {all1 Q1 f''}.
+Let allQ1l f'' h' := {all1 Q1l f'' h'}.
+Let allQ2 f'' := {all2 Q2 f''}.
+
+Lemma on1W : allQ1 f -> {on D2, allQ1 f}. Proof. by move=> ? ?. Qed.
+
+Lemma on1lW : allQ1l f h -> {on D2, allQ1l f & h}. Proof. by move=> ? ?. Qed.
+
+Lemma on2W : allQ2 f -> {on D2 &, allQ2 f}. Proof. by move=> ? ?. Qed.
+
+Lemma on1T : {on T2, allQ1 f} -> allQ1 f. Proof. by move=> ? ?; auto. Qed.
+
+Lemma on1lT : {on T2, allQ1l f & h} -> allQ1l f h.
+Proof. by move=> ? ?; auto. Qed.
+
+Lemma on2T : {on T2 &, allQ2 f} -> allQ2 f.
+Proof. by move=> ? ?; auto. Qed.
+
+Lemma subon1 (Phf : ph (allQ1 f)) (Ph : ph (allQ1 f)) :
+ prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph.
+Proof. by move=> allQ x /sub2; apply: allQ. Qed.
+
+Lemma subon1l (Phf : ph (allQ1l f)) (Ph : ph (allQ1l f h)) :
+ prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph.
+Proof. by move=> allQ x /sub2; apply: allQ. Qed.
+
+Lemma subon2 (Phf : ph (allQ2 f)) (Ph : ph (allQ2 f)) :
+ prop_on2 d2' Phf Ph -> prop_on2 d2 Phf Ph.
+Proof. by move=> allQ x y /sub2=> d2fx /sub2; apply: allQ. Qed.
+
+Lemma can_in_inj : {in D1, cancel f g} -> {in D1 &, injective f}.
+Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed.
+
+Lemma canLR_in x y : {in D1, cancel f g} -> y \in D1 -> x = f y -> g x = y.
+Proof. by move=> fK D1y ->; rewrite fK. Qed.
+
+Lemma canRL_in x y : {in D1, cancel f g} -> x \in D1 -> f x = y -> x = g y.
+Proof. by move=> fK D1x <-; rewrite fK. Qed.
+
+Lemma on_can_inj : {on D2, cancel f & g} -> {on D2 &, injective f}.
+Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed.
+
+Lemma canLR_on x y : {on D2, cancel f & g} -> f y \in D2 -> x = f y -> g x = y.
+Proof. by move=> fK D2fy ->; rewrite fK. Qed.
+
+Lemma canRL_on x y : {on D2, cancel f & g} -> f x \in D2 -> f x = y -> x = g y.
+Proof. by move=> fK D2fx <-; rewrite fK. Qed.
+
+Lemma inW_bij : bijective f -> {in D1, bijective f}.
+Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed.
+
+Lemma onW_bij : bijective f -> {on D2, bijective f}.
+Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed.
+
+Lemma inT_bij : {in T1, bijective f} -> bijective f.
+Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed.
+
+Lemma onT_bij : {on T2, bijective f} -> bijective f.
+Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed.
+
+Lemma sub_in_bij (D1' : pred T1) :
+ {subset D1 <= D1'} -> {in D1', bijective f} -> {in D1, bijective f}.
+Proof.
+by move=> subD [g' fK g'K]; exists g' => x; move/subD; [apply: fK | apply: g'K].
+Qed.
+
+Lemma subon_bij (D2' : pred T2) :
+ {subset D2 <= D2'} -> {on D2', bijective f} -> {on D2, bijective f}.
+Proof.
+by move=> subD [g' fK g'K]; exists g' => x; move/subD; [apply: fK | apply: g'K].
+Qed.
+
+End LocalGlobal.
+
+Lemma sub_in2 T d d' (P : T -> T -> Prop) :
+ sub_mem d d' -> forall Ph : ph {all2 P}, prop_in2 d' Ph -> prop_in2 d Ph.
+Proof. by move=> /= sub_dd'; apply: sub_in11. Qed.
+
+Lemma sub_in3 T d d' (P : T -> T -> T -> Prop) :
+ sub_mem d d' -> forall Ph : ph {all3 P}, prop_in3 d' Ph -> prop_in3 d Ph.
+Proof. by move=> /= sub_dd'; apply: sub_in111. Qed.
+
+Lemma sub_in12 T1 T d1 d1' d d' (P : T1 -> T -> T -> Prop) :
+ sub_mem d1 d1' -> sub_mem d d' ->
+ forall Ph : ph {all3 P}, prop_in12 d1' d' Ph -> prop_in12 d1 d Ph.
+Proof. by move=> /= sub1 sub; apply: sub_in111. Qed.
+
+Lemma sub_in21 T T3 d d' d3 d3' (P : T -> T -> T3 -> Prop) :
+ sub_mem d d' -> sub_mem d3 d3' ->
+ forall Ph : ph {all3 P}, prop_in21 d' d3' Ph -> prop_in21 d d3 Ph.
+Proof. by move=> /= sub sub3; apply: sub_in111. Qed.
+
+Lemma equivalence_relP_in T (R : rel T) (A : pred T) :
+ {in A & &, equivalence_rel R}
+ <-> {in A, reflexive R} /\ {in A &, forall x y, R x y -> {in A, R x =1 R y}}.
+Proof.
+split=> [eqiR | [Rxx trR] x y z *]; last by split=> [|/trR-> //]; apply: Rxx.
+by split=> [x Ax|x y Ax Ay Rxy z Az]; [rewrite (eqiR x x) | rewrite (eqiR x y)].
+Qed.
+
+Section MonoHomoMorphismTheory.
+
+Variables (aT rT sT : Type) (f : aT -> rT) (g : rT -> aT).
+Variables (aP : pred aT) (rP : pred rT) (aR : rel aT) (rR : rel rT).
+
+Lemma monoW : {mono f : x / aP x >-> rP x} -> {homo f : x / aP x >-> rP x}.
+Proof. by move=> hf x ax; rewrite hf. Qed.
+
+Lemma mono2W :
+ {mono f : x y / aR x y >-> rR x y} -> {homo f : x y / aR x y >-> rR x y}.
+Proof. by move=> hf x y axy; rewrite hf. Qed.
+
+Hypothesis fgK : cancel g f.
+
+Lemma homoRL :
+ {homo f : x y / aR x y >-> rR x y} -> forall x y, aR (g x) y -> rR x (f y).
+Proof. by move=> Hf x y /Hf; rewrite fgK. Qed.
+
+Lemma homoLR :
+ {homo f : x y / aR x y >-> rR x y} -> forall x y, aR x (g y) -> rR (f x) y.
+Proof. by move=> Hf x y /Hf; rewrite fgK. Qed.
+
+Lemma homo_mono :
+ {homo f : x y / aR x y >-> rR x y} -> {homo g : x y / rR x y >-> aR x y} ->
+ {mono g : x y / rR x y >-> aR x y}.
+Proof.
+move=> mf mg x y; case: (boolP (rR _ _))=> [/mg //|].
+by apply: contraNF=> /mf; rewrite !fgK.
+Qed.
+
+Lemma monoLR :
+ {mono f : x y / aR x y >-> rR x y} -> forall x y, rR (f x) y = aR x (g y).
+Proof. by move=> mf x y; rewrite -{1}[y]fgK mf. Qed.
+
+Lemma monoRL :
+ {mono f : x y / aR x y >-> rR x y} -> forall x y, rR x (f y) = aR (g x) y.
+Proof. by move=> mf x y; rewrite -{1}[x]fgK mf. Qed.
+
+Lemma can_mono :
+ {mono f : x y / aR x y >-> rR x y} -> {mono g : x y / rR x y >-> aR x y}.
+Proof. by move=> mf x y /=; rewrite -mf !fgK. Qed.
+
+End MonoHomoMorphismTheory.
+
+Section MonoHomoMorphismTheory_in.
+
+Variables (aT rT sT : predArgType) (f : aT -> rT) (g : rT -> aT).
+Variable (aD : pred aT).
+Variable (aP : pred aT) (rP : pred rT) (aR : rel aT) (rR : rel rT).
+
+Notation rD := [pred x | g x \in aD].
+
+Lemma monoW_in :
+ {in aD &, {mono f : x y / aR x y >-> rR x y}} ->
+ {in aD &, {homo f : x y / aR x y >-> rR x y}}.
+Proof. by move=> hf x y hx hy axy; rewrite hf. Qed.
+
+Lemma mono2W_in :
+ {in aD, {mono f : x / aP x >-> rP x}} ->
+ {in aD, {homo f : x / aP x >-> rP x}}.
+Proof. by move=> hf x hx ax; rewrite hf. Qed.
+
+Hypothesis fgK_on : {on aD, cancel g & f}.
+
+Lemma homoRL_in :
+ {in aD &, {homo f : x y / aR x y >-> rR x y}} ->
+ {in rD & aD, forall x y, aR (g x) y -> rR x (f y)}.
+Proof. by move=> Hf x y hx hy /Hf; rewrite fgK_on //; apply. Qed.
+
+Lemma homoLR_in :
+ {in aD &, {homo f : x y / aR x y >-> rR x y}} ->
+ {in aD & rD, forall x y, aR x (g y) -> rR (f x) y}.
+Proof. by move=> Hf x y hx hy /Hf; rewrite fgK_on //; apply. Qed.
+
+Lemma homo_mono_in :
+ {in aD &, {homo f : x y / aR x y >-> rR x y}} ->
+ {in rD &, {homo g : x y / rR x y >-> aR x y}} ->
+ {in rD &, {mono g : x y / rR x y >-> aR x y}}.
+Proof.
+move=> mf mg x y hx hy; case: (boolP (rR _ _))=> [/mg //|]; first exact.
+by apply: contraNF=> /mf; rewrite !fgK_on //; apply.
+Qed.
+
+Lemma monoLR_in :
+ {in aD &, {mono f : x y / aR x y >-> rR x y}} ->
+ {in aD & rD, forall x y, rR (f x) y = aR x (g y)}.
+Proof. by move=> mf x y hx hy; rewrite -{1}[y]fgK_on // mf. Qed.
+
+Lemma monoRL_in :
+ {in aD &, {mono f : x y / aR x y >-> rR x y}} ->
+ {in rD & aD, forall x y, rR x (f y) = aR (g x) y}.
+Proof. by move=> mf x y hx hy; rewrite -{1}[x]fgK_on // mf. Qed.
+
+Lemma can_mono_in :
+ {in aD &, {mono f : x y / aR x y >-> rR x y}} ->
+ {in rD &, {mono g : x y / rR x y >-> aR x y}}.
+Proof. by move=> mf x y hx hy /=; rewrite -mf // !fgK_on. Qed.
+
+End MonoHomoMorphismTheory_in.
diff --git a/plugins/ssr/ssrbwd.ml b/plugins/ssr/ssrbwd.ml
new file mode 100644
index 000000000..cc0e86684
--- /dev/null
+++ b/plugins/ssr/ssrbwd.ml
@@ -0,0 +1,126 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Printer
+open Pretyping
+open Globnames
+open Glob_term
+open Tacmach
+
+open Ssrmatching_plugin
+open Ssrmatching
+
+open Ssrast
+open Ssrprinters
+open Ssrcommon
+
+let char_to_kind = function
+ | '(' -> xInParens
+ | '@' -> xWithAt
+ | ' ' -> xNoFlag
+ | 'x' -> xCpattern
+ | _ -> assert false
+
+(** Backward chaining tactics: apply, exact, congr. *)
+
+(** The "apply" tactic *)
+
+let interp_agen ist gl ((goclr, _), (k, gc as c)) (clr, rcs) =
+(* ppdebug(lazy(str"sigma@interp_agen=" ++ pr_evar_map None (project gl))); *)
+ let k = char_to_kind k in
+ let rc = pf_intern_term ist gl c in
+ let rcs' = rc :: rcs in
+ match goclr with
+ | None -> clr, rcs'
+ | Some ghyps ->
+ let clr' = snd (interp_hyps ist gl ghyps) @ clr in
+ if k <> xNoFlag then clr', rcs' else
+ let open CAst in
+ match rc with
+ | { loc; v = GVar id } when not_section_id id -> SsrHyp (Loc.tag ?loc id) :: clr', rcs'
+ | { loc; v = GRef (VarRef id, _) } when not_section_id id ->
+ SsrHyp (Loc.tag ?loc id) :: clr', rcs'
+ | _ -> clr', rcs'
+
+let pf_pr_glob_constr gl = pr_glob_constr_env (pf_env gl)
+
+let interp_agens ist gl gagens =
+ match List.fold_right (interp_agen ist gl) gagens ([], []) with
+ | clr, rlemma :: args ->
+ let n = interp_nbargs ist gl rlemma - List.length args in
+ let rec loop i =
+ if i > n then
+ errorstrm Pp.(str "Cannot apply lemma " ++ pf_pr_glob_constr gl rlemma)
+ else
+ try interp_refine ist gl (mkRApp rlemma (mkRHoles i @ args))
+ with _ -> loop (i + 1) in
+ clr, loop 0
+ | _ -> assert false
+
+let pf_match = pf_apply (fun e s c t -> understand_tcc e s ~expected_type:t c)
+
+let apply_rconstr ?ist t gl =
+(* ppdebug(lazy(str"sigma@apply_rconstr=" ++ pr_evar_map None (project gl))); *)
+ let open CAst in
+ let n = match ist, t with
+ | None, { v = GVar id | GRef (VarRef id,_) } -> pf_nbargs gl (EConstr.mkVar id)
+ | Some ist, _ -> interp_nbargs ist gl t
+ | _ -> anomaly "apply_rconstr without ist and not RVar" in
+ let mkRlemma i = mkRApp t (mkRHoles i) in
+ let cl = pf_concl gl in
+ let rec loop i =
+ if i > n then
+ errorstrm Pp.(str"Cannot apply lemma "++pf_pr_glob_constr gl t)
+ else try pf_match gl (mkRlemma i) (OfType cl) with _ -> loop (i + 1) in
+ refine_with (loop 0) gl
+
+let mkRAppView ist gl rv gv =
+ let nb_view_imps = interp_view_nbimps ist gl rv in
+ mkRApp rv (mkRHoles (abs nb_view_imps))
+
+let prof_apply_interp_with = mk_profiler "ssrapplytac.interp_with";;
+
+let refine_interp_apply_view i ist gl gv =
+ let pair i = List.map (fun x -> i, x) in
+ let rv = pf_intern_term ist gl gv in
+ let v = mkRAppView ist gl rv gv in
+ let interp_with (i, hint) =
+ interp_refine ist gl (mkRApp hint (v :: mkRHoles i)) in
+ let interp_with x = prof_apply_interp_with.profile interp_with x in
+ let rec loop = function
+ | [] -> (try apply_rconstr ~ist rv gl with _ -> view_error "apply" gv)
+ | h :: hs -> (try refine_with (snd (interp_with h)) gl with _ -> loop hs) in
+ loop (pair i Ssrview.viewtab.(i) @
+ if i = 2 then pair 1 Ssrview.viewtab.(1) else [])
+
+let apply_top_tac gl =
+ Tacticals.tclTHENLIST [introid top_id; apply_rconstr (mkRVar top_id); Proofview.V82.of_tactic (Tactics.clear [top_id])] gl
+
+let inner_ssrapplytac gviews ggenl gclr ist gl =
+ let _, clr = interp_hyps ist gl gclr in
+ let vtac gv i gl' = refine_interp_apply_view i ist gl' gv in
+ let ggenl, tclGENTAC =
+ if gviews <> [] && ggenl <> [] then
+ let ggenl= List.map (fun (x,g) -> x, cpattern_of_term g) (List.hd ggenl) in
+ [], Tacticals.tclTHEN (genstac (ggenl,[]) ist)
+ else ggenl, Tacticals.tclTHEN Tacticals.tclIDTAC in
+ tclGENTAC (fun gl ->
+ match gviews, ggenl with
+ | v :: tl, [] ->
+ let dbl = if List.length tl = 1 then 2 else 1 in
+ Tacticals.tclTHEN
+ (List.fold_left (fun acc v -> Tacticals.tclTHENLAST acc (vtac v dbl)) (vtac v 1) tl)
+ (cleartac clr) gl
+ | [], [agens] ->
+ let clr', (sigma, lemma) = interp_agens ist gl agens in
+ let gl = pf_merge_uc_of sigma gl in
+ Tacticals.tclTHENLIST [cleartac clr; refine_with ~beta:true lemma; cleartac clr'] gl
+ | _, _ -> Tacticals.tclTHEN apply_top_tac (cleartac clr) gl) gl
+
diff --git a/plugins/ssr/ssrbwd.mli b/plugins/ssr/ssrbwd.mli
new file mode 100644
index 000000000..8bf785a21
--- /dev/null
+++ b/plugins/ssr/ssrbwd.mli
@@ -0,0 +1,20 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+val apply_top_tac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val inner_ssrapplytac :
+ Ssrast.ssrterm list ->
+ ((Ssrast.ssrhyps option * Ssrmatching_plugin.Ssrmatching.occ) *
+ (Ssrast.ssrtermkind * Tacexpr.glob_constr_and_expr))
+ list list ->
+ Ssrast.ssrhyps ->
+ Ssrast.ist ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
new file mode 100644
index 000000000..e90be92cf
--- /dev/null
+++ b/plugins/ssr/ssrcommon.ml
@@ -0,0 +1,1297 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Util
+open Names
+open Evd
+open Constr
+open Termops
+open Printer
+open Locusops
+
+open Ltac_plugin
+open Tacmach
+open Refiner
+open Libnames
+open Ssrmatching_plugin
+open Ssrmatching
+open Ssrast
+open Ssrprinters
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
+(* 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 () ;;
+
+let errorstrm x = CErrors.user_err ~hdr:"ssreflect" x
+
+let allocc = Some(false,[])
+
+(** Bound assumption argument *)
+
+(* The Ltac API does have a type for assumptions but it is level-dependent *)
+(* and therefore impractical to use for complex arguments, so we substitute *)
+(* our own to have a uniform representation. Also, we refuse to intern *)
+(* idents that match global/section constants, since this would lead to *)
+(* fragile Ltac scripts. *)
+
+let hyp_id (SsrHyp (_, id)) = id
+
+let hyp_err ?loc msg id =
+ CErrors.user_err ?loc ~hdr:"ssrhyp" Pp.(str msg ++ Id.print id)
+
+let not_section_id id = not (Termops.is_section_variable id)
+
+let hyps_ids = List.map hyp_id
+
+let rec check_hyps_uniq ids = function
+ | SsrHyp (loc, id) :: _ when List.mem id ids ->
+ hyp_err ?loc "Duplicate assumption " id
+ | SsrHyp (_, id) :: hyps -> check_hyps_uniq (id :: ids) hyps
+ | [] -> ()
+
+let check_hyp_exists hyps (SsrHyp(_, id)) =
+ try ignore(Context.Named.lookup id hyps)
+ with Not_found -> errorstrm Pp.(str"No assumption is named " ++ Id.print id)
+
+let test_hypname_exists hyps id =
+ try ignore(Context.Named.lookup id hyps); true
+ with Not_found -> false
+
+let hoik f = function Hyp x -> f x | Id x -> f x
+let hoi_id = hoik hyp_id
+
+let mk_hint tac = false, [Some tac]
+let mk_orhint tacs = true, tacs
+let nullhint = true, []
+let nohint = false, []
+
+type 'a tac_a = (goal * 'a) sigma -> (goal * 'a) list sigma
+
+let push_ctx a gl = re_sig (sig_it gl, a) (project gl)
+let push_ctxs a gl =
+ re_sig (List.map (fun x -> x,a) (sig_it gl)) (project gl)
+let pull_ctx gl = let g, a = sig_it gl in re_sig g (project gl), a
+let pull_ctxs gl = let g, a = List.split (sig_it gl) in re_sig g (project gl), a
+
+let with_ctx f gl =
+ let gl, ctx = pull_ctx gl in
+ let rc, ctx = f ctx in
+ rc, push_ctx ctx gl
+let without_ctx f gl =
+ let gl, _ctx = pull_ctx gl in
+ f gl
+let tac_ctx t gl =
+ let gl, a = pull_ctx gl in
+ let gl = t gl in
+ push_ctxs a gl
+
+let tclTHEN_ia t1 t2 gl =
+ let gal = t1 gl in
+ let goals, sigma = sig_it gal, project gal in
+ let _, opened, sigma =
+ List.fold_left (fun (i,opened,sigma) g ->
+ let gl = t2 i (re_sig g sigma) in
+ i+1, sig_it gl :: opened, project gl)
+ (1,[],sigma) goals in
+ re_sig (List.flatten (List.rev opened)) sigma
+
+let tclTHEN_a t1 t2 gl = tclTHEN_ia t1 (fun _ -> t2) gl
+
+let tclTHENS_a t1 tl gl = tclTHEN_ia t1
+ (fun i -> List.nth tl (i-1)) gl
+
+let rec tclTHENLIST_a = function
+ | [] -> tac_ctx tclIDTAC
+ | t1::tacl -> tclTHEN_a t1 (tclTHENLIST_a tacl)
+
+(* like tclTHEN_i but passes to the tac "i of n" and not just i *)
+let tclTHEN_i_max tac taci gl =
+ let maxi = ref 0 in
+ tclTHEN_ia (tclTHEN_ia tac (fun i -> maxi := max i !maxi; tac_ctx tclIDTAC))
+ (fun i gl -> taci i !maxi gl) gl
+
+let tac_on_all gl tac =
+ let goals = sig_it gl in
+ let opened, sigma =
+ List.fold_left (fun (opened,sigma) g ->
+ let gl = tac (re_sig g sigma) in
+ sig_it gl :: opened, project gl)
+ ([],project gl) goals in
+ re_sig (List.flatten (List.rev opened)) sigma
+
+(* Used to thread data between intro patterns at run time *)
+type tac_ctx = {
+ tmp_ids : (Id.t * name ref) list;
+ wild_ids : Id.t list;
+ delayed_clears : Id.t list;
+}
+
+let new_ctx () =
+ { tmp_ids = []; wild_ids = []; delayed_clears = [] }
+
+let with_fresh_ctx t gl =
+ let gl = push_ctx (new_ctx()) gl in
+ let gl = t gl in
+ fst (pull_ctxs gl)
+
+open Genarg
+open Stdarg
+open Pp
+
+let errorstrm x = CErrors.user_err ~hdr:"ssreflect" x
+let anomaly s = CErrors.anomaly (str s)
+
+(* Tentative patch from util.ml *)
+
+let array_fold_right_from n f v a =
+ let rec fold n =
+ if n >= Array.length v then a else f v.(n) (fold (succ n))
+ in
+ fold n
+
+let array_app_tl v l =
+ if Array.length v = 0 then invalid_arg "array_app_tl";
+ array_fold_right_from 1 (fun e l -> e::l) v l
+
+let array_list_of_tl v =
+ if Array.length v = 0 then invalid_arg "array_list_of_tl";
+ array_fold_right_from 1 (fun e l -> e::l) v []
+
+(* end patch *)
+
+
+(** Constructors for rawconstr *)
+open Glob_term
+open Globnames
+open Misctypes
+open Decl_kinds
+
+let mkRHole = CAst.make @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None)
+
+let rec mkRHoles n = if n > 0 then mkRHole :: mkRHoles (n - 1) else []
+let rec isRHoles = function { CAst.v = GHole _ } :: cl -> isRHoles cl | cl -> cl = []
+let mkRApp f args = if args = [] then f else CAst.make @@ GApp (f, args)
+let mkRVar id = CAst.make @@ GRef (VarRef id,None)
+let mkRltacVar id = CAst.make @@ GVar (id)
+let mkRCast rc rt = CAst.make @@ GCast (rc, CastConv rt)
+let mkRType = CAst.make @@ GSort (GType [])
+let mkRProp = CAst.make @@ GSort (GProp)
+let mkRArrow rt1 rt2 = CAst.make @@ GProd (Anonymous, Explicit, rt1, rt2)
+let mkRConstruct c = CAst.make @@ GRef (ConstructRef c,None)
+let mkRInd mind = CAst.make @@ GRef (IndRef mind,None)
+let mkRLambda n s t = CAst.make @@ GLambda (n, Explicit, s, t)
+
+let rec mkRnat n =
+ if n <= 0 then CAst.make @@ GRef (Coqlib.glob_O, None) else
+ mkRApp (CAst.make @@ GRef (Coqlib.glob_S, None)) [mkRnat (n - 1)]
+
+let glob_constr ist genv = function
+ | _, Some ce ->
+ let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.Tacinterp.lfun Id.Set.empty in
+ let ltacvars = {
+ Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in
+ Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv ce
+ | rc, None -> rc
+
+let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c
+let intern_term ist env (_, c) = glob_constr ist env c
+
+(* Estimate a bound on the number of arguments of a raw constr. *)
+(* This is not perfect, because the unifier may fail to *)
+(* typecheck the partial application, so we use a minimum of 5. *)
+(* Also, we don't handle delayed or iterated coercions to *)
+(* FUNCLASS, which is probably just as well since these can *)
+(* lead to infinite arities. *)
+
+let splay_open_constr gl (sigma, c) =
+ let env = pf_env gl in let t = Retyping.get_type_of env sigma c in
+ Reductionops.splay_prod env sigma t
+
+let isAppInd gl c =
+ try ignore (pf_reduce_to_atomic_ind gl c); true with _ -> false
+
+(** Generic argument-based globbing/typing utilities *)
+
+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
+ } in
+ let kind = Pretyping.OfType (pf_concl gl) in
+ let flags = {
+ Pretyping.use_typeclasses = true;
+ solve_unification_constraints = true;
+ use_hook = None;
+ fail_evar = false;
+ expand_evars = true }
+ in
+ let sigma, c = Pretyping.understand_ltac flags (pf_env gl) (project gl) vars kind rc in
+(* ppdebug(lazy(str"sigma@interp_refine=" ++ pr_evar_map None sigma)); *)
+ ppdebug(lazy(str"c@interp_refine=" ++ Printer.pr_econstr c));
+ (sigma, (sigma, c))
+
+
+let interp_open_constr ist gl gc =
+ let (sigma, (c, _)) = Tacinterp.interp_open_constr_with_bindings ist (pf_env gl) (project gl) (gc, Misctypes.NoBindings) in
+ (project gl, (sigma, c))
+
+let interp_term ist gl (_, c) = snd (interp_open_constr ist gl c)
+
+let of_ftactic ftac gl =
+ let r = ref None in
+ let tac = Ftactic.run ftac (fun ans -> r := Some ans; Proofview.tclUNIT ()) in
+ let tac = Proofview.V82.of_tactic tac in
+ let { sigma = sigma } = tac gl in
+ let ans = match !r with
+ | None -> assert false (** If the tactic failed we should not reach this point *)
+ | Some ans -> ans
+ in
+ (sigma, ans)
+
+let interp_wit wit ist gl x =
+ let globarg = in_gen (glbwit wit) x in
+ let arg = Tacinterp.interp_genarg ist globarg in
+ let (sigma, arg) = of_ftactic arg gl in
+ sigma, Tacinterp.Value.cast (topwit wit) arg
+
+let interp_hyp ist gl (SsrHyp (loc, id)) =
+ let s, id' = interp_wit wit_var ist gl (loc, id) in
+ if not_section_id id' then s, SsrHyp (loc, id') else
+ hyp_err ?loc "Can't clear section hypothesis " id'
+
+let interp_hyps ist gl ghyps =
+ let hyps = List.map snd (List.map (interp_hyp ist gl) ghyps) in
+ check_hyps_uniq [] hyps; Tacmach.project gl, hyps
+
+let mk_term k c = k, (mkRHole, Some c)
+let mk_lterm c = mk_term xNoFlag c
+
+let interp_view_nbimps ist gl rc =
+ try
+ let sigma, t = interp_open_constr ist gl (rc, None) in
+ let si = sig_it gl in
+ let gl = re_sig si sigma in
+ let pl, c = splay_open_constr gl t in
+ if isAppInd gl c then List.length pl else (-(List.length pl))
+ with _ -> 0
+
+let nbargs_open_constr gl oc =
+ let pl, _ = splay_open_constr gl oc in List.length pl
+
+let interp_nbargs ist gl rc =
+ try
+ let rc6 = mkRApp rc (mkRHoles 6) in
+ let sigma, t = interp_open_constr ist gl (rc6, None) in
+ let si = sig_it gl in
+ let gl = re_sig si sigma in
+ 6 + nbargs_open_constr gl t
+ with _ -> 5
+
+let pf_nbargs gl c = nbargs_open_constr gl (project gl, c)
+
+let internal_names = ref []
+let add_internal_name pt = internal_names := pt :: !internal_names
+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)
+let new_tmp_id ctx =
+ let id = mk_tmp_id (1 + List.length ctx.tmp_ids) in
+ let orig = ref Anonymous in
+ (id, orig), { ctx with tmp_ids = (id, orig) :: ctx.tmp_ids }
+;;
+
+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'
+
+let same_prefix s t n =
+ let rec loop i = i = n || s.[i] = t.[i] && loop (i + 1) in loop 0
+
+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 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
+
+let evar_tag = "_evar_"
+let _ = add_internal_name (is_tagged evar_tag)
+let mk_evar_name n = Name (mk_tagged_id evar_tag n)
+
+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)
+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
+ n < m + m' + 2 && same_prefix s wildcard_tag m &&
+ String.sub s (n - m') m' = wildcard_post &&
+ skip_digits s m = n - m' - 2
+let _ = add_internal_name has_wildcard_tag
+
+let new_wild_id ctx =
+ let i = 1 + List.length ctx.wild_ids in
+ let id = mk_wildcard_id i in
+ id, { ctx with wild_ids = id :: ctx.wild_ids }
+
+let discharged_tag = "_discharged_"
+let mk_discharged_id id =
+ id_of_string (Printf.sprintf "%s%s_" discharged_tag (string_of_id 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 max_suffix m (t, j0 as tj0) id =
+ let s = string_of_id 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 =
+ if i < i0 && s.[i] = '0' then loop (i + 1) else
+ if (if i < i0 then skip_digits s i = n else le_s_t i) then s, i else tj0
+ and le_s_t i =
+ let ds = s.[i] and dt = t.[i + dn] in
+ if ds = dt then i = n || le_s_t (i + 1) else
+ dt < ds && skip_digits s i = n in
+ loop m
+
+let mk_anon_id t gl =
+ let m, si0, id0 =
+ let s = ref (Printf.sprintf "_%s_" t) in
+ if is_internal_name !s then s := "_" ^ !s;
+ let n = String.length !s - 1 in
+ 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 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
+ let open Bytes in
+ let s = of_string s in
+ let n = length s - 1 in
+ let rec loop i =
+ if get s i = '9' then (set s i '0'; loop (i - 1)) else
+ if i < m then (set s n '0'; set s m '1'; cat s (of_string "_")) else
+ (set s i (Char.chr (Char.code (get s i) + 1)); s) in
+ Id.of_bytes (loop (n - 1))
+
+let convert_concl_no_check t = Tactics.convert_concl_no_check t Term.DEFAULTcast
+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) ->
+ Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkProd (!orig_name_ref,src,tgt))) gl
+ | _ -> CErrors.anomaly (str "gentac creates no product")
+
+(* Reduction that preserves the Prod/Let spine of the "in" tactical. *)
+
+let inc_safe n = if n = 0 then n else n + 1
+let rec safe_depth s c = match EConstr.kind s c with
+| LetIn (Name x, _, _, c') when is_discharged_id x -> safe_depth s c' + 1
+| LetIn (_, _, _, c') | Prod (_, _, c') -> inc_safe (safe_depth s c')
+| _ -> 0
+
+let red_safe (r : Reductionops.reduction_function) e s c0 =
+ let rec red_to e c n = match EConstr.kind s c with
+ | Prod (x, t, c') when n > 0 ->
+ let t' = r e s t in let e' = EConstr.push_rel (RelDecl.LocalAssum (x, t')) e in
+ EConstr.mkProd (x, t', red_to e' c' (n - 1))
+ | LetIn (x, b, t, c') when n > 0 ->
+ let t' = r e s t in let e' = EConstr.push_rel (RelDecl.LocalAssum (x, t')) e in
+ EConstr.mkLetIn (x, r e s b, t', red_to e' c' (n - 1))
+ | _ -> r e s c in
+ red_to e c0 (safe_depth s c0)
+
+let is_id_constr sigma c = match EConstr.kind sigma c with
+ | Lambda(_,_,c) when EConstr.isRel sigma c -> 1 = EConstr.destRel sigma c
+ | _ -> false
+
+let red_product_skip_id env sigma c = match EConstr.kind sigma c with
+ | App(hd,args) when Array.length args = 1 && is_id_constr sigma hd -> args.(0)
+ | _ -> try Tacred.red_product env sigma c with _ -> c
+
+let ssrevaltac ist gtac =
+ Proofview.V82.of_tactic (Tacinterp.tactic_of_value ist gtac)
+(** Open term to lambda-term coercion {{{ ************************************)
+
+(* This operation takes a goal gl and an open term (sigma, t), and *)
+(* returns a term t' where all the new evars in sigma are abstracted *)
+(* with the mkAbs argument, i.e., for mkAbs = mkLambda then there is *)
+(* some duplicate-free array args of evars of sigma such that the *)
+(* term mkApp (t', args) is convertible to t. *)
+(* This makes a useful shorthand for local definitions in proofs, *)
+(* i.e., pose succ := _ + 1 means pose succ := fun n : nat => n + 1, *)
+(* and, in context of the the 4CT library, pose mid := maps id means *)
+(* pose mid := fun d : detaSet => @maps d d (@id (datum d)) *)
+(* Note that this facility does not extend to set, which tries *)
+(* instead to fill holes by matching a goal subterm. *)
+(* The argument to "have" et al. uses product abstraction, e.g. *)
+(* have Hmid: forall s, (maps id s) = s. *)
+(* stands for *)
+(* have Hmid: forall (d : dataSet) (s : seq d), (maps id s) = s. *)
+(* We also use this feature for rewrite rules, so that, e.g., *)
+(* rewrite: (plus_assoc _ 3). *)
+(* will execute as *)
+(* rewrite (fun n => plus_assoc n 3) *)
+(* i.e., it will rewrite some subterm .. + (3 + ..) to .. + 3 + ... *)
+(* The convention is also used for the argument of the congr tactic, *)
+(* e.g., congr (x + _ * 1). *)
+
+(* Replace new evars with lambda variables, retaining local dependencies *)
+(* but stripping global ones. We use the variable names to encode the *)
+(* the number of dependencies, so that the transformation is reversible. *)
+
+open Term
+let env_size env = List.length (Environ.named_context env)
+
+let pf_concl gl = EConstr.Unsafe.to_constr (pf_concl gl)
+let pf_get_hyp gl x = EConstr.Unsafe.to_named_decl (pf_get_hyp gl x)
+
+let pf_e_type_of gl t =
+ let sigma, env, it = project gl, pf_env gl, sig_it gl in
+ let sigma, ty = Typing.type_of env sigma t in
+ re_sig it sigma, ty
+
+let nf_evar sigma t =
+ EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr t))
+
+let pf_abs_evars2 gl rigid (sigma, c0) =
+ let c0 = EConstr.Unsafe.to_constr c0 in
+ let sigma0, ucst = project gl, Evd.evar_universe_context sigma in
+ let nenv = env_size (pf_env gl) in
+ let abs_evar n k =
+ let evi = Evd.find sigma k in
+ let dc = CList.firstn n (evar_filtered_context evi) in
+ let abs_dc c = function
+ | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c)
+ | NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
+ let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in
+ nf_evar sigma t in
+ let rec put evlist c = match kind_of_term c with
+ | Evar (k, a) ->
+ if List.mem_assoc k evlist || Evd.mem sigma0 k || List.mem k rigid then evlist else
+ let n = max 0 (Array.length a - nenv) in
+ let t = abs_evar n k in (k, (n, t)) :: put evlist t
+ | _ -> fold_constr put evlist c in
+ let evlist = put [] c0 in
+ if evlist = [] then 0, EConstr.of_constr c0,[], ucst else
+ let rec lookup k i = function
+ | [] -> 0, 0
+ | (k', (n, _)) :: evl -> if k = k' then i, n else lookup k (i + 1) evl in
+ let rec get i c = match kind_of_term c with
+ | Evar (ev, a) ->
+ let j, n = lookup ev i evlist in
+ if j = 0 then map_constr (get i) c else if n = 0 then mkRel j else
+ mkApp (mkRel j, Array.init n (fun k -> get i a.(n - 1 - k)))
+ | _ -> map_constr_with_binders ((+) 1) get i c in
+ let rec loop c i = function
+ | (_, (n, t)) :: evl ->
+ loop (mkLambda (mk_evar_name n, get (i - 1) t, c)) (i - 1) evl
+ | [] -> c in
+ List.length evlist, EConstr.of_constr (loop (get 1 c0) 1 evlist), List.map fst evlist, ucst
+
+let pf_abs_evars gl t = pf_abs_evars2 gl [] t
+
+
+(* As before but if (?i : T(?j)) and (?j : P : Prop), then the lambda for i
+ * looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all
+ * occurrences of evar_i are replaced by (evar_i evar_j) thanks to "app".
+ *
+ * If P can be solved by ssrautoprop (that defaults to trivial), then
+ * the corresponding lambda looks like (fun evar_i : T(c)) where c is
+ * the solution found by ssrautoprop.
+ *)
+let ssrautoprop_tac = ref (fun gl -> assert false)
+
+(* Thanks to Arnaud Spiwack for this snippet *)
+let call_on_evar tac e s =
+ let { it = gs ; sigma = s } =
+ tac { it = e ; sigma = s; } in
+ gs, s
+
+open Pp
+let pp _ = () (* FIXME *)
+module Intset = Evar.Set
+
+let pf_abs_evars_pirrel gl (sigma, c0) =
+ pp(lazy(str"==PF_ABS_EVARS_PIRREL=="));
+ pp(lazy(str"c0= " ++ Printer.pr_constr c0));
+ let sigma0 = project gl in
+ let c0 = nf_evar sigma0 (nf_evar sigma c0) in
+ let nenv = env_size (pf_env gl) in
+ let abs_evar n k =
+ let evi = Evd.find sigma k in
+ let dc = CList.firstn n (evar_filtered_context evi) in
+ let abs_dc c = function
+ | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c)
+ | NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
+ let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in
+ nf_evar sigma0 (nf_evar sigma t) in
+ let rec put evlist c = match kind_of_term c with
+ | Evar (k, a) ->
+ if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else
+ let n = max 0 (Array.length a - nenv) in
+ let k_ty =
+ Retyping.get_sort_family_of
+ (pf_env gl) sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k))) in
+ let is_prop = k_ty = InProp in
+ let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t
+ | _ -> fold_constr put evlist c in
+ let evlist = put [] c0 in
+ if evlist = [] then 0, c0 else
+ let pr_constr t = Printer.pr_econstr (Reductionops.nf_beta (project gl) (EConstr.of_constr t)) in
+ pp(lazy(str"evlist=" ++ pr_list (fun () -> str";")
+ (fun (k,_) -> str(Evd.string_of_existential k)) evlist));
+ let evplist =
+ let depev = List.fold_left (fun evs (_,(_,t,_)) ->
+ let t = EConstr.of_constr t in
+ Intset.union evs (Evarutil.undefined_evars_of_term sigma t)) Intset.empty evlist in
+ List.filter (fun (i,(_,_,b)) -> b && Intset.mem i depev) evlist in
+ let evlist, evplist, sigma =
+ if evplist = [] then evlist, [], sigma else
+ List.fold_left (fun (ev, evp, sigma) (i, (_,t,_) as p) ->
+ try
+ let ng, sigma = call_on_evar !ssrautoprop_tac i sigma in
+ if (ng <> []) then errorstrm (str "Should we tell the user?");
+ List.filter (fun (j,_) -> j <> i) ev, evp, sigma
+ with _ -> ev, p::evp, sigma) (evlist, [], sigma) (List.rev evplist) in
+ let c0 = nf_evar sigma c0 in
+ let evlist =
+ List.map (fun (x,(y,t,z)) -> x,(y,nf_evar sigma t,z)) evlist in
+ let evplist =
+ List.map (fun (x,(y,t,z)) -> x,(y,nf_evar sigma t,z)) evplist in
+ pp(lazy(str"c0= " ++ pr_constr c0));
+ let rec lookup k i = function
+ | [] -> 0, 0
+ | (k', (n,_,_)) :: evl -> if k = k' then i,n else lookup k (i + 1) evl in
+ let rec get evlist i c = match kind_of_term c with
+ | Evar (ev, a) ->
+ let j, n = lookup ev i evlist in
+ if j = 0 then map_constr (get evlist i) c else if n = 0 then mkRel j else
+ mkApp (mkRel j, Array.init n (fun k -> get evlist i a.(n - 1 - k)))
+ | _ -> map_constr_with_binders ((+) 1) (get evlist) i c in
+ let rec app extra_args i c = match decompose_app c with
+ | hd, args when isRel hd && destRel hd = i ->
+ let j = destRel hd in
+ mkApp (mkRel j, Array.of_list (List.map (Vars.lift (i-1)) extra_args @ args))
+ | _ -> map_constr_with_binders ((+) 1) (app extra_args) i c in
+ 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
+ loopP evlist (mkProd (n, t, c)) (i - 1) evl
+ | [] -> c in
+ let rec loop c i = function
+ | (_, (n, t, _)) :: evl ->
+ let evs = Evarutil.undefined_evars_of_term sigma (EConstr.of_constr t) in
+ let t_evplist = List.filter (fun (k,_) -> Intset.mem k evs) evplist in
+ let t = loopP t_evplist (get t_evplist 1 t) 1 t_evplist in
+ let t = get evlist (i - 1) t in
+ let extra_args =
+ List.map (fun (k,_) -> mkRel (fst (lookup k i evlist)))
+ (List.rev t_evplist) in
+ let c = if extra_args = [] then c else app extra_args 1 c in
+ loop (mkLambda (mk_evar_name n, t, c)) (i - 1) evl
+ | [] -> c in
+ let res = loop (get evlist 1 c0) 1 evlist in
+ pp(lazy(str"res= " ++ pr_constr res));
+ List.length evlist, res
+
+(* Strip all non-essential dependencies from an abstracted term, generating *)
+(* standard names for the abstracted holes. *)
+
+let nb_evar_deps = function
+ | Name id ->
+ let s = string_of_id 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 pfe_type_of gl t =
+ let sigma, ty = pf_type_of gl t in
+ re_sig (sig_it gl) sigma, ty
+let pf_type_of gl t =
+ let sigma, ty = pf_type_of gl (EConstr.of_constr t) in
+ re_sig (sig_it gl) sigma, EConstr.Unsafe.to_constr ty
+
+let pf_abs_cterm gl n c0 =
+ if n <= 0 then c0 else
+ let c0 = EConstr.Unsafe.to_constr c0 in
+ let noargs = [|0|] in
+ let eva = Array.make n noargs in
+ let rec strip i c = match kind_of_term c with
+ | App (f, a) when isRel f ->
+ let j = i - destRel f in
+ if j >= n || eva.(j) = noargs then mkApp (f, Array.map (strip i) a) else
+ let dp = eva.(j) in
+ let nd = Array.length dp - 1 in
+ let mkarg k = strip i a.(if k < nd then dp.(k + 1) - j else k + dp.(0)) in
+ mkApp (f, Array.init (Array.length a - dp.(0)) mkarg)
+ | _ -> map_constr_with_binders ((+) 1) strip i c in
+ let rec strip_ndeps j i c = match kind_of_term c with
+ | Prod (x, t, c1) when i < j ->
+ let dl, c2 = strip_ndeps j (i + 1) c1 in
+ if Vars.noccurn 1 c2 then dl, Vars.lift (-1) c2 else
+ i :: dl, mkProd (x, strip i t, c2)
+ | LetIn (x, b, t, c1) when i < j ->
+ let _, _, c1' = destProd c1 in
+ let dl, c2 = strip_ndeps j (i + 1) c1' in
+ if Vars.noccurn 1 c2 then dl, Vars.lift (-1) c2 else
+ i :: dl, mkLetIn (x, strip i b, strip i t, c2)
+ | _ -> [], strip i c in
+ let rec strip_evars i c = match kind_of_term c with
+ | Lambda (x, t1, c1) when i < n ->
+ let na = nb_evar_deps x in
+ let dl, t2 = strip_ndeps (i + na) i t1 in
+ let na' = List.length dl in
+ eva.(i) <- Array.of_list (na - na' :: dl);
+ let x' =
+ if na' = 0 then Name (pf_type_id gl (EConstr.of_constr t2)) else mk_evar_name na' in
+ mkLambda (x', t2, strip_evars (i + 1) c1)
+(* if noccurn 1 c2 then lift (-1) c2 else
+ mkLambda (Name (pf_type_id gl t2), t2, c2) *)
+ | _ -> strip i c in
+ EConstr.of_constr (strip_evars 0 c0)
+
+(* }}} *)
+
+let pf_merge_uc uc gl =
+ re_sig (sig_it gl) (Evd.merge_universe_context (Refiner.project gl) uc)
+let pf_merge_uc_of sigma gl =
+ let ucst = Evd.evar_universe_context sigma in
+ pf_merge_uc ucst 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))
+ | App (c', _) -> constr_name sigma c'
+ | _ -> Anonymous
+
+let pf_mkprod gl c ?(name=constr_name (project gl) c) cl =
+ let gl, t = pfe_type_of gl c in
+ if name <> Anonymous || EConstr.Vars.noccurn (project gl) 1 cl then gl, EConstr.mkProd (name, t, cl) else
+ gl, EConstr.mkProd (Name (pf_type_id gl t), t, 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 locate_reference qid =
+ Smartlocate.global_of_extended_global (Nametab.locate_extended qid)
+let mkSsrRef name =
+ try locate_reference (ssrqid name) with Not_found ->
+ try locate_reference (ssrtopqid name) with Not_found ->
+ CErrors.user_err (Pp.str "Small scale reflection library not loaded")
+let mkSsrRRef name = (CAst.make @@ GRef (mkSsrRef name,None)), None
+let mkSsrConst name env sigma =
+ EConstr.fresh_global env sigma (mkSsrRef name)
+let pf_mkSsrConst name gl =
+ let sigma, env, it = project gl, pf_env gl, sig_it gl in
+ let (sigma, t) = mkSsrConst name env sigma in
+ t, re_sig it sigma
+let pf_fresh_global name gl =
+ let sigma, env, it = project gl, pf_env gl, sig_it gl in
+ let sigma,t = Evd.fresh_global env sigma name in
+ t, re_sig it sigma
+
+let mkProt t c gl =
+ let prot, gl = pf_mkSsrConst "protect_term" gl in
+ EConstr.mkApp (prot, [|t; c|]), gl
+
+let mkEtaApp c n imin =
+ let open EConstr in
+ if n = 0 then c else
+ let nargs, mkarg =
+ if n < 0 then -n, (fun i -> mkRel (imin + i)) else
+ let imax = imin + n - 1 in n, (fun i -> mkRel (imax - i)) in
+ mkApp (c, Array.init nargs mkarg)
+
+let mkRefl t c gl =
+ let sigma = project gl in
+ let (sigma, refl) = EConstr.fresh_global (pf_env gl) sigma Coqlib.((build_coq_eq_data()).refl) in
+ EConstr.mkApp (refl, [|t; c|]), { gl with sigma }
+
+let discharge_hyp (id', (id, mode)) gl =
+ let cl' = Vars.subst_var id (pf_concl gl) in
+ match pf_get_hyp gl id, mode with
+ | NamedDecl.LocalAssum (_, t), _ | NamedDecl.LocalDef (_, _, t), "(" ->
+ Proofview.V82.of_tactic (Tactics.apply_type (EConstr.of_constr (mkProd (Name id', t, cl')))
+ [EConstr.of_constr (mkVar id)]) gl
+ | NamedDecl.LocalDef (_, v, t), _ ->
+ Proofview.V82.of_tactic
+ (convert_concl (EConstr.of_constr (mkLetIn (Name id', v, t, cl')))) gl
+
+(* wildcard names *)
+let clear_wilds wilds gl =
+ Proofview.V82.of_tactic (Tactics.clear (List.filter (fun id -> List.mem id wilds) (pf_ids_of_hyps gl))) gl
+
+let clear_with_wilds wilds clr0 gl =
+ let extend_clr clr nd =
+ let id = NamedDecl.get_id nd in
+ if List.mem id clr || not (List.mem id wilds) then clr else
+ let vars = Termops.global_vars_set_of_decl (pf_env gl) (project gl) nd in
+ let occurs id' = Idset.mem id' vars in
+ if List.exists occurs clr then id :: clr else clr in
+ Proofview.V82.of_tactic (Tactics.clear (Context.Named.fold_inside extend_clr ~init:clr0 (Tacmach.pf_hyps gl))) gl
+
+let clear_wilds_and_tmp_and_delayed_ids gl =
+ let _, ctx = pull_ctx gl in
+ tac_ctx
+ (tclTHEN
+ (clear_with_wilds ctx.wild_ids ctx.delayed_clears)
+ (clear_wilds (List.map fst ctx.tmp_ids @ ctx.wild_ids))) gl
+
+let rec is_name_in_ipats name = function
+ | IPatClear clr :: tl ->
+ List.exists (function SsrHyp(_,id) -> id = name) clr
+ || is_name_in_ipats name tl
+ | IPatId id :: tl -> id = name || is_name_in_ipats name tl
+ | IPatCase l :: tl -> List.exists (is_name_in_ipats name) l || is_name_in_ipats name tl
+ | _ :: tl -> is_name_in_ipats name tl
+ | [] -> false
+
+let view_error s gv =
+ errorstrm (str ("Cannot " ^ s ^ " view ") ++ pr_term gv)
+
+
+open Locus
+(****************************** tactics ***********************************)
+
+let rewritetac dir c =
+ (* Due to the new optional arg ?tac, application shouldn't be too partial *)
+ Proofview.V82.of_tactic begin
+ Equality.general_rewrite (dir = L2R) AllOccurrences true false c
+ end
+
+(**********************`:********* hooks ************************************)
+
+type name_hint = (int * EConstr.types array) option ref
+
+let pf_abs_ssrterm ?(resolve_typeclasses=false) ist gl t =
+ let sigma, ct as t = interp_term ist gl t in
+ let sigma, _ as t =
+ let env = pf_env gl in
+ if not resolve_typeclasses then t
+ else
+ let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in
+ sigma, Evarutil.nf_evar sigma ct in
+ let n, c, abstracted_away, ucst = pf_abs_evars gl t in
+ List.fold_left Evd.remove sigma abstracted_away, pf_abs_cterm gl n c, ucst, n
+
+let top_id = mk_internal_id "top assumption"
+
+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))
+ with Not_found -> try Nametab.locate_tactic (ssrqid name)
+ with Not_found ->
+ if n = -1 then fail "The ssreflect library was not loaded"
+ else fail ("The tactic "^name^" was not found") in
+ let tacexpr = Loc.tag @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in
+ Proofview.V82.of_tactic (Tacinterp.eval_tactic (Tacexpr.TacArg tacexpr)) gl
+
+let donetac n gl = ssr_n_tac "done" n gl
+
+open Constrexpr
+open Util
+
+(** Constructors for constr_expr *)
+let mkCProp loc = CAst.make ?loc @@ CSort GProp
+let mkCType loc = CAst.make ?loc @@ CSort (GType [])
+let mkCVar ?loc id = CAst.make ?loc @@ CRef (Ident (Loc.tag ?loc id), None)
+let rec mkCHoles ?loc n =
+ if n <= 0 then [] else (CAst.make ?loc @@ CHole (None, IntroAnonymous, None)) :: mkCHoles ?loc (n - 1)
+let mkCHole loc = CAst.make ?loc @@ CHole (None, IntroAnonymous, None)
+let mkCLambda ?loc name ty t = CAst.make ?loc @@
+ CLambdaN ([[loc, name], Default Explicit, ty], t)
+let mkCArrow ?loc ty t = CAst.make ?loc @@
+ CProdN ([[Loc.tag Anonymous], Default Explicit, ty], t)
+let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, CastConv ty)
+
+let rec isCHoles = function { CAst.v = CHole _ } :: cl -> isCHoles cl | cl -> cl = []
+let rec isCxHoles = function ({ CAst.v = CHole _ }, None) :: ch -> isCxHoles ch | _ -> false
+
+let pf_interp_ty ?(resolve_typeclasses=false) ist gl ty =
+ let n_binders = ref 0 in
+ let ty = match ty with
+ | a, (t, None) ->
+ let rec force_type ty = CAst.(map (function
+ | GProd (x, k, s, t) -> incr n_binders; GProd (x, k, s, force_type t)
+ | GLetIn (x, v, oty, t) -> incr n_binders; GLetIn (x, v, oty, force_type t)
+ | _ -> (mkRCast ty mkRType).v)) ty in
+ a, (force_type t, None)
+ | _, (_, Some ty) ->
+ let rec force_type ty = CAst.(map (function
+ | CProdN (abs, t) ->
+ n_binders := !n_binders + List.length (List.flatten (List.map pi1 abs));
+ CProdN (abs, force_type t)
+ | CLetIn (n, v, oty, t) -> incr n_binders; CLetIn (n, v, oty, force_type t)
+ | _ -> (mkCCast ty (mkCType None)).v)) ty in
+ mk_term ' ' (force_type ty) in
+ let strip_cast (sigma, t) =
+ let rec aux t = match EConstr.kind_of_type sigma t with
+ | CastType (t, ty) when !n_binders = 0 && EConstr.isSort sigma ty -> t
+ | ProdType(n,s,t) -> decr n_binders; EConstr.mkProd (n, s, aux t)
+ | LetInType(n,v,ty,t) -> decr n_binders; EConstr.mkLetIn (n, v, ty, aux t)
+ | _ -> anomaly "pf_interp_ty: ssr Type cast deleted by typecheck" in
+ sigma, aux t in
+ let sigma, cty as ty = strip_cast (interp_term ist gl ty) in
+ let ty =
+ let env = pf_env gl in
+ if not resolve_typeclasses then ty
+ else
+ let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in
+ sigma, Evarutil.nf_evar sigma cty in
+ let n, c, _, ucst = pf_abs_evars gl ty in
+ let lam_c = pf_abs_cterm gl n c in
+ let ctx, c = EConstr.decompose_lam_n_assum sigma n lam_c in
+ n, EConstr.it_mkProd_or_LetIn c ctx, lam_c, ucst
+;;
+
+(* TASSI: given (c : ty), generates (c ??? : ty[???/...]) with m evars *)
+exception NotEnoughProducts
+let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_of env sigma c) m
+=
+ let rec loop ty args sigma n =
+ if n = 0 then
+ let args = List.rev args in
+ (if beta then Reductionops.whd_beta sigma else fun x -> x)
+ (EConstr.mkApp (c, Array.of_list (List.map snd args))), ty, args, sigma
+ else match EConstr.kind_of_type sigma ty with
+ | ProdType (_, src, tgt) ->
+ let sigma = create_evar_defs sigma in
+ let (sigma, x) =
+ Evarutil.new_evar env sigma
+ (if bi_types then Reductionops.nf_betaiota sigma src else src) in
+ loop (EConstr.Vars.subst1 x tgt) ((m - n,x) :: args) sigma (n-1)
+ | CastType (t, _) -> loop t args sigma n
+ | LetInType (_, v, _, t) -> loop (EConstr.Vars.subst1 v t) args sigma n
+ | SortType _ -> assert false
+ | AtomicType _ ->
+ let ty = (* FIXME *)
+ (Reductionops.whd_all env sigma) ty in
+ match EConstr.kind_of_type sigma ty with
+ | ProdType _ -> loop ty args sigma n
+ | _ -> raise NotEnoughProducts
+ in
+ loop ty [] sigma m
+
+let pf_saturate ?beta ?bi_types gl c ?ty m =
+ let env, sigma, si = pf_env gl, project gl, sig_it gl in
+ let t, ty, args, sigma = saturate ?beta ?bi_types env sigma c ?ty m in
+ t, ty, args, re_sig si sigma
+
+let pf_partial_solution gl t evl =
+ let sigma, g = project gl, sig_it gl in
+ let sigma = Goal.V82.partial_solution sigma g t in
+ re_sig (List.map (fun x -> (fst (EConstr.destEvar sigma x))) evl) sigma
+
+let dependent_apply_error =
+ try CErrors.user_err (Pp.str "Could not fill dependent hole in \"apply\"")
+ with err -> err
+
+(* TASSI: Sometimes Coq's apply fails. According to my experience it may be
+ * related to goals that are products and with beta redexes. In that case it
+ * guesses the wrong number of implicit arguments for your lemma. What follows
+ * is just like apply, but with a user-provided number n of implicits.
+ *
+ * Refine.refine function that handles type classes and evars but fails to
+ * handle "dependently typed higher order evars".
+ *
+ * Refiner.refiner that does not handle metas with a non ground type but works
+ * with dependently typed higher order metas. *)
+let applyn ~with_evars ?beta ?(with_shelve=false) n t gl =
+ if with_evars then
+ let refine gl =
+ let t, ty, args, gl = pf_saturate ?beta ~bi_types:true gl t n in
+(* pp(lazy(str"sigma@saturate=" ++ pr_evar_map None (project gl))); *)
+ let gl = pf_unify_HO gl ty (Tacmach.pf_concl gl) in
+ let gs = CList.map_filter (fun (_, e) ->
+ if EConstr.isEvar (project gl) e then Some e else None)
+ args in
+ pf_partial_solution gl t gs
+ in
+ Proofview.(V82.of_tactic
+ (tclTHEN (V82.tactic refine)
+ (if with_shelve then shelve_unifiable else tclUNIT ()))) gl
+ else
+ let t, gl = if n = 0 then t, gl else
+ let sigma, si = project gl, sig_it gl in
+ let rec loop sigma bo args = function (* saturate with metas *)
+ | 0 -> EConstr.mkApp (t, Array.of_list (List.rev args)), re_sig si sigma
+ | n -> match EConstr.kind sigma bo with
+ | Lambda (_, ty, bo) ->
+ if not (EConstr.Vars.closed0 sigma ty) then
+ raise dependent_apply_error;
+ let m = Evarutil.new_meta () in
+ loop (meta_declare m (EConstr.Unsafe.to_constr ty) sigma) bo ((EConstr.mkMeta m)::args) (n-1)
+ | _ -> assert false
+ in loop sigma t [] n in
+ pp(lazy(str"Refiner.refiner " ++ Printer.pr_econstr t));
+ Refiner.refiner (Proof_type.Refine (EConstr.Unsafe.to_constr t)) gl
+
+let refine_with ?(first_goes_last=false) ?beta ?(with_evars=true) oc gl =
+ let rec mkRels = function 1 -> [] | n -> mkRel n :: mkRels (n-1) in
+ let uct = Evd.evar_universe_context (fst oc) in
+ let n, oc = pf_abs_evars_pirrel gl (fst oc, EConstr.Unsafe.to_constr (snd oc)) in
+ let gl = pf_unsafe_merge_uc uct gl in
+ let oc = if not first_goes_last || n <= 1 then oc else
+ let l, c = decompose_lam oc in
+ if not (List.for_all_i (fun i (_,t) -> Vars.closedn ~-i t) (1-n) l) then oc else
+ compose_lam (let xs,y = List.chop (n-1) l in y @ xs)
+ (mkApp (compose_lam l c, Array.of_list (mkRel 1 :: mkRels n)))
+ in
+ pp(lazy(str"after: " ++ Printer.pr_constr oc));
+ try applyn ~with_evars ~with_shelve:true ?beta n (EConstr.of_constr oc) gl
+ with e when CErrors.noncritical e -> raise dependent_apply_error
+
+(** Profiling {{{ *************************************************************)
+type profiler = {
+ profile : 'a 'b. ('a -> 'b) -> 'a -> 'b;
+ reset : unit -> unit;
+ print : unit -> unit }
+let profile_now = ref false
+let something_profiled = ref false
+let profilers = ref []
+let add_profiler f = profilers := f :: !profilers;;
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "ssreflect profiling";
+ Goptions.optkey = ["SsrProfiling"];
+ Goptions.optread = (fun _ -> !profile_now);
+ Goptions.optdepr = false;
+ Goptions.optwrite = (fun b ->
+ Ssrmatching.profile b;
+ profile_now := b;
+ if b then List.iter (fun f -> f.reset ()) !profilers;
+ if not b then List.iter (fun f -> f.print ()) !profilers) }
+let () =
+ let prof_total =
+ let init = ref 0.0 in {
+ profile = (fun f x -> assert false);
+ reset = (fun () -> init := Unix.gettimeofday ());
+ print = (fun () -> if !something_profiled then
+ prerr_endline
+ (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f"
+ "total" 0 (Unix.gettimeofday() -. !init) 0.0 0.0)) } in
+ let prof_legenda = {
+ profile = (fun f x -> assert false);
+ reset = (fun () -> ());
+ print = (fun () -> if !something_profiled then begin
+ prerr_endline
+ (Printf.sprintf "!! %39s ---------- --------- --------- ---------"
+ (String.make 39 '-'));
+ prerr_endline
+ (Printf.sprintf "!! %-39s %10s %9s %9s %9s"
+ "function" "#calls" "total" "max" "average") end) } in
+ add_profiler prof_legenda;
+ add_profiler prof_total
+;;
+
+let mk_profiler s =
+ let total, calls, max = ref 0.0, ref 0, ref 0.0 in
+ let reset () = total := 0.0; calls := 0; max := 0.0 in
+ let profile f x =
+ if not !profile_now then f x else
+ let before = Unix.gettimeofday () in
+ try
+ incr calls;
+ let res = f x in
+ let after = Unix.gettimeofday () in
+ let delta = after -. before in
+ total := !total +. delta;
+ if delta > !max then max := delta;
+ res
+ with exc ->
+ let after = Unix.gettimeofday () in
+ let delta = after -. before in
+ total := !total +. delta;
+ if delta > !max then max := delta;
+ raise exc in
+ let print () =
+ if !calls <> 0 then begin
+ something_profiled := true;
+ prerr_endline
+ (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f"
+ s !calls !total !max (!total /. (float_of_int !calls))) end in
+ let prof = { profile = profile; reset = reset; print = print } in
+ add_profiler prof;
+ prof
+;;
+(* }}} *)
+
+(* We wipe out all the keywords generated by the grammar rules we defined. *)
+(* The user is supposed to Require Import ssreflect or Require ssreflect *)
+(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *)
+(* consequence the extended ssreflect grammar. *)
+let () = CLexer.set_keyword_state frozen_lexer ;;
+
+(** Basic tactics *)
+
+let rec fst_prod red tac = Proofview.Goal.nf_enter begin fun gl ->
+ let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
+ match EConstr.kind (Proofview.Goal.sigma gl) concl with
+ | Prod (id,_,tgt) | LetIn(id,_,_,tgt) -> tac id
+ | _ -> if red then Tacticals.New.tclZEROMSG (str"No product even after head-reduction.")
+ else Tacticals.New.tclTHEN Tactics.hnf_in_concl (fst_prod true tac)
+end
+
+let introid ?(orig=ref Anonymous) name = tclTHEN (fun gl ->
+ let g, env = Tacmach.pf_concl gl, pf_env gl in
+ let sigma = project gl in
+ match EConstr.kind sigma g with
+ | App (hd, _) when EConstr.isLambda sigma hd ->
+ Proofview.V82.of_tactic (convert_concl_no_check (Reductionops.whd_beta sigma g)) gl
+ | _ -> tclIDTAC gl)
+ (Proofview.V82.of_tactic
+ (fst_prod false (fun id -> orig := id; Tactics.intro_mustbe_force name)))
+;;
+
+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
+ | _ -> mk_anon_id ssr_anon_hyp gl in
+ introid id gl
+
+let intro_all gl =
+ let dc, _ = EConstr.decompose_prod_assum (project gl) (Tacmach.pf_concl gl) in
+ tclTHENLIST (List.map anontac (List.rev dc)) gl
+
+let rec intro_anon gl =
+ try anontac (List.hd (fst (EConstr.decompose_prod_n_assum (project gl) 1 (Tacmach.pf_concl gl)))) gl
+ with err0 -> try tclTHEN (Proofview.V82.of_tactic Tactics.red_in_concl) intro_anon gl with e when CErrors.noncritical e -> raise err0
+ (* with _ -> CErrors.error "No product even after reduction" *)
+
+let is_pf_var sigma c =
+ EConstr.isVar sigma c && not_section_id (EConstr.destVar sigma c)
+
+let hyp_of_var sigma v = SsrHyp (Loc.tag @@ EConstr.destVar sigma v)
+
+let interp_clr sigma = function
+| Some clr, (k, c)
+ when (k = xNoFlag || k = xWithAt) && is_pf_var sigma c ->
+ hyp_of_var sigma c :: clr
+| Some clr, _ -> clr
+| None, _ -> []
+
+(** Basic tacticals *)
+
+(** Multipliers {{{ ***********************************************************)
+
+(* tactical *)
+
+let tclID tac = tac
+
+let tclDOTRY n tac =
+ if n <= 0 then tclIDTAC else
+ let rec loop i gl =
+ if i = n then tclTRY tac gl else
+ tclTRY (tclTHEN tac (loop (i + 1))) gl in
+ loop 1
+
+let tclDO n tac =
+ let prefix i = str"At iteration " ++ int i ++ str": " in
+ let tac_err_at i gl =
+ try tac gl
+ with
+ | CErrors.UserError (l, s) as e ->
+ let _, info = CErrors.push e in
+ let e' = CErrors.UserError (l, prefix i ++ s) in
+ Util.iraise (e', info)
+ | Ploc.Exc(loc, CErrors.UserError (l, s)) ->
+ raise (Ploc.Exc(loc, CErrors.UserError (l, prefix i ++ s))) in
+ let rec loop i gl =
+ if i = n then tac_err_at i gl else
+ (tclTHEN (tac_err_at i) (loop (i + 1))) gl in
+ loop 1
+
+let tclMULT = function
+ | 0, May -> tclREPEAT
+ | 1, May -> tclTRY
+ | n, May -> tclDOTRY n
+ | 0, Must -> tclAT_LEAST_ONCE
+ | n, Must when n > 1 -> tclDO n
+ | _ -> tclID
+
+let cleartac clr = check_hyps_uniq [] clr; Proofview.V82.of_tactic (Tactics.clear (hyps_ids clr))
+
+(** }}} *)
+
+(** Generalize tactic *)
+
+(* XXX the k of the redex should percolate out *)
+let pf_interp_gen_aux ist gl to_ind ((oclr, occ), t) =
+ let pat = interp_cpattern ist gl t None in (* UGLY API *)
+ let cl, env, sigma = Tacmach.pf_concl gl, pf_env gl, project gl in
+ let (c, ucst), cl =
+ try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr cl) pat occ 1
+ with NoMatch -> redex_of_pattern env pat, (EConstr.Unsafe.to_constr cl) in
+ let c = EConstr.of_constr c in
+ let cl = EConstr.of_constr cl in
+ let clr = interp_clr sigma (oclr, (tag_of_cpattern t, c)) in
+ if not(occur_existential sigma c) then
+ if tag_of_cpattern t = xWithAt then
+ if not (EConstr.isVar sigma c) then
+ errorstrm (str "@ can be used with variables only")
+ else match Tacmach.pf_get_hyp gl (EConstr.destVar sigma c) with
+ | NamedDecl.LocalAssum _ -> errorstrm (str "@ can be used with let-ins only")
+ | NamedDecl.LocalDef (name, b, ty) -> true, pat, EConstr.mkLetIn (Name name,b,ty,cl),c,clr,ucst,gl
+ else let gl, ccl = pf_mkprod gl c cl in false, pat, ccl, c, clr,ucst,gl
+ else if to_ind && occ = None then
+ let nv, p, _, ucst' = pf_abs_evars gl (fst pat, c) in
+ let ucst = Evd.union_evar_universe_context ucst ucst' in
+ if nv = 0 then anomaly "occur_existential but no evars" else
+ let gl, pty = pfe_type_of gl p in
+ false, pat, EConstr.mkProd (constr_name (project gl) c, pty, Tacmach.pf_concl gl), p, clr,ucst,gl
+ else CErrors.user_err ?loc:(loc_of_cpattern t) (str "generalized term didn't match")
+
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+
+let genclrtac cl cs clr =
+ let tclmyORELSE tac1 tac2 gl =
+ try tac1 gl
+ with e when CErrors.noncritical e -> tac2 e gl in
+ (* apply_type may give a type error, but the useful message is
+ * the one of clear. You type "move: x" and you get
+ * "x is used in hyp H" instead of
+ * "The term H has type T x but is expected to have type T x0". *)
+ tclTHEN
+ (tclmyORELSE
+ (apply_type cl cs)
+ (fun type_err gl ->
+ tclTHEN
+ (tclTHEN (Proofview.V82.of_tactic (Tactics.elim_type (EConstr.of_constr
+ (Universes.constr_of_global @@ Coqlib.build_coq_False ())))) (cleartac clr))
+ (fun gl -> raise type_err)
+ gl))
+ (cleartac clr)
+
+let gentac ist gen gl =
+(* ppdebug(lazy(str"sigma@gentac=" ++ pr_evar_map None (project gl))); *)
+ let conv, _, cl, c, clr, ucst,gl = pf_interp_gen_aux ist gl false gen in
+ ppdebug(lazy(str"c@gentac=" ++ pr_econstr c));
+ let gl = pf_merge_uc ucst gl in
+ if conv
+ then tclTHEN (Proofview.V82.of_tactic (convert_concl cl)) (cleartac clr) gl
+ else genclrtac cl [c] clr gl
+
+let genstac (gens, clr) ist =
+ tclTHENLIST (cleartac clr :: List.rev_map (gentac ist) gens)
+
+let gen_tmp_ids
+ ?(ist=Geninterp.({ lfun = Id.Map.empty; extra = Tacinterp.TacStore.empty })) gl
+=
+ let gl, ctx = pull_ctx gl in
+ push_ctxs ctx
+ (tclTHENLIST
+ (List.map (fun (id,orig_ref) ->
+ tclTHEN
+ (gentac ist ((None,Some(false,[])),cpattern_of_id id))
+ (rename_hd_prod orig_ref))
+ ctx.tmp_ids) gl)
+;;
+
+let pf_interp_gen ist gl to_ind gen =
+ let _, _, a, b, c, ucst,gl = pf_interp_gen_aux ist gl to_ind gen in
+ a, b ,c, pf_merge_uc ucst gl
+
+(* TASSI: This version of unprotects inlines the unfold tactic definition,
+ * since we don't want to wipe out let-ins, and it seems there is no flag
+ * to change that behaviour in the standard unfold code *)
+let unprotecttac gl =
+ let c, gl = pf_mkSsrConst "protect_term" gl in
+ let prot, _ = EConstr.destConst (project gl) c in
+ Tacticals.onClause (fun idopt ->
+ let hyploc = Option.map (fun id -> id, InHyp) idopt in
+ Proofview.V82.of_tactic (Tactics.reduct_option
+ (Reductionops.clos_norm_flags
+ (CClosure.RedFlags.mkflags
+ [CClosure.RedFlags.fBETA;
+ CClosure.RedFlags.fCONST prot;
+ CClosure.RedFlags.fMATCH;
+ CClosure.RedFlags.fFIX;
+ CClosure.RedFlags.fCOFIX]), DEFAULTcast) hyploc))
+ allHypsAndConcl gl
+
+let abs_wgen keep_let ist f gen (gl,args,c) =
+ let sigma, env = project gl, pf_env gl in
+ let evar_closed t p =
+ if occur_existential sigma t then
+ CErrors.user_err ?loc:(loc_of_cpattern p) ~hdr:"ssreflect"
+ (pr_constr_pat (EConstr.Unsafe.to_constr t) ++
+ str" contains holes and matches no subterm of the goal") in
+ match gen with
+ | _, Some ((x, mode), None) when mode = "@" || (mode = " " && keep_let) ->
+ let x = hoi_id x in
+ let decl = Tacmach.pf_get_hyp gl x in
+ gl,
+ (if NamedDecl.is_local_def decl then args else EConstr.mkVar x :: args),
+ EConstr.mkProd_or_LetIn (decl |> NamedDecl.to_rel_decl |> RelDecl.set_name (Name (f x)))
+ (EConstr.Vars.subst_var x c)
+ | _, Some ((x, _), None) ->
+ let x = hoi_id x in
+ gl, EConstr.mkVar x :: args, EConstr.mkProd (Name (f x),Tacmach.pf_get_hyp_typ gl x, EConstr.Vars.subst_var x c)
+ | _, Some ((x, "@"), Some p) ->
+ let x = hoi_id x in
+ let cp = interp_cpattern ist gl p None in
+ let (t, ucst), c =
+ try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr c) cp None 1
+ with NoMatch -> redex_of_pattern env cp, (EConstr.Unsafe.to_constr c) in
+ let c = EConstr.of_constr c in
+ let t = EConstr.of_constr t in
+ evar_closed t p;
+ let ut = red_product_skip_id env sigma t in
+ let gl, ty = pfe_type_of gl t in
+ pf_merge_uc ucst gl, args, EConstr.mkLetIn(Name (f x), ut, ty, c)
+ | _, Some ((x, _), Some p) ->
+ let x = hoi_id x in
+ let cp = interp_cpattern ist gl p None in
+ let (t, ucst), c =
+ try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr c) cp None 1
+ with NoMatch -> redex_of_pattern env cp, (EConstr.Unsafe.to_constr c) in
+ let c = EConstr.of_constr c in
+ let t = EConstr.of_constr t in
+ evar_closed t p;
+ let gl, ty = pfe_type_of gl t in
+ pf_merge_uc ucst gl, t :: args, EConstr.mkProd(Name (f x), ty, c)
+ | _ -> gl, args, c
+
+let clr_of_wgen gen clrs = match gen with
+ | clr, Some ((x, _), None) ->
+ let x = hoi_id x in
+ cleartac clr :: cleartac [SsrHyp(Loc.tag x)] :: clrs
+ | clr, _ -> cleartac clr :: clrs
+
+
+(* vim: set filetype=ocaml foldmethod=marker: *)
diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli
new file mode 100644
index 000000000..834b7b722
--- /dev/null
+++ b/plugins/ssr/ssrcommon.mli
@@ -0,0 +1,410 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Environ
+open Proof_type
+open Evd
+open Constrexpr
+open Ssrast
+
+open Ltac_plugin
+open Genarg
+
+val allocc : ssrocc
+
+(******************************** hyps ************************************)
+
+val hyp_id : ssrhyp -> Id.t
+val hyps_ids : ssrhyps -> Id.t list
+val check_hyp_exists : ('a, 'b) Context.Named.pt -> ssrhyp -> unit
+val test_hypname_exists : ('a, 'b) Context.Named.pt -> Id.t -> bool
+val check_hyps_uniq : Id.t list -> ssrhyps -> unit
+val not_section_id : Id.t -> bool
+val hyp_err : ?loc:Loc.t -> string -> Id.t -> 'a
+val hoik : (ssrhyp -> 'a) -> ssrhyp_or_id -> 'a
+val hoi_id : ssrhyp_or_id -> Id.t
+
+(******************************* hints ***********************************)
+
+val mk_hint : 'a -> 'a ssrhint
+val mk_orhint : 'a -> bool * 'a
+val nullhint : bool * 'a list
+val nohint : 'a ssrhint
+
+(******************************** misc ************************************)
+
+val errorstrm : Pp.std_ppcmds -> 'a
+val anomaly : string -> 'a
+
+val array_app_tl : 'a array -> 'a list -> 'a list
+val array_list_of_tl : 'a array -> 'a list
+val array_fold_right_from : int -> ('a -> 'b -> 'b) -> 'a array -> 'b -> 'b
+
+(**************************** lifted tactics ******************************)
+
+(* tactics with extra data attached to each goals, e.g. the list of
+ * temporary variables to be cleared *)
+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;
+ wild_ids : Id.t list;
+ (* List of variables to be cleared at the end of the sentence *)
+ delayed_clears : Id.t list;
+}
+
+val new_ctx : unit -> tac_ctx (* REMOVE *)
+val pull_ctxs : ('a * tac_ctx) list sigma -> 'a list sigma * tac_ctx list (* REMOVE *)
+
+val with_fresh_ctx : tac_ctx tac_a -> tactic
+
+val pull_ctx : ('a * tac_ctx) sigma -> 'a sigma * tac_ctx
+val push_ctx : tac_ctx -> 'a sigma -> ('a * tac_ctx) sigma
+val push_ctxs : tac_ctx -> 'a list sigma -> ('a * tac_ctx) list sigma
+val tac_ctx : tactic -> tac_ctx tac_a
+val with_ctx :
+ (tac_ctx -> 'b * tac_ctx) -> ('a * tac_ctx) sigma -> 'b * ('a * tac_ctx) sigma
+val without_ctx : ('a sigma -> 'b) -> ('a * tac_ctx) sigma -> 'b
+
+(* Standard tacticals lifted to the tac_a type *)
+val tclTHENLIST_a : tac_ctx tac_a list -> tac_ctx tac_a
+val tclTHEN_i_max :
+ tac_ctx tac_a -> (int -> int -> tac_ctx tac_a) -> tac_ctx tac_a
+val tclTHEN_a : tac_ctx tac_a -> tac_ctx tac_a -> tac_ctx tac_a
+val tclTHENS_a : tac_ctx tac_a -> tac_ctx tac_a list -> tac_ctx tac_a
+
+val tac_on_all :
+ (goal * tac_ctx) list sigma -> tac_ctx tac_a -> (goal * tac_ctx) list sigma
+(************************ ssr tactic arguments ******************************)
+
+
+(*********************** Misc helpers *****************************)
+val mkRHole : Glob_term.glob_constr
+val mkRHoles : int -> Glob_term.glob_constr list
+val isRHoles : Glob_term.glob_constr list -> bool
+val mkRApp : Glob_term.glob_constr -> Glob_term.glob_constr list -> Glob_term.glob_constr
+val mkRVar : Id.t -> Glob_term.glob_constr
+val mkRltacVar : Id.t -> Glob_term.glob_constr
+val mkRCast : Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr
+val mkRType : Glob_term.glob_constr
+val mkRProp : Glob_term.glob_constr
+val mkRArrow : Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr
+val mkRConstruct : Names.constructor -> Glob_term.glob_constr
+val mkRInd : Names.inductive -> Glob_term.glob_constr
+val mkRLambda : Name.t -> Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr
+val mkRnat : int -> Glob_term.glob_constr
+
+
+val mkCHole : Loc.t option -> constr_expr
+val mkCHoles : ?loc:Loc.t -> int -> constr_expr list
+val mkCVar : ?loc:Loc.t -> Id.t -> constr_expr
+val mkCCast : ?loc:Loc.t -> constr_expr -> constr_expr -> constr_expr
+val mkCType : Loc.t option -> constr_expr
+val mkCProp : Loc.t option -> constr_expr
+val mkCArrow : ?loc:Loc.t -> constr_expr -> constr_expr -> constr_expr
+val mkCLambda : ?loc:Loc.t -> Name.t -> constr_expr -> constr_expr -> constr_expr
+
+val isCHoles : constr_expr list -> bool
+val isCxHoles : (constr_expr * 'a option) list -> bool
+
+val intern_term :
+ Tacinterp.interp_sign -> env ->
+ ssrterm -> Glob_term.glob_constr
+
+val pf_intern_term :
+ Tacinterp.interp_sign -> Proof_type.goal Evd.sigma ->
+ ssrterm -> Glob_term.glob_constr
+
+val interp_term :
+ Tacinterp.interp_sign -> Proof_type.goal Evd.sigma ->
+ ssrterm -> evar_map * EConstr.t
+
+val interp_wit :
+ ('a, 'b, 'c) genarg_type -> ist -> goal sigma -> 'b -> evar_map * 'c
+
+val interp_hyp : ist -> goal sigma -> ssrhyp -> evar_map * ssrhyp
+val interp_hyps : ist -> goal sigma -> ssrhyps -> evar_map * ssrhyps
+
+val interp_refine :
+ Tacinterp.interp_sign -> Proof_type.goal Evd.sigma ->
+ Glob_term.glob_constr -> evar_map * (evar_map * EConstr.constr)
+
+val interp_open_constr :
+ Tacinterp.interp_sign -> Proof_type.goal Evd.sigma ->
+ Tacexpr.glob_constr_and_expr -> evar_map * (evar_map * EConstr.t)
+
+val pf_e_type_of :
+ Proof_type.goal Evd.sigma ->
+ EConstr.constr -> Proof_type.goal Evd.sigma * EConstr.types
+
+val splay_open_constr :
+ Proof_type.goal Evd.sigma ->
+ evar_map * EConstr.t ->
+ (Names.Name.t * EConstr.t) list * EConstr.t
+val isAppInd : Proof_type.goal Evd.sigma -> EConstr.types -> bool
+val interp_view_nbimps :
+ Tacinterp.interp_sign ->
+ Proof_type.goal Evd.sigma -> Glob_term.glob_constr -> int
+val interp_nbargs :
+ Tacinterp.interp_sign ->
+ Proof_type.goal Evd.sigma -> Glob_term.glob_constr -> int
+
+
+val mk_term : ssrtermkind -> 'b -> ssrtermkind * (Glob_term.glob_constr * 'b option)
+val mk_lterm : 'a -> ssrtermkind * (Glob_term.glob_constr * 'a option)
+
+val is_internal_name : string -> bool
+val add_internal_name : (string -> bool) -> unit
+val mk_internal_id : string -> Id.t
+val mk_tagged_id : string -> int -> Id.t
+val mk_evar_name : int -> Name.t
+val ssr_anon_hyp : string
+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
+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
+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
+val pf_merge_uc_of :
+ evar_map -> 'a Evd.sigma -> 'a Evd.sigma
+val constr_name : evar_map -> EConstr.t -> Name.t
+val pf_type_of :
+ Proof_type.goal Evd.sigma ->
+ Term.constr -> Proof_type.goal Evd.sigma * Term.types
+val pfe_type_of :
+ Proof_type.goal Evd.sigma ->
+ EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types
+val pf_abs_prod :
+ Names.name ->
+ 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 ->
+ EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types
+
+val mkSsrRRef : string -> Glob_term.glob_constr * 'a option
+val mkSsrRef : string -> Globnames.global_reference
+val mkSsrConst :
+ string ->
+ env -> evar_map -> evar_map * EConstr.t
+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 pf_fresh_global :
+ Globnames.global_reference ->
+ Proof_type.goal Evd.sigma ->
+ Term.constr * Proof_type.goal Evd.sigma
+
+val is_discharged_id : Id.t -> bool
+val mk_discharged_id : Id.t -> Id.t
+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
+val mk_anon_id : string -> Proof_type.goal Evd.sigma -> Id.t
+val pf_abs_evars_pirrel :
+ Proof_type.goal Evd.sigma ->
+ evar_map * Term.constr -> int * Term.constr
+val pf_nbargs : Proof_type.goal Evd.sigma -> EConstr.t -> int
+val gen_tmp_ids :
+ ?ist:Geninterp.interp_sign ->
+ (Proof_type.goal * tac_ctx) Evd.sigma ->
+ (Proof_type.goal * tac_ctx) list Evd.sigma
+
+val ssrevaltac : Tacinterp.interp_sign -> Tacinterp.Value.t -> Proofview.V82.tac
+
+val convert_concl_no_check : EConstr.t -> unit Proofview.tactic
+val convert_concl : EConstr.t -> unit Proofview.tactic
+
+val red_safe :
+ Reductionops.reduction_function ->
+ env -> evar_map -> EConstr.t -> EConstr.t
+
+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
+
+val mkProt :
+ EConstr.t ->
+ EConstr.t ->
+ Proof_type.goal Evd.sigma ->
+ EConstr.t * Proof_type.goal Evd.sigma
+
+val mkEtaApp : EConstr.t -> int -> int -> EConstr.t
+
+val mkRefl :
+ EConstr.t ->
+ EConstr.t ->
+ Proof_type.goal Evd.sigma -> EConstr.t * Proof_type.goal Evd.sigma
+
+val discharge_hyp :
+ Id.t * (Id.t * string) ->
+ Proof_type.goal Evd.sigma -> Evar.t list Evd.sigma
+
+val clear_wilds_and_tmp_and_delayed_ids :
+ (Proof_type.goal * tac_ctx) Evd.sigma ->
+ (Proof_type.goal * tac_ctx) list Evd.sigma
+
+val view_error : string -> ssrterm -> 'a
+
+
+val top_id : Id.t
+
+val pf_abs_ssrterm :
+ ?resolve_typeclasses:bool ->
+ ist ->
+ Proof_type.goal Evd.sigma ->
+ ssrterm ->
+ evar_map * EConstr.t * Evd.evar_universe_context * int
+
+val pf_interp_ty :
+ ?resolve_typeclasses:bool ->
+ Tacinterp.interp_sign ->
+ Proof_type.goal Evd.sigma ->
+ Ssrast.ssrtermkind *
+ (Glob_term.glob_constr * Constrexpr.constr_expr option) ->
+ int * EConstr.t * EConstr.t * Evd.evar_universe_context
+
+val ssr_n_tac : string -> int -> v82tac
+val donetac : int -> v82tac
+
+val applyn :
+ with_evars:bool ->
+ ?beta:bool ->
+ ?with_shelve:bool ->
+ int ->
+ EConstr.t -> v82tac
+exception NotEnoughProducts
+val pf_saturate :
+ ?beta:bool ->
+ ?bi_types:bool ->
+ Proof_type.goal Evd.sigma ->
+ EConstr.constr ->
+ ?ty:EConstr.types ->
+ int ->
+ EConstr.constr * EConstr.types * (int * EConstr.constr) list *
+ Proof_type.goal Evd.sigma
+val saturate :
+ ?beta:bool ->
+ ?bi_types:bool ->
+ env ->
+ evar_map ->
+ EConstr.constr ->
+ ?ty:EConstr.types ->
+ int ->
+ EConstr.constr * EConstr.types * (int * EConstr.constr) list * evar_map
+val refine_with :
+ ?first_goes_last:bool ->
+ ?beta:bool ->
+ ?with_evars:bool ->
+ evar_map * EConstr.t -> v82tac
+(*********************** Wrapped Coq tactics *****************************)
+
+val rewritetac : ssrdir -> EConstr.t -> tactic
+
+type name_hint = (int * EConstr.types array) option ref
+
+val gentac :
+ (Geninterp.interp_sign ->
+ (Ssrast.ssrdocc) *
+ Ssrmatching_plugin.Ssrmatching.cpattern -> Proof_type.tactic)
+
+val genstac :
+ ((Ssrast.ssrhyp list option * Ssrmatching_plugin.Ssrmatching.occ) *
+ Ssrmatching_plugin.Ssrmatching.cpattern)
+ list * Ssrast.ssrhyp list ->
+ Tacinterp.interp_sign -> Proof_type.tactic
+
+val pf_interp_gen :
+ Tacinterp.interp_sign ->
+ Proof_type.goal Evd.sigma ->
+ bool ->
+ (Ssrast.ssrhyp list option * Ssrmatching_plugin.Ssrmatching.occ) *
+ Ssrmatching_plugin.Ssrmatching.cpattern ->
+ EConstr.t * EConstr.t * Ssrast.ssrhyp list *
+ Proof_type.goal Evd.sigma
+
+val pf_interp_gen_aux :
+ Tacinterp.interp_sign ->
+ Proof_type.goal Evd.sigma ->
+ bool ->
+ (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 *
+ Proof_type.goal Evd.sigma
+
+val is_name_in_ipats :
+ Id.t -> ssripats -> bool
+
+type profiler = {
+ profile : 'a 'b. ('a -> 'b) -> 'a -> 'b;
+ reset : unit -> unit;
+ print : unit -> unit }
+
+val mk_profiler : string -> profiler
+
+(** Basic tactics *)
+
+val introid : ?orig:name ref -> Id.t -> v82tac
+val intro_anon : v82tac
+val intro_all : v82tac
+
+val interp_clr :
+ evar_map -> ssrhyps option * (ssrtermkind * EConstr.t) -> ssrhyps
+
+val genclrtac :
+ EConstr.constr ->
+ EConstr.constr list -> Ssrast.ssrhyp list -> Proof_type.tactic
+val cleartac : ssrhyps -> v82tac
+
+val tclMULT : int * ssrmmod -> Proof_type.tactic -> Proof_type.tactic
+
+val unprotecttac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val abs_wgen :
+ bool ->
+ Tacinterp.interp_sign ->
+ (Id.t -> Id.t) ->
+ 'a *
+ ((Ssrast.ssrhyp_or_id * string) *
+ Ssrmatching_plugin.Ssrmatching.cpattern option)
+ option ->
+ Proof_type.goal Evd.sigma * EConstr.t list * EConstr.t ->
+ Proof_type.goal Evd.sigma * EConstr.t list * EConstr.t
+
+val clr_of_wgen :
+ ssrhyps * ((ssrhyp_or_id * 'a) * 'b option) option ->
+ Proofview.V82.tac list -> Proofview.V82.tac list
+
+
diff --git a/plugins/ssr/ssreflect.v b/plugins/ssr/ssreflect.v
new file mode 100644
index 000000000..1c599ac8c
--- /dev/null
+++ b/plugins/ssr/ssreflect.v
@@ -0,0 +1,451 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+Require Import Bool. (* For bool_scope delimiter 'bool'. *)
+Require Import ssrmatching.
+Declare ML Module "ssreflect_plugin".
+
+(******************************************************************************)
+(* This file is the Gallina part of the ssreflect plugin implementation. *)
+(* Files that use the ssreflect plugin should always Require ssreflect and *)
+(* either Import ssreflect or Import ssreflect.SsrSyntax. *)
+(* Part of the contents of this file is technical and will only interest *)
+(* advanced developers; in addition the following are defined: *)
+(* [the str of v by f] == the Canonical s : str such that f s = v. *)
+(* [the str of v] == the Canonical s : str that coerces to v. *)
+(* argumentType c == the T such that c : forall x : T, P x. *)
+(* returnType c == the R such that c : T -> R. *)
+(* {type of c for s} == P s where c : forall x : T, P x. *)
+(* phantom T v == singleton type with inhabitant Phantom T v. *)
+(* phant T == singleton type with inhabitant Phant v. *)
+(* =^~ r == the converse of rewriting rule r (e.g., in a *)
+(* rewrite multirule). *)
+(* unkeyed t == t, but treated as an unkeyed matching pattern by *)
+(* the ssreflect matching algorithm. *)
+(* nosimpl t == t, but on the right-hand side of Definition C := *)
+(* nosimpl disables expansion of C by /=. *)
+(* locked t == t, but locked t is not convertible to t. *)
+(* locked_with k t == t, but not convertible to t or locked_with k' t *)
+(* unless k = k' (with k : unit). Coq type-checking *)
+(* will be much more efficient if locked_with with a *)
+(* bespoke k is used for sealed definitions. *)
+(* unlockable v == interface for sealed constant definitions of v. *)
+(* Unlockable def == the unlockable that registers def : C = v. *)
+(* [unlockable of C] == a clone for C of the canonical unlockable for the *)
+(* definition of C (e.g., if it uses locked_with). *)
+(* [unlockable fun C] == [unlockable of C] with the expansion forced to be *)
+(* an explicit lambda expression. *)
+(* -> The usage pattern for ADT operations is: *)
+(* Definition foo_def x1 .. xn := big_foo_expression. *)
+(* Fact foo_key : unit. Proof. by []. Qed. *)
+(* Definition foo := locked_with foo_key foo_def. *)
+(* Canonical foo_unlockable := [unlockable fun foo]. *)
+(* This minimizes the comparison overhead for foo, while still allowing *)
+(* rewrite unlock to expose big_foo_expression. *)
+(* More information about these definitions and their use can be found in the *)
+(* ssreflect manual, and in specific comments below. *)
+(******************************************************************************)
+
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Module SsrSyntax.
+
+(* Declare Ssr keywords: 'is' 'of' '//' '/=' and '//='. We also declare the *)
+(* parsing level 8, as a workaround for a notation grammar factoring problem. *)
+(* Arguments of application-style notations (at level 10) should be declared *)
+(* at level 8 rather than 9 or the camlp5 grammar will not factor properly. *)
+
+Reserved Notation "(* x 'is' y 'of' z 'isn't' // /= //= *)" (at level 8).
+Reserved Notation "(* 69 *)" (at level 69).
+
+(* Non ambiguous keyword to check if the SsrSyntax module is imported *)
+Reserved Notation "(* Use to test if 'SsrSyntax_is_Imported' *)" (at level 8).
+
+Reserved Notation "<hidden n >" (at level 200).
+Reserved Notation "T (* n *)" (at level 200, format "T (* n *)").
+
+End SsrSyntax.
+
+Export SsrMatchingSyntax.
+Export SsrSyntax.
+
+(* Make the general "if" into a notation, so that we can override it below. *)
+(* The notations are "only parsing" because the Coq decompiler will not *)
+(* recognize the expansion of the boolean if; using the default printer *)
+(* avoids a spurrious trailing %GEN_IF. *)
+
+Delimit Scope general_if_scope with GEN_IF.
+
+Notation "'if' c 'then' v1 'else' v2" :=
+ (if c then v1 else v2)
+ (at level 200, c, v1, v2 at level 200, only parsing) : general_if_scope.
+
+Notation "'if' c 'return' t 'then' v1 'else' v2" :=
+ (if c return t then v1 else v2)
+ (at level 200, c, t, v1, v2 at level 200, only parsing) : general_if_scope.
+
+Notation "'if' c 'as' x 'return' t 'then' v1 'else' v2" :=
+ (if c as x return t then v1 else v2)
+ (at level 200, c, t, v1, v2 at level 200, x ident, only parsing)
+ : general_if_scope.
+
+(* Force boolean interpretation of simple if expressions. *)
+
+Delimit Scope boolean_if_scope with BOOL_IF.
+
+Notation "'if' c 'return' t 'then' v1 'else' v2" :=
+ (if c%bool is true in bool return t then v1 else v2) : boolean_if_scope.
+
+Notation "'if' c 'then' v1 'else' v2" :=
+ (if c%bool is true in bool return _ then v1 else v2) : boolean_if_scope.
+
+Notation "'if' c 'as' x 'return' t 'then' v1 'else' v2" :=
+ (if c%bool is true as x in bool return t then v1 else v2) : boolean_if_scope.
+
+Open Scope boolean_if_scope.
+
+(* To allow a wider variety of notations without reserving a large number of *)
+(* of identifiers, the ssreflect library systematically uses "forms" to *)
+(* enclose complex mixfix syntax. A "form" is simply a mixfix expression *)
+(* enclosed in square brackets and introduced by a keyword: *)
+(* [keyword ... ] *)
+(* Because the keyword follows a bracket it does not need to be reserved. *)
+(* Non-ssreflect libraries that do not respect the form syntax (e.g., the Coq *)
+(* Lists library) should be loaded before ssreflect so that their notations *)
+(* do not mask all ssreflect forms. *)
+Delimit Scope form_scope with FORM.
+Open Scope form_scope.
+
+(* Allow overloading of the cast (x : T) syntax, put whitespace around the *)
+(* ":" symbol to avoid lexical clashes (and for consistency with the parsing *)
+(* precedence of the notation, which binds less tightly than application), *)
+(* and put printing boxes that print the type of a long definition on a *)
+(* separate line rather than force-fit it at the right margin. *)
+Notation "x : T" := (x : T)
+ (at level 100, right associativity,
+ format "'[hv' x '/ ' : T ']'") : core_scope.
+
+(* Allow the casual use of notations like nat * nat for explicit Type *)
+(* declarations. Note that (nat * nat : Type) is NOT equivalent to *)
+(* (nat * nat)%type, whose inferred type is legacy type "Set". *)
+Notation "T : 'Type'" := (T%type : Type)
+ (at level 100, only parsing) : core_scope.
+(* Allow similarly Prop annotation for, e.g., rewrite multirules. *)
+Notation "P : 'Prop'" := (P%type : Prop)
+ (at level 100, only parsing) : core_scope.
+
+(* Constants for abstract: and [: name ] intro pattern *)
+Definition abstract_lock := unit.
+Definition abstract_key := tt.
+
+Definition abstract (statement : Type) (id : nat) (lock : abstract_lock) :=
+ let: tt := lock in statement.
+
+Notation "<hidden n >" := (abstract _ n _).
+Notation "T (* n *)" := (abstract T n abstract_key).
+
+(* Constants for tactic-views *)
+Inductive external_view : Type := tactic_view of Type.
+
+(* Syntax for referring to canonical structures: *)
+(* [the struct_type of proj_val by proj_fun] *)
+(* This form denotes the Canonical instance s of the Structure type *)
+(* struct_type whose proj_fun projection is proj_val, i.e., such that *)
+(* proj_fun s = proj_val. *)
+(* Typically proj_fun will be A record field accessors of struct_type, but *)
+(* this need not be the case; it can be, for instance, a field of a record *)
+(* type to which struct_type coerces; proj_val will likewise be coerced to *)
+(* the return type of proj_fun. In all but the simplest cases, proj_fun *)
+(* should be eta-expanded to allow for the insertion of implicit arguments. *)
+(* In the common case where proj_fun itself is a coercion, the "by" part *)
+(* can be omitted entirely; in this case it is inferred by casting s to the *)
+(* inferred type of proj_val. Obviously the latter can be fixed by using an *)
+(* explicit cast on proj_val, and it is highly recommended to do so when the *)
+(* return type intended for proj_fun is "Type", as the type inferred for *)
+(* proj_val may vary because of sort polymorphism (it could be Set or Prop). *)
+(* Note when using the [the _ of _] form to generate a substructure from a *)
+(* telescopes-style canonical hierarchy (implementing inheritance with *)
+(* coercions), one should always project or coerce the value to the BASE *)
+(* structure, because Coq will only find a Canonical derived structure for *)
+(* the Canonical base structure -- not for a base structure that is specific *)
+(* to proj_value. *)
+
+Module TheCanonical.
+
+CoInductive put vT sT (v1 v2 : vT) (s : sT) := Put.
+
+Definition get vT sT v s (p : @put vT sT v v s) := let: Put _ _ _ := p in s.
+
+Definition get_by vT sT of sT -> vT := @get vT sT.
+
+End TheCanonical.
+
+Import TheCanonical. (* Note: no export. *)
+
+Local Arguments get_by _%type_scope _%type_scope _ _ _ _.
+
+Notation "[ 'the' sT 'of' v 'by' f ]" :=
+ (@get_by _ sT f _ _ ((fun v' (s : sT) => Put v' (f s) s) v _))
+ (at level 0, only parsing) : form_scope.
+
+Notation "[ 'the' sT 'of' v ]" := (get ((fun s : sT => Put v (*coerce*)s s) _))
+ (at level 0, only parsing) : form_scope.
+
+(* The following are "format only" versions of the above notations. Since Coq *)
+(* doesn't provide this facility, we fake it by splitting the "the" keyword. *)
+(* We need to do this to prevent the formatter from being be thrown off by *)
+(* application collapsing, coercion insertion and beta reduction in the right *)
+(* hand side of the notations above. *)
+
+Notation "[ 'th' 'e' sT 'of' v 'by' f ]" := (@get_by _ sT f v _ _)
+ (at level 0, format "[ 'th' 'e' sT 'of' v 'by' f ]") : form_scope.
+
+Notation "[ 'th' 'e' sT 'of' v ]" := (@get _ sT v _ _)
+ (at level 0, format "[ 'th' 'e' sT 'of' v ]") : form_scope.
+
+(* We would like to recognize
+Notation "[ 'th' 'e' sT 'of' v : 'Type' ]" := (@get Type sT v _ _)
+ (at level 0, format "[ 'th' 'e' sT 'of' v : 'Type' ]") : form_scope.
+*)
+
+(* Helper notation for canonical structure inheritance support. *)
+(* This is a workaround for the poor interaction between delta reduction and *)
+(* canonical projections in Coq's unification algorithm, by which transparent *)
+(* definitions hide canonical instances, i.e., in *)
+(* Canonical a_type_struct := @Struct a_type ... *)
+(* Definition my_type := a_type. *)
+(* my_type doesn't effectively inherit the struct structure from a_type. Our *)
+(* solution is to redeclare the instance as follows *)
+(* Canonical my_type_struct := Eval hnf in [struct of my_type]. *)
+(* The special notation [str of _] must be defined for each Strucure "str" *)
+(* with constructor "Str", typically as follows *)
+(* Definition clone_str s := *)
+(* let: Str _ x y ... z := s return {type of Str for s} -> str in *)
+(* fun k => k _ x y ... z. *)
+(* Notation "[ 'str' 'of' T 'for' s ]" := (@clone_str s (@Str T)) *)
+(* (at level 0, format "[ 'str' 'of' T 'for' s ]") : form_scope. *)
+(* Notation "[ 'str' 'of' T ]" := (repack_str (fun x => @Str T x)) *)
+(* (at level 0, format "[ 'str' 'of' T ]") : form_scope. *)
+(* The notation for the match return predicate is defined below; the eta *)
+(* expansion in the second form serves both to distinguish it from the first *)
+(* and to avoid the delta reduction problem. *)
+(* There are several variations on the notation and the definition of the *)
+(* the "clone" function, for telescopes, mixin classes, and join (multiple *)
+(* inheritance) classes. We describe a different idiom for clones in ssrfun; *)
+(* it uses phantom types (see below) and static unification; see fintype and *)
+(* ssralg for examples. *)
+
+Definition argumentType T P & forall x : T, P x := T.
+Definition dependentReturnType T P & forall x : T, P x := P.
+Definition returnType aT rT & aT -> rT := rT.
+
+Notation "{ 'type' 'of' c 'for' s }" := (dependentReturnType c s)
+ (at level 0, format "{ 'type' 'of' c 'for' s }") : type_scope.
+
+(* A generic "phantom" type (actually, a unit type with a phantom parameter). *)
+(* This type can be used for type definitions that require some Structure *)
+(* on one of their parameters, to allow Coq to infer said structure so it *)
+(* does not have to be supplied explicitly or via the "[the _ of _]" notation *)
+(* (the latter interacts poorly with other Notation). *)
+(* The definition of a (co)inductive type with a parameter p : p_type, that *)
+(* needs to use the operations of a structure *)
+(* Structure p_str : Type := p_Str {p_repr :> p_type; p_op : p_repr -> ...} *)
+(* should be given as *)
+(* Inductive indt_type (p : p_str) := Indt ... . *)
+(* Definition indt_of (p : p_str) & phantom p_type p := indt_type p. *)
+(* Notation "{ 'indt' p }" := (indt_of (Phantom p)). *)
+(* Definition indt p x y ... z : {indt p} := @Indt p x y ... z. *)
+(* Notation "[ 'indt' x y ... z ]" := (indt x y ... z). *)
+(* That is, the concrete type and its constructor should be shadowed by *)
+(* definitions that use a phantom argument to infer and display the true *)
+(* value of p (in practice, the "indt" constructor often performs additional *)
+(* functions, like "locking" the representation -- see below). *)
+(* We also define a simpler version ("phant" / "Phant") of phantom for the *)
+(* common case where p_type is Type. *)
+
+CoInductive phantom T (p : T) := Phantom.
+Arguments phantom : clear implicits.
+Arguments Phantom : clear implicits.
+CoInductive phant (p : Type) := Phant.
+
+(* Internal tagging used by the implementation of the ssreflect elim. *)
+
+Definition protect_term (A : Type) (x : A) : A := x.
+
+(* The ssreflect idiom for a non-keyed pattern: *)
+(* - unkeyed t wiil match any subterm that unifies with t, regardless of *)
+(* whether it displays the same head symbol as t. *)
+(* - unkeyed t a b will match any application of a term f unifying with t, *)
+(* to two arguments unifying with with a and b, repectively, regardless of *)
+(* apparent head symbols. *)
+(* - unkeyed x where x is a variable will match any subterm with the same *)
+(* type as x (when x would raise the 'indeterminate pattern' error). *)
+
+Notation unkeyed x := (let flex := x in flex).
+
+(* Ssreflect converse rewrite rule rule idiom. *)
+Definition ssr_converse R (r : R) := (Logic.I, r).
+Notation "=^~ r" := (ssr_converse r) (at level 100) : form_scope.
+
+(* Term tagging (user-level). *)
+(* The ssreflect library uses four strengths of term tagging to restrict *)
+(* convertibility during type checking: *)
+(* nosimpl t simplifies to t EXCEPT in a definition; more precisely, given *)
+(* Definition foo := nosimpl bar, foo (or foo t') will NOT be expanded by *)
+(* the /= and //= switches unless it is in a forcing context (e.g., in *)
+(* match foo t' with ... end, foo t' will be reduced if this allows the *)
+(* match to be reduced). Note that nosimpl bar is simply notation for a *)
+(* a term that beta-iota reduces to bar; hence rewrite /foo will replace *)
+(* foo by bar, and rewrite -/foo will replace bar by foo. *)
+(* CAVEAT: nosimpl should not be used inside a Section, because the end of *)
+(* section "cooking" removes the iota redex. *)
+(* locked t is provably equal to t, but is not convertible to t; 'locked' *)
+(* provides support for selective rewriting, via the lock t : t = locked t *)
+(* Lemma, and the ssreflect unlock tactic. *)
+(* locked_with k t is equal but not convertible to t, much like locked t, *)
+(* but supports explicit tagging with a value k : unit. This is used to *)
+(* mitigate a flaw in the term comparison heuristic of the Coq kernel, *)
+(* which treats all terms of the form locked t as equal and conpares their *)
+(* arguments recursively, leading to an exponential blowup of comparison. *)
+(* For this reason locked_with should be used rather than locked when *)
+(* defining ADT operations. The unlock tactic does not support locked_with *)
+(* but the unlock rewrite rule does, via the unlockable interface. *)
+(* we also use Module Type ascription to create truly opaque constants, *)
+(* because simple expansion of constants to reveal an unreducible term *)
+(* doubles the time complexity of a negative comparison. Such opaque *)
+(* constants can be expanded generically with the unlock rewrite rule. *)
+(* See the definition of card and subset in fintype for examples of this. *)
+
+Notation nosimpl t := (let: tt := tt in t).
+
+Lemma master_key : unit. Proof. exact tt. Qed.
+Definition locked A := let: tt := master_key in fun x : A => x.
+
+Lemma lock A x : x = locked x :> A. Proof. unlock; reflexivity. Qed.
+
+(* Needed for locked predicates, in particular for eqType's. *)
+Lemma not_locked_false_eq_true : locked false <> true.
+Proof. unlock; discriminate. Qed.
+
+(* The basic closing tactic "done". *)
+Ltac done :=
+ trivial; hnf; intros; solve
+ [ do ![solve [trivial | apply: sym_equal; trivial]
+ | discriminate | contradiction | split]
+ | case not_locked_false_eq_true; assumption
+ | match goal with H : ~ _ |- _ => solve [case H; trivial] end ].
+
+(* Quicker done tactic not including split, syntax: /0/ *)
+Ltac ssrdone0 :=
+ trivial; hnf; intros; solve
+ [ do ![solve [trivial | apply: sym_equal; trivial]
+ | discriminate | contradiction ]
+ | case not_locked_false_eq_true; assumption
+ | match goal with H : ~ _ |- _ => solve [case H; trivial] end ].
+
+(* To unlock opaque constants. *)
+Structure unlockable T v := Unlockable {unlocked : T; _ : unlocked = v}.
+Lemma unlock T x C : @unlocked T x C = x. Proof. by case: C. Qed.
+
+Notation "[ 'unlockable' 'of' C ]" := (@Unlockable _ _ C (unlock _))
+ (at level 0, format "[ 'unlockable' 'of' C ]") : form_scope.
+
+Notation "[ 'unlockable' 'fun' C ]" := (@Unlockable _ (fun _ => _) C (unlock _))
+ (at level 0, format "[ 'unlockable' 'fun' C ]") : form_scope.
+
+(* Generic keyed constant locking. *)
+
+(* The argument order ensures that k is always compared before T. *)
+Definition locked_with k := let: tt := k in fun T x => x : T.
+
+(* This can be used as a cheap alternative to cloning the unlockable instance *)
+(* below, but with caution as unkeyed matching can be expensive. *)
+Lemma locked_withE T k x : unkeyed (locked_with k x) = x :> T.
+Proof. by case: k. Qed.
+
+(* Intensionaly, this instance will not apply to locked u. *)
+Canonical locked_with_unlockable T k x :=
+ @Unlockable T x (locked_with k x) (locked_withE k x).
+
+(* More accurate variant of unlock, and safer alternative to locked_withE. *)
+Lemma unlock_with T k x : unlocked (locked_with_unlockable k x) = x :> T.
+Proof. exact: unlock. Qed.
+
+(* The internal lemmas for the have tactics. *)
+
+Definition ssr_have Plemma Pgoal (step : Plemma) rest : Pgoal := rest step.
+Arguments ssr_have Plemma [Pgoal].
+
+Definition ssr_have_let Pgoal Plemma step
+ (rest : let x : Plemma := step in Pgoal) : Pgoal := rest.
+Arguments ssr_have_let [Pgoal].
+
+Definition ssr_suff Plemma Pgoal step (rest : Plemma) : Pgoal := step rest.
+Arguments ssr_suff Plemma [Pgoal].
+
+Definition ssr_wlog := ssr_suff.
+Arguments ssr_wlog Plemma [Pgoal].
+
+(* Internal N-ary congruence lemmas for the congr tactic. *)
+
+Fixpoint nary_congruence_statement (n : nat)
+ : (forall B, (B -> B -> Prop) -> Prop) -> Prop :=
+ match n with
+ | O => fun k => forall B, k B (fun x1 x2 : B => x1 = x2)
+ | S n' =>
+ let k' A B e (f1 f2 : A -> B) :=
+ forall x1 x2, x1 = x2 -> (e (f1 x1) (f2 x2) : Prop) in
+ fun k => forall A, nary_congruence_statement n' (fun B e => k _ (k' A B e))
+ end.
+
+Lemma nary_congruence n (k := fun B e => forall y : B, (e y y : Prop)) :
+ nary_congruence_statement n k.
+Proof.
+have: k _ _ := _; rewrite {1}/k.
+elim: n k => [|n IHn] k k_P /= A; first exact: k_P.
+by apply: IHn => B e He; apply: k_P => f x1 x2 <-.
+Qed.
+
+Lemma ssr_congr_arrow Plemma Pgoal : Plemma = Pgoal -> Plemma -> Pgoal.
+Proof. by move->. Qed.
+Arguments ssr_congr_arrow : clear implicits.
+
+(* View lemmas that don't use reflection. *)
+
+Section ApplyIff.
+
+Variables P Q : Prop.
+Hypothesis eqPQ : P <-> Q.
+
+Lemma iffLR : P -> Q. Proof. by case: eqPQ. Qed.
+Lemma iffRL : Q -> P. Proof. by case: eqPQ. Qed.
+
+Lemma iffLRn : ~P -> ~Q. Proof. by move=> nP tQ; case: nP; case: eqPQ tQ. Qed.
+Lemma iffRLn : ~Q -> ~P. Proof. by move=> nQ tP; case: nQ; case: eqPQ tP. Qed.
+
+End ApplyIff.
+
+Hint View for move/ iffLRn|2 iffRLn|2 iffLR|2 iffRL|2.
+Hint View for apply/ iffRLn|2 iffLRn|2 iffRL|2 iffLR|2.
+
+(* To focus non-ssreflect tactics on a subterm, eg vm_compute. *)
+(* Usage: *)
+(* elim/abstract_context: (pattern) => G defG. *)
+(* vm_compute; rewrite {}defG {G}. *)
+(* Note that vm_cast are not stored in the proof term *)
+(* for reductions occuring in the context, hence *)
+(* set here := pattern; vm_compute in (value of here) *)
+(* blows up at Qed time. *)
+Lemma abstract_context T (P : T -> Type) x :
+ (forall Q, Q = P -> Q x) -> P x.
+Proof. by move=> /(_ P); apply. Qed.
diff --git a/plugins/ssr/ssreflect_plugin.mlpack b/plugins/ssr/ssreflect_plugin.mlpack
new file mode 100644
index 000000000..824348fee
--- /dev/null
+++ b/plugins/ssr/ssreflect_plugin.mlpack
@@ -0,0 +1,13 @@
+Ssrast
+Ssrprinters
+Ssrcommon
+Ssrtacticals
+Ssrelim
+Ssrview
+Ssrbwd
+Ssrequality
+Ssripats
+Ssrfwd
+Ssrparser
+Ssrvernac
+
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
new file mode 100644
index 000000000..832044909
--- /dev/null
+++ b/plugins/ssr/ssrelim.ml
@@ -0,0 +1,441 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Util
+open Names
+open Printer
+open Term
+open Termops
+open Globnames
+open Misctypes
+open Tacmach
+
+open Ssrmatching_plugin
+open Ssrmatching
+
+open Ssrast
+open Ssrprinters
+open Ssrcommon
+
+module RelDecl = Context.Rel.Declaration
+
+(** The "case" and "elim" tactic *)
+
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+
+(* TASSI: given the type of an elimination principle, it finds the higher order
+ * argument (index), it computes it's arity and the arity of the eliminator and
+ * checks if the eliminator is recursive or not *)
+let analyze_eliminator elimty env sigma =
+ let rec loop ctx t = match EConstr.kind_of_type sigma t with
+ | AtomicType (hd, args) when EConstr.isRel sigma hd ->
+ ctx, EConstr.destRel sigma hd, not (EConstr.Vars.noccurn sigma 1 t), Array.length args, t
+ | CastType (t, _) -> loop ctx t
+ | ProdType (x, ty, t) -> loop (RelDecl.LocalAssum (x, ty) :: ctx) t
+ | LetInType (x,b,ty,t) -> loop (RelDecl.LocalDef (x, b, ty) :: ctx) (EConstr.Vars.subst1 b t)
+ | _ ->
+ let env' = EConstr.push_rel_context ctx env in
+ let t' = Reductionops.whd_all env' sigma t in
+ if not (EConstr.eq_constr sigma t t') then loop ctx t' else
+ errorstrm Pp.(str"The eliminator has the wrong shape."++spc()++
+ str"A (applied) bound variable was expected as the conclusion of "++
+ str"the eliminator's"++Pp.cut()++str"type:"++spc()++pr_econstr elimty) in
+ let ctx, pred_id, elim_is_dep, n_pred_args,concl = loop [] elimty in
+ let n_elim_args = Context.Rel.nhyps ctx in
+ let is_rec_elim =
+ let count_occurn n term =
+ let count = ref 0 in
+ let rec occur_rec n c = match EConstr.kind sigma c with
+ | Rel m -> if m = n then incr count
+ | _ -> EConstr.iter_with_binders sigma succ occur_rec n c
+ in
+ occur_rec n term; !count in
+ let occurr2 n t = count_occurn n t > 1 in
+ not (List.for_all_i
+ (fun i (_,rd) -> pred_id <= i || not (occurr2 (pred_id - i) rd))
+ 1 (assums_of_rel_context ctx))
+ in
+ n_elim_args - pred_id, n_elim_args, is_rec_elim, elim_is_dep, n_pred_args,
+ (ctx,concl)
+
+let subgoals_tys sigma (relctx, concl) =
+ let rec aux cur_depth acc = function
+ | hd :: rest ->
+ let ty = Context.Rel.Declaration.get_type hd in
+ if EConstr.Vars.noccurn sigma cur_depth concl &&
+ List.for_all_i (fun i -> function
+ | Context.Rel.Declaration.LocalAssum(_, t) ->
+ EConstr.Vars.noccurn sigma i t
+ | Context.Rel.Declaration.LocalDef (_, b, t) ->
+ EConstr.Vars.noccurn sigma i t && EConstr.Vars.noccurn sigma i b) 1 rest
+ then aux (cur_depth - 1) (ty :: acc) rest
+ else aux (cur_depth - 1) acc rest
+ | [] -> Array.of_list (List.rev acc)
+ in
+ aux (List.length relctx) [] (List.rev relctx)
+
+(* A case without explicit dependent terms but with both a view and an *)
+(* occurrence switch and/or an equation is treated as dependent, with the *)
+(* viewed term as the dependent term (the occurrence switch would be *)
+(* meaningless otherwise). When both a view and explicit dependents are *)
+(* present, it is forbidden to put a (meaningless) occurrence switch on *)
+(* the viewed term. *)
+
+(* This is both elim and case (defaulting to the former). If ~elim is omitted
+ * the standard eliminator is chosen. The code is made of 4 parts:
+ * 1. find the eliminator if not given as ~elim and analyze it
+ * 2. build the patterns to be matched against the conclusion, looking at
+ * (occ, c), deps and the pattern inferred from the type of the eliminator
+ * 3. build the new predicate matching the patterns, and the tactic to
+ * generalize the equality in case eqid is not None
+ * 4. build the tactic handle intructions and clears as required in ipats and
+ * by eqid *)
+let ssrelim ?(ind=ref None) ?(is_case=false) ?ist deps what ?elim eqid elim_intro_tac gl =
+ (* some sanity checks *)
+ let oc, orig_clr, occ, c_gen, gl = match what with
+ | `EConstr(_,_,t) when EConstr.isEvar (project gl) t ->
+ anomaly "elim called on a constr evar"
+ | `EGen _ when ist = None ->
+ anomaly "no ist and non simple elimination"
+ | `EGen (_, g) when elim = None && is_wildcard g ->
+ errorstrm Pp.(str"Indeterminate pattern and no eliminator")
+ | `EGen ((Some clr,occ), g) when is_wildcard g ->
+ None, clr, occ, None, gl
+ | `EGen ((None, occ), g) when is_wildcard g -> None,[],occ,None,gl
+ | `EGen ((_, occ), p as gen) ->
+ let _, c, clr,gl = pf_interp_gen (Option.get ist) gl true gen in
+ Some c, clr, occ, Some p,gl
+ | `EConstr (clr, occ, c) -> Some c, clr, occ, None,gl in
+ let orig_gl, concl, env = gl, pf_concl gl, pf_env gl in
+ ppdebug(lazy(Pp.str(if is_case then "==CASE==" else "==ELIM==")));
+ let fire_subst gl t = Reductionops.nf_evar (project gl) t in
+ let eq, gl = pf_fresh_global (Coqlib.build_coq_eq ()) gl in
+ let eq = EConstr.of_constr eq in
+ let is_undef_pat = function
+ | sigma, T t -> EConstr.isEvar sigma (EConstr.of_constr t)
+ | _ -> false in
+ let match_pat env p occ h cl =
+ let sigma0 = project orig_gl in
+ ppdebug(lazy Pp.(str"matching: " ++ pr_occ occ ++ pp_pattern p));
+ let (c,ucst), cl =
+ fill_occ_pattern ~raise_NoMatch:true env sigma0 (EConstr.Unsafe.to_constr cl) p occ h in
+ ppdebug(lazy Pp.(str" got: " ++ pr_constr c));
+ c, EConstr.of_constr cl, ucst in
+ let mkTpat gl t = (* takes a term, refreshes it and makes a T pattern *)
+ let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in
+ let t, _, _, sigma = saturate ~beta:true env (project gl) t n in
+ Evd.merge_universe_context sigma ucst, T (EConstr.Unsafe.to_constr t) in
+ let unif_redex gl (sigma, r as p) t = (* t is a hint for the redex of p *)
+ let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in
+ let t, _, _, sigma = saturate ~beta:true env sigma t n in
+ let sigma = Evd.merge_universe_context sigma ucst in
+ match r with
+ | X_In_T (e, p) -> sigma, E_As_X_In_T (EConstr.Unsafe.to_constr t, e, p)
+ | _ ->
+ try unify_HO env sigma t (EConstr.of_constr (fst (redex_of_pattern env p))), r
+ with e when CErrors.noncritical e -> p in
+ (* finds the eliminator applies it to evars and c saturated as needed *)
+ (* obtaining "elim ??? (c ???)". pred is the higher order evar *)
+ (* cty is None when the user writes _ (hence we can't make a pattern *)
+ let cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl =
+ match elim with
+ | Some elim ->
+ let gl, elimty = pf_e_type_of gl elim in
+ let pred_id, n_elim_args, is_rec, elim_is_dep, n_pred_args,ctx_concl =
+ analyze_eliminator elimty env (project gl) in
+ ind := Some (0, subgoals_tys (project gl) ctx_concl);
+ let elim, elimty, elim_args, gl =
+ pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in
+ let pred = List.assoc pred_id elim_args in
+ let elimty = Reductionops.whd_all env (project gl) elimty in
+ let cty, gl =
+ if Option.is_empty oc then None, gl
+ else
+ let c = Option.get oc in let gl, c_ty = pfe_type_of gl c in
+ let pc = match c_gen with
+ | Some p -> interp_cpattern (Option.get ist) orig_gl p None
+ | _ -> mkTpat gl c in
+ Some(c, c_ty, pc), gl in
+ cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl
+ | None ->
+ let c = Option.get oc in let gl, c_ty = pfe_type_of gl c in
+ let ((kn, i),_ as indu), unfolded_c_ty =
+ pf_reduce_to_quantified_ind gl c_ty in
+ let sort = Tacticals.elimination_sort_of_goal gl in
+ let gl, elim =
+ if not is_case then
+ let t,gl= pf_fresh_global (Indrec.lookup_eliminator (kn,i) sort) gl in
+ gl, t
+ else
+ Tacmach.pf_eapply (fun env sigma () ->
+ let indu = (fst indu, EConstr.EInstance.kind sigma (snd indu)) in
+ let (sigma, ind) = Indrec.build_case_analysis_scheme env sigma indu true sort in
+ (sigma, ind)) gl () in
+ let elim = EConstr.of_constr elim in
+ let gl, elimty = pfe_type_of gl elim in
+ let pred_id,n_elim_args,is_rec,elim_is_dep,n_pred_args,ctx_concl =
+ analyze_eliminator elimty env (project gl) in
+ if is_case then
+ let mind,indb = Inductive.lookup_mind_specif env (kn,i) in
+ ind := Some(mind.Declarations.mind_nparams,Array.map EConstr.of_constr indb.Declarations.mind_nf_lc);
+ else
+ ind := Some (0, subgoals_tys (project gl) ctx_concl);
+ let rctx = fst (EConstr.decompose_prod_assum (project gl) unfolded_c_ty) in
+ let n_c_args = Context.Rel.length rctx in
+ let c, c_ty, t_args, gl = pf_saturate gl c ~ty:c_ty n_c_args in
+ let elim, elimty, elim_args, gl =
+ pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in
+ let pred = List.assoc pred_id elim_args in
+ let pc = match n_c_args, c_gen with
+ | 0, Some p -> interp_cpattern (Option.get ist) orig_gl p None
+ | _ -> mkTpat gl c in
+ let cty = Some (c, c_ty, pc) in
+ let elimty = Reductionops.whd_all env (project gl) elimty in
+ cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl
+ in
+ ppdebug(lazy Pp.(str"elim= "++ pr_constr_pat (EConstr.Unsafe.to_constr elim)));
+ ppdebug(lazy Pp.(str"elimty= "++ pr_constr_pat (EConstr.Unsafe.to_constr elimty)));
+ let inf_deps_r = match EConstr.kind_of_type (project gl) elimty with
+ | AtomicType (_, args) -> List.rev (Array.to_list args)
+ | _ -> assert false in
+ let saturate_until gl c c_ty f =
+ let rec loop n = try
+ let c, c_ty, _, gl = pf_saturate gl c ~ty:c_ty n in
+ let gl' = f c c_ty gl in
+ Some (c, c_ty, gl, gl')
+ with
+ | NotEnoughProducts -> None
+ | e when CErrors.noncritical e -> loop (n+1) in loop 0 in
+ (* Here we try to understand if the main pattern/term the user gave is
+ * the first pattern to be matched (i.e. if elimty ends in P t1 .. tn,
+ * weather tn is the t the user wrote in 'elim: t' *)
+ let c_is_head_p, gl = match cty with
+ | None -> true, gl (* The user wrote elim: _ *)
+ | Some (c, c_ty, _) ->
+ let res =
+ (* we try to see if c unifies with the last arg of elim *)
+ if elim_is_dep then None else
+ let arg = List.assoc (n_elim_args - 1) elim_args in
+ let gl, arg_ty = pfe_type_of gl arg in
+ match saturate_until gl c c_ty (fun c c_ty gl ->
+ pf_unify_HO (pf_unify_HO gl c_ty arg_ty) arg c) with
+ | Some (c, _, _, gl) -> Some (false, gl)
+ | None -> None in
+ match res with
+ | Some x -> x
+ | None ->
+ (* we try to see if c unifies with the last inferred pattern *)
+ let inf_arg = List.hd inf_deps_r in
+ let gl, inf_arg_ty = pfe_type_of gl inf_arg in
+ match saturate_until gl c c_ty (fun _ c_ty gl ->
+ pf_unify_HO gl c_ty inf_arg_ty) with
+ | Some (c, _, _,gl) -> true, gl
+ | None ->
+ errorstrm Pp.(str"Unable to apply the eliminator to the term"++
+ spc()++pr_econstr c++spc()++str"or to unify it's type with"++
+ pr_econstr inf_arg_ty) in
+ ppdebug(lazy Pp.(str"c_is_head_p= " ++ bool c_is_head_p));
+ let gl, predty = pfe_type_of gl pred in
+ (* Patterns for the inductive types indexes to be bound in pred are computed
+ * looking at the ones provided by the user and the inferred ones looking at
+ * the type of the elimination principle *)
+ let pp_pat (_,p,_,occ) = Pp.(pr_occ occ ++ pp_pattern p) in
+ let pp_inf_pat gl (_,_,t,_) = pr_constr_pat (EConstr.Unsafe.to_constr (fire_subst gl t)) in
+ let patterns, clr, gl =
+ let rec loop patterns clr i = function
+ | [],[] -> patterns, clr, gl
+ | ((oclr, occ), t):: deps, inf_t :: inf_deps ->
+ let ist = match ist with Some x -> x | None -> assert false in
+ let p = interp_cpattern ist orig_gl t None in
+ let clr_t =
+ interp_clr (project gl) (oclr,(tag_of_cpattern t,EConstr.of_constr (fst (redex_of_pattern env p)))) in
+ (* if we are the index for the equation we do not clear *)
+ let clr_t = if deps = [] && eqid <> None then [] else clr_t in
+ let p = if is_undef_pat p then mkTpat gl inf_t else p in
+ loop (patterns @ [i, p, inf_t, occ])
+ (clr_t @ clr) (i+1) (deps, inf_deps)
+ | [], c :: inf_deps ->
+ ppdebug(lazy Pp.(str"adding inf pattern " ++ pr_constr_pat (EConstr.Unsafe.to_constr c)));
+ loop (patterns @ [i, mkTpat gl c, c, allocc])
+ clr (i+1) ([], inf_deps)
+ | _::_, [] -> errorstrm Pp.(str "Too many dependent abstractions") in
+ let deps, head_p, inf_deps_r = match what, c_is_head_p, cty with
+ | `EConstr _, _, None -> anomaly "Simple elim with no term"
+ | _, false, _ -> deps, [], inf_deps_r
+ | `EGen gen, true, None -> deps @ [gen], [], inf_deps_r
+ | _, true, Some (c, _, pc) ->
+ let occ = if occ = None then allocc else occ in
+ let inf_p, inf_deps_r = List.hd inf_deps_r, List.tl inf_deps_r in
+ deps, [1, pc, inf_p, occ], inf_deps_r in
+ let patterns, clr, gl =
+ loop [] orig_clr (List.length head_p+1) (List.rev deps, inf_deps_r) in
+ head_p @ patterns, Util.List.uniquize clr, gl
+ in
+ ppdebug(lazy Pp.(pp_concat (str"patterns=") (List.map pp_pat patterns)));
+ ppdebug(lazy Pp.(pp_concat (str"inf. patterns=") (List.map (pp_inf_pat gl) patterns)));
+ (* Predicate generation, and (if necessary) tactic to generalize the
+ * equation asked by the user *)
+ let elim_pred, gen_eq_tac, clr, gl =
+ let error gl t inf_t = errorstrm Pp.(str"The given pattern matches the term"++
+ spc()++pp_term gl t++spc()++str"while the inferred pattern"++
+ spc()++pr_constr_pat (EConstr.Unsafe.to_constr (fire_subst gl inf_t))++spc()++ str"doesn't") in
+ let match_or_postpone (cl, gl, post) (h, p, inf_t, occ) =
+ let p = unif_redex gl p inf_t in
+ if is_undef_pat p then
+ let () = ppdebug(lazy Pp.(str"postponing " ++ pp_pattern p)) in
+ cl, gl, post @ [h, p, inf_t, occ]
+ else try
+ let c, cl, ucst = match_pat env p occ h cl in
+ let gl = pf_merge_uc ucst gl in
+ let c = EConstr.of_constr c in
+ let gl = try pf_unify_HO gl inf_t c with _ -> error gl c inf_t in
+ cl, gl, post
+ with
+ | NoMatch | NoProgress ->
+ let e, ucst = redex_of_pattern env p in
+ let gl = pf_merge_uc ucst gl in
+ let e = EConstr.of_constr e in
+ let n, e, _, _ucst = pf_abs_evars gl (fst p, e) in
+ let e, _, _, gl = pf_saturate ~beta:true gl e n in
+ let gl = try pf_unify_HO gl inf_t e with _ -> error gl e inf_t in
+ cl, gl, post
+ in
+ let rec match_all concl gl patterns =
+ let concl, gl, postponed =
+ List.fold_left match_or_postpone (concl, gl, []) patterns in
+ if postponed = [] then concl, gl
+ else if List.length postponed = List.length patterns then
+ errorstrm Pp.(str "Some patterns are undefined even after all"++spc()++
+ str"the defined ones matched")
+ else match_all concl gl postponed in
+ let concl, gl = match_all concl gl patterns in
+ let pred_rctx, _ = EConstr.decompose_prod_assum (project gl) (fire_subst gl predty) in
+ let concl, gen_eq_tac, clr, gl = match eqid with
+ | Some (IPatId _) when not is_rec ->
+ let k = List.length deps in
+ let c = fire_subst gl (List.assoc (n_elim_args - k - 1) elim_args) in
+ let gl, t = pfe_type_of gl c in
+ let gen_eq_tac, gl =
+ let refl = EConstr.mkApp (eq, [|t; c; c|]) in
+ let new_concl = EConstr.mkArrow refl (EConstr.Vars.lift 1 (pf_concl orig_gl)) in
+ let new_concl = fire_subst gl new_concl in
+ let erefl, gl = mkRefl t c gl in
+ let erefl = fire_subst gl erefl in
+ apply_type new_concl [erefl], gl in
+ let rel = k + if c_is_head_p then 1 else 0 in
+ let src, gl = mkProt EConstr.mkProp EConstr.(mkApp (eq,[|t; c; mkRel rel|])) gl in
+ let concl = EConstr.mkArrow src (EConstr.Vars.lift 1 concl) in
+ let clr = if deps <> [] then clr else [] in
+ concl, gen_eq_tac, clr, gl
+ | _ -> concl, Tacticals.tclIDTAC, clr, gl in
+ let mk_lam t r = EConstr.mkLambda_or_LetIn r t in
+ let concl = List.fold_left mk_lam concl pred_rctx in
+ let gl, concl =
+ if eqid <> None && is_rec then
+ let gl, concls = pfe_type_of gl concl in
+ let concl, gl = mkProt concls concl gl in
+ let gl, _ = pfe_type_of gl concl in
+ gl, concl
+ else gl, concl in
+ concl, gen_eq_tac, clr, gl in
+ let gl, pty = pf_e_type_of gl elim_pred in
+ ppdebug(lazy Pp.(str"elim_pred=" ++ pp_term gl elim_pred));
+ ppdebug(lazy Pp.(str"elim_pred_ty=" ++ pp_term gl pty));
+ let gl = pf_unify_HO gl pred elim_pred in
+ let elim = fire_subst gl elim in
+ let gl, _ = pf_e_type_of gl elim in
+ (* check that the patterns do not contain non instantiated dependent metas *)
+ let () =
+ let evars_of_term = Evarutil.undefined_evars_of_term (project gl) in
+ let patterns = List.map (fun (_,_,t,_) -> fire_subst gl t) patterns in
+ let patterns_ev = List.map evars_of_term patterns in
+ let ev = List.fold_left Evar.Set.union Evar.Set.empty patterns_ev in
+ let ty_ev = Evar.Set.fold (fun i e ->
+ let ex = i in
+ let i_ty = EConstr.of_constr (Evd.evar_concl (Evd.find (project gl) ex)) in
+ Evar.Set.union e (evars_of_term i_ty))
+ ev Evar.Set.empty in
+ let inter = Evar.Set.inter ev ty_ev in
+ if not (Evar.Set.is_empty inter) then begin
+ let i = Evar.Set.choose inter in
+ let pat = List.find (fun t -> Evar.Set.mem i (evars_of_term t)) patterns in
+ errorstrm Pp.(str"Pattern"++spc()++pr_constr_pat (EConstr.Unsafe.to_constr pat)++spc()++
+ str"was not completely instantiated and one of its variables"++spc()++
+ str"occurs in the type of another non-instantiated pattern variable");
+ end
+ in
+ (* the elim tactic, with the eliminator and the predicated we computed *)
+ let elim = project gl, elim in
+ let elim_tac gl =
+ Tacticals.tclTHENLIST [refine_with ~with_evars:false elim; cleartac clr] gl in
+ Tacticals.tclTHENLIST [gen_eq_tac; elim_intro_tac ?ist what eqid elim_tac is_rec clr] orig_gl
+
+let no_intro ?ist what eqid elim_tac is_rec clr = elim_tac
+
+let elimtac x = ssrelim ~is_case:false [] (`EConstr ([],None,x)) None no_intro
+let casetac x = ssrelim ~is_case:true [] (`EConstr ([],None,x)) None no_intro
+
+let pf_nb_prod gl = nb_prod (project gl) (pf_concl gl)
+
+let rev_id = mk_internal_id "rev concl"
+let injecteq_id = mk_internal_id "injection equation"
+
+let revtoptac n0 gl =
+ let n = pf_nb_prod gl - n0 in
+ let dc, cl = EConstr.decompose_prod_n_assum (project gl) n (pf_concl gl) in
+ let dc' = dc @ [Context.Rel.Declaration.LocalAssum(Name rev_id, EConstr.it_mkProd_or_LetIn cl (List.rev dc))] in
+ let f = EConstr.it_mkLambda_or_LetIn (mkEtaApp (EConstr.mkRel (n + 1)) (-n) 1) dc' in
+ refine (EConstr.mkApp (f, [|Evarutil.mk_new_meta ()|])) gl
+
+let equality_inj l b id c gl =
+ let msg = ref "" in
+ try Proofview.V82.of_tactic (Equality.inj l b None c) gl
+ with
+ | Ploc.Exc(_,CErrors.UserError (_,s))
+ | CErrors.UserError (_,s)
+ when msg := Pp.string_of_ppcmds s;
+ !msg = "Not a projectable equality but a discriminable one." ||
+ !msg = "Nothing to inject." ->
+ Feedback.msg_warning (Pp.str !msg);
+ discharge_hyp (id, (id, "")) gl
+
+let injectidl2rtac id c gl =
+ Tacticals.tclTHEN (equality_inj None true id c) (revtoptac (pf_nb_prod gl)) gl
+
+let injectl2rtac sigma c = match EConstr.kind sigma c with
+| Var id -> injectidl2rtac id (EConstr.mkVar id, NoBindings)
+| _ ->
+ let id = injecteq_id in
+ let xhavetac id c = Proofview.V82.of_tactic (Tactics.pose_proof (Name id) c) in
+ Tacticals.tclTHENLIST [xhavetac id c; injectidl2rtac id (EConstr.mkVar id, NoBindings); Proofview.V82.of_tactic (Tactics.clear [id])]
+
+let is_injection_case c gl =
+ let gl, cty = pfe_type_of gl c in
+ let (mind,_), _ = pf_reduce_to_quantified_ind gl cty in
+ eq_gr (IndRef mind) (Coqlib.build_coq_eq ())
+
+let perform_injection c gl =
+ let gl, cty = pfe_type_of gl c in
+ let mind, t = pf_reduce_to_quantified_ind gl cty in
+ let dc, eqt = EConstr.decompose_prod (project gl) t in
+ if dc = [] then injectl2rtac (project gl) c gl else
+ if not (EConstr.Vars.closed0 (project gl) eqt) then
+ CErrors.user_err (Pp.str "can't decompose a quantified equality") else
+ let cl = pf_concl gl in let n = List.length dc in
+ let c_eq = mkEtaApp c n 2 in
+ let cl1 = EConstr.mkLambda EConstr.(Anonymous, mkArrow eqt cl, mkApp (mkRel 1, [|c_eq|])) in
+ let id = injecteq_id in
+ let id_with_ebind = (EConstr.mkVar id, NoBindings) in
+ let injtac = Tacticals.tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in
+ Tacticals.tclTHENLAST (Proofview.V82.of_tactic (Tactics.apply (EConstr.compose_lam dc cl1))) injtac gl
+
+let ssrscasetac force_inj c gl =
+ if force_inj || is_injection_case c gl then perform_injection c gl
+ else casetac c gl
diff --git a/plugins/ssr/ssrelim.mli b/plugins/ssr/ssrelim.mli
new file mode 100644
index 000000000..fb1b58ac3
--- /dev/null
+++ b/plugins/ssr/ssrelim.mli
@@ -0,0 +1,53 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Ssrmatching_plugin
+
+val ssrelim :
+ ?ind:(int * EConstr.types array) option ref ->
+ ?is_case:bool ->
+ ?ist:Ltac_plugin.Tacinterp.interp_sign ->
+ ((Ssrast.ssrhyps option * Ssrast.ssrocc) *
+ Ssrmatching.cpattern)
+ list ->
+ ([< `EConstr of
+ Ssrast.ssrhyp list * Ssrmatching.occ *
+ EConstr.constr &
+ 'b
+ | `EGen of
+ (Ssrast.ssrhyp list option *
+ Ssrmatching.occ) *
+ Ssrmatching.cpattern ]
+ as 'a) ->
+ ?elim:EConstr.constr ->
+ Ssrast.ssripat option ->
+ (?ist:Ltac_plugin.Tacinterp.interp_sign ->
+ 'a ->
+ Ssrast.ssripat option ->
+ (Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma) ->
+ bool -> Ssrast.ssrhyp list -> Proof_type.tactic) ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val elimtac :
+ EConstr.constr ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+val casetac :
+ EConstr.constr ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val is_injection_case : EConstr.t -> Proof_type.goal Evd.sigma -> bool
+val perform_injection :
+ EConstr.constr ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val ssrscasetac :
+ bool ->
+ EConstr.constr ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
new file mode 100644
index 000000000..af315aac5
--- /dev/null
+++ b/plugins/ssr/ssrequality.ml
@@ -0,0 +1,663 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Ltac_plugin
+open Util
+open Names
+open Vars
+open Locus
+open Printer
+open Globnames
+open Termops
+open Tacinterp
+open Term
+
+open Ssrmatching_plugin
+open Ssrmatching
+
+open Ssrast
+open Ssrprinters
+open Ssrcommon
+open Tacticals
+open Tacmach
+
+let ssroldreworder = Summary.ref ~name:"SSR:oldreworder" false
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "ssreflect 1.3 compatibility flag";
+ Goptions.optkey = ["SsrOldRewriteGoalsOrder"];
+ Goptions.optread = (fun _ -> !ssroldreworder);
+ Goptions.optdepr = false;
+ Goptions.optwrite = (fun b -> ssroldreworder := b) }
+
+(** The "simpl" tactic *)
+
+(* We must avoid zeta-converting any "let"s created by the "in" tactical. *)
+
+let tacred_simpl gl =
+ let simpl_expr =
+ Genredexpr.(
+ Simpl(Redops.make_red_flag[FBeta;FMatch;FZeta;FDeltaBut []],None)) in
+ let esimpl, _ = Redexpr.reduction_of_red_expr (pf_env gl) simpl_expr in
+ let esimpl e sigma c =
+ let (_,t) = esimpl e sigma c in
+ t in
+ let simpl env sigma c = (esimpl env sigma c) in
+ simpl
+
+let safe_simpltac n gl =
+ if n = ~-1 then
+ let cl= red_safe (tacred_simpl gl) (pf_env gl) (project gl) (pf_concl gl) in
+ Proofview.V82.of_tactic (convert_concl_no_check cl) gl
+ else
+ ssr_n_tac "simpl" n gl
+
+let simpltac = function
+ | Simpl n -> safe_simpltac n
+ | Cut n -> tclTRY (donetac n)
+ | SimplCut (n,m) -> tclTHEN (safe_simpltac m) (tclTRY (donetac n))
+ | Nop -> tclIDTAC
+
+(** The "congr" tactic *)
+
+let interp_congrarg_at ist gl n rf ty m =
+ ppdebug(lazy Pp.(str"===interp_congrarg_at==="));
+ let congrn, _ = mkSsrRRef "nary_congruence" in
+ let args1 = mkRnat n :: mkRHoles n @ [ty] in
+ let args2 = mkRHoles (3 * n) in
+ let rec loop i =
+ if i + n > m then None else
+ try
+ let rt = mkRApp congrn (args1 @ mkRApp rf (mkRHoles i) :: args2) in
+ ppdebug(lazy Pp.(str"rt=" ++ Printer.pr_glob_constr rt));
+ Some (interp_refine ist gl rt)
+ with _ -> loop (i + 1) in
+ loop 0
+
+let pattern_id = mk_internal_id "pattern value"
+
+let congrtac ((n, t), ty) ist gl =
+ ppdebug(lazy (Pp.str"===congr==="));
+ ppdebug(lazy Pp.(str"concl=" ++ Printer.pr_econstr (Tacmach.pf_concl gl)));
+ let sigma, _ as it = interp_term ist gl t in
+ let gl = pf_merge_uc_of sigma gl in
+ let _, f, _, _ucst = pf_abs_evars gl it in
+ let ist' = {ist with lfun =
+ Id.Map.add pattern_id (Tacinterp.Value.of_constr f) Id.Map.empty } in
+ let rf = mkRltacVar pattern_id in
+ let m = pf_nbargs gl f in
+ let _, cf = if n > 0 then
+ match interp_congrarg_at ist' gl n rf ty m with
+ | Some cf -> cf
+ | None -> errorstrm Pp.(str "No " ++ int n ++ str "-congruence with "
+ ++ pr_term t)
+ else let rec loop i =
+ if i > m then errorstrm Pp.(str "No congruence with " ++ pr_term t)
+ else match interp_congrarg_at ist' gl i rf ty m with
+ | Some cf -> cf
+ | None -> loop (i + 1) in
+ loop 1 in
+ tclTHEN (refine_with cf) (tclTRY (Proofview.V82.of_tactic Tactics.reflexivity)) gl
+
+let newssrcongrtac arg ist gl =
+ ppdebug(lazy Pp.(str"===newcongr==="));
+ ppdebug(lazy Pp.(str"concl=" ++ Printer.pr_econstr (pf_concl gl)));
+ (* utils *)
+ let fs gl t = Reductionops.nf_evar (project gl) t in
+ let tclMATCH_GOAL (c, gl_c) proj t_ok t_fail gl =
+ match try Some (pf_unify_HO gl_c (pf_concl gl) c) with _ -> None with
+ | Some gl_c ->
+ tclTHEN (Proofview.V82.of_tactic (convert_concl (fs gl_c c)))
+ (t_ok (proj gl_c)) gl
+ | None -> t_fail () gl in
+ let mk_evar gl ty =
+ let env, sigma, si = pf_env gl, project gl, sig_it gl in
+ let sigma = Evd.create_evar_defs sigma in
+ let (sigma, x) = Evarutil.new_evar env sigma ty in
+ x, re_sig si sigma in
+ let arr, gl = pf_mkSsrConst "ssr_congr_arrow" gl in
+ let ssr_congr lr = EConstr.mkApp (arr, lr) in
+ (* here thw two cases: simple equality or arrow *)
+ let equality, _, eq_args, gl' =
+ let eq, gl = pf_fresh_global (Coqlib.build_coq_eq ()) gl in
+ pf_saturate gl (EConstr.of_constr eq) 3 in
+ tclMATCH_GOAL (equality, gl') (fun gl' -> fs gl' (List.assoc 0 eq_args))
+ (fun ty -> congrtac (arg, Detyping.detype false [] (pf_env gl) (project gl) ty) ist)
+ (fun () ->
+ let lhs, gl' = mk_evar gl EConstr.mkProp in let rhs, gl' = mk_evar gl' EConstr.mkProp in
+ let arrow = EConstr.mkArrow lhs (EConstr.Vars.lift 1 rhs) in
+ tclMATCH_GOAL (arrow, gl') (fun gl' -> [|fs gl' lhs;fs gl' rhs|])
+ (fun lr -> tclTHEN (Proofview.V82.of_tactic (Tactics.apply (ssr_congr lr))) (congrtac (arg, mkRType) ist))
+ (fun _ _ -> errorstrm Pp.(str"Conclusion is not an equality nor an arrow")))
+ gl
+
+(** 7. Rewriting tactics (rewrite, unlock) *)
+
+(** Coq rewrite compatibility flag *)
+
+let ssr_strict_match = ref false
+
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "strict redex matching";
+ Goptions.optkey = ["Match"; "Strict"];
+ Goptions.optread = (fun () -> !ssr_strict_match);
+ Goptions.optdepr = false;
+ Goptions.optwrite = (fun b -> ssr_strict_match := b) }
+
+(** Rewrite rules *)
+
+type ssrwkind = RWred of ssrsimpl | RWdef | RWeq
+type ssrrule = ssrwkind * ssrterm
+
+(** Rewrite arguments *)
+
+type ssrrwarg = (ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule)
+
+let notimes = 0
+let nomult = 1, Once
+
+let mkocc occ = None, occ
+let noclr = mkocc None
+let mkclr clr = Some clr, None
+let nodocc = mkclr []
+
+let is_rw_cut = function RWred (Cut _) -> true | _ -> false
+
+let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) : ssrrwarg =
+ if rt <> RWeq then begin
+ if rt = RWred Nop && not (m = nomult && occ = None && rx = None)
+ && (clr = None || clr = Some []) then
+ anomaly "Improper rewrite clear switch";
+ if d = R2L && rt <> RWdef then
+ CErrors.user_err (Pp.str "Right-to-left switch on simplification");
+ if n <> 1 && is_rw_cut rt then
+ CErrors.user_err (Pp.str "Bad or useless multiplier");
+ if occ <> None && rx = None && rt <> RWdef then
+ CErrors.user_err (Pp.str "Missing redex for simplification occurrence")
+ end; (d, m), ((docc, rx), r)
+
+let norwmult = L2R, nomult
+let norwocc = noclr, None
+
+let simplintac occ rdx sim gl =
+ let simptac m gl =
+ if m <> ~-1 then
+ CErrors.user_err (Pp.str "Localized custom simpl tactic not supported");
+ let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in
+ let simp env c _ _ = EConstr.Unsafe.to_constr (red_safe Tacred.simpl env sigma0 (EConstr.of_constr c)) in
+ Proofview.V82.of_tactic
+ (convert_concl_no_check (EConstr.of_constr (eval_pattern env0 sigma0 (EConstr.Unsafe.to_constr concl0) rdx occ simp)))
+ gl in
+ match sim with
+ | Simpl m -> simptac m gl
+ | SimplCut (n,m) -> tclTHEN (simptac m) (tclTRY (donetac n)) gl
+ | _ -> simpltac sim gl
+
+let rec get_evalref sigma c = match EConstr.kind sigma c with
+ | Var id -> EvalVarRef id
+ | Const (k,_) -> EvalConstRef k
+ | App (c', _) -> get_evalref sigma c'
+ | Cast (c', _, _) -> get_evalref sigma c'
+ | Proj(c,_) -> EvalConstRef(Projection.constant c)
+ | _ -> errorstrm Pp.(str "The term " ++ pr_constr_pat (EConstr.Unsafe.to_constr c) ++ str " is not unfoldable")
+
+(* Strip a pattern generated by a prenex implicit to its constant. *)
+let strip_unfold_term _ ((sigma, t) as p) kt = match EConstr.kind sigma t with
+ | App (f, a) when kt = xNoFlag && Array.for_all (EConstr.isEvar sigma) a && EConstr.isConst sigma f ->
+ (sigma, f), true
+ | Const _ | Var _ -> p, true
+ | Proj _ -> p, true
+ | _ -> p, false
+
+let same_proj sigma t1 t2 =
+ match EConstr.kind sigma t1, EConstr.kind sigma t2 with
+ | Proj(c1,_), Proj(c2, _) -> Projection.equal c1 c2
+ | _ -> false
+
+let all_ok _ _ = true
+
+let fake_pmatcher_end () =
+ mkProp, L2R, (Evd.empty, Evd.empty_evar_universe_context, mkProp)
+
+let unfoldintac occ rdx t (kt,_) gl =
+ let fs sigma x = Reductionops.nf_evar sigma x in
+ let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in
+ let (sigma, t), const = strip_unfold_term env0 t kt in
+ let body env t c =
+ Tacred.unfoldn [AllOccurrences, get_evalref sigma t] env sigma0 c in
+ let easy = occ = None && rdx = None in
+ let red_flags = if easy then CClosure.betaiotazeta else CClosure.betaiota in
+ let beta env = Reductionops.clos_norm_flags red_flags env sigma0 in
+ let unfold, conclude = match rdx with
+ | Some (_, (In_T _ | In_X_In_T _)) | None ->
+ let ise = Evd.create_evar_defs sigma in
+ let ise, u = mk_tpattern env0 sigma0 (ise,EConstr.Unsafe.to_constr t) all_ok L2R (EConstr.Unsafe.to_constr t) in
+ let find_T, end_T =
+ mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in
+ (fun env c _ h ->
+ try find_T env c h ~k:(fun env c _ _ -> EConstr.Unsafe.to_constr (body env t (EConstr.of_constr c)))
+ with NoMatch when easy -> c
+ | NoMatch | NoProgress -> errorstrm Pp.(str"No occurrence of "
+ ++ pr_constr_pat (EConstr.Unsafe.to_constr t) ++ spc() ++ str "in " ++ Printer.pr_constr c)),
+ (fun () -> try end_T () with
+ | NoMatch when easy -> fake_pmatcher_end ()
+ | NoMatch -> anomaly "unfoldintac")
+ | _ ->
+ (fun env (c as orig_c) _ h ->
+ if const then
+ let rec aux c =
+ match EConstr.kind sigma0 c with
+ | Const _ when EConstr.eq_constr sigma0 c t -> body env t t
+ | App (f,a) when EConstr.eq_constr sigma0 f t -> EConstr.mkApp (body env f f,a)
+ | Proj _ when same_proj sigma0 c t -> body env t c
+ | _ ->
+ let c = Reductionops.whd_betaiotazeta sigma0 c in
+ match EConstr.kind sigma0 c with
+ | Const _ when EConstr.eq_constr sigma0 c t -> body env t t
+ | App (f,a) when EConstr.eq_constr sigma0 f t -> EConstr.mkApp (body env f f,a)
+ | Proj _ when same_proj sigma0 c t -> body env t c
+ | Const f -> aux (body env c c)
+ | App (f, a) -> aux (EConstr.mkApp (body env f f, a))
+ | _ -> errorstrm Pp.(str "The term "++pr_constr orig_c++
+ str" contains no " ++ pr_econstr t ++ str" even after unfolding")
+ in EConstr.Unsafe.to_constr @@ aux (EConstr.of_constr c)
+ else
+ try EConstr.Unsafe.to_constr @@ body env t (fs (unify_HO env sigma (EConstr.of_constr c) t) t)
+ with _ -> errorstrm Pp.(str "The term " ++
+ pr_constr c ++spc()++ str "does not unify with " ++ pr_constr_pat (EConstr.Unsafe.to_constr t))),
+ fake_pmatcher_end in
+ let concl =
+ let concl0 = EConstr.Unsafe.to_constr concl0 in
+ try beta env0 (EConstr.of_constr (eval_pattern env0 sigma0 concl0 rdx occ unfold))
+ with Option.IsNone -> errorstrm Pp.(str"Failed to unfold " ++ pr_constr_pat (EConstr.Unsafe.to_constr t)) in
+ let _ = conclude () in
+ Proofview.V82.of_tactic (convert_concl concl) gl
+;;
+
+let foldtac occ rdx ft gl =
+ let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in
+ let sigma, t = ft in
+ let t = EConstr.to_constr sigma t in
+ let fold, conclude = match rdx with
+ | Some (_, (In_T _ | In_X_In_T _)) | None ->
+ let ise = Evd.create_evar_defs sigma in
+ let ut = EConstr.Unsafe.to_constr (red_product_skip_id env0 sigma (EConstr.of_constr t)) in
+ let ise, ut = mk_tpattern env0 sigma0 (ise,t) all_ok L2R ut in
+ let find_T, end_T =
+ mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[ut]) in
+ (fun env c _ h -> try find_T env c h ~k:(fun env t _ _ -> t) with NoMatch ->c),
+ (fun () -> try end_T () with NoMatch -> fake_pmatcher_end ())
+ | _ ->
+ (fun env c _ h -> try let sigma = unify_HO env sigma (EConstr.of_constr c) (EConstr.of_constr t) in EConstr.to_constr sigma (EConstr.of_constr t)
+ with _ -> errorstrm Pp.(str "fold pattern " ++ pr_constr_pat t ++ spc ()
+ ++ str "does not match redex " ++ pr_constr_pat c)),
+ fake_pmatcher_end in
+ let concl0 = EConstr.Unsafe.to_constr concl0 in
+ let concl = eval_pattern env0 sigma0 concl0 rdx occ fold in
+ let _ = conclude () in
+ Proofview.V82.of_tactic (convert_concl (EConstr.of_constr concl)) gl
+;;
+
+let converse_dir = function L2R -> R2L | R2L -> L2R
+
+let rw_progress rhs lhs ise = not (EConstr.eq_constr ise lhs (Evarutil.nf_evar ise rhs))
+
+(* Coq has a more general form of "equation" (any type with a single *)
+(* constructor with no arguments with_rect_r elimination lemmas). *)
+(* However there is no clear way of determining the LHS and RHS of *)
+(* such a generic Leibnitz equation -- short of inspecting the type *)
+(* of the elimination lemmas. *)
+
+let rec strip_prod_assum c = match Term.kind_of_term c with
+ | Prod (_, _, c') -> strip_prod_assum c'
+ | LetIn (_, v, _, c') -> strip_prod_assum (subst1 v c)
+ | Cast (c', _, _) -> strip_prod_assum c'
+ | _ -> c
+
+let rule_id = mk_internal_id "rewrite rule"
+
+exception PRtype_error
+
+let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl =
+(* ppdebug(lazy(str"sigma@pirrel_rewrite=" ++ pr_evar_map None sigma)); *)
+ let env = pf_env gl in
+ let beta = Reductionops.clos_norm_flags CClosure.beta env sigma in
+ let sigma, p =
+ let sigma = Evd.create_evar_defs sigma in
+ let (sigma, ev) = Evarutil.new_evar env sigma (beta (EConstr.Vars.subst1 new_rdx pred)) in
+ (sigma, ev)
+ in
+ let pred = EConstr.mkNamedLambda pattern_id rdx_ty pred in
+ let elim, gl =
+ let ((kn, i) as ind, _), unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in
+ let sort = elimination_sort_of_goal gl in
+ 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
+ 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
+ (* We check the proof is well typed *)
+ let sigma, proof_ty =
+ try Typing.type_of env sigma proof with _ -> raise PRtype_error in
+ ppdebug(lazy Pp.(str"pirrel_rewrite proof term of type: " ++ pr_econstr proof_ty));
+ try refine_with
+ ~first_goes_last:(not !ssroldreworder) ~with_evars:false (sigma, proof) gl
+ with _ ->
+ (* we generate a msg like: "Unable to find an instance for the variable" *)
+ let hd_ty, miss = match EConstr.kind sigma c with
+ | App (hd, args) ->
+ let hd_ty = Retyping.get_type_of env sigma hd in
+ let names = let rec aux t = function 0 -> [] | n ->
+ let t = Reductionops.whd_all env sigma t in
+ match EConstr.kind_of_type sigma t with
+ | ProdType (name, _, t) -> name :: aux t (n-1)
+ | _ -> assert false in aux hd_ty (Array.length args) in
+ hd_ty, Util.List.map_filter (fun (t, name) ->
+ let evs = Evar.Set.elements (Evarutil.undefined_evars_of_term sigma t) in
+ let open_evs = List.filter (fun k ->
+ Sorts.InProp <> Retyping.get_sort_family_of
+ env sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k))))
+ evs in
+ if open_evs <> [] then Some name else None)
+ (List.combine (Array.to_list args) names)
+ | _ -> anomaly "rewrite rule not an application" in
+ errorstrm Pp.(Himsg.explain_refiner_error (Logic.UnresolvedBindings miss)++
+ (Pp.fnl()++str"Rule's type:" ++ spc() ++ pr_econstr hd_ty))
+;;
+
+let is_construct_ref sigma c r =
+ EConstr.isConstruct sigma c && eq_gr (ConstructRef (fst(EConstr.destConstruct sigma c))) r
+let is_ind_ref sigma c r = EConstr.isInd sigma c && eq_gr (IndRef (fst(EConstr.destInd sigma c))) r
+
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+
+let rwcltac cl rdx dir sr gl =
+ let n, r_n,_, ucst = pf_abs_evars gl sr in
+ let r_n' = pf_abs_cterm gl n r_n in
+ let r' = EConstr.Vars.subst_var pattern_id r_n' in
+ let gl = pf_unsafe_merge_uc ucst gl in
+ let rdxt = Retyping.get_type_of (pf_env gl) (fst sr) rdx in
+(* ppdebug(lazy(str"sigma@rwcltac=" ++ pr_evar_map None (fst sr))); *)
+ ppdebug(lazy Pp.(str"r@rwcltac=" ++ pr_econstr (snd sr)));
+ let cvtac, rwtac, gl =
+ if EConstr.Vars.closed0 (project gl) r' then
+ let env, sigma, c, c_eq = pf_env gl, fst sr, snd sr, Coqlib.build_coq_eq () in
+ let sigma, c_ty = Typing.type_of env sigma c in
+ ppdebug(lazy Pp.(str"c_ty@rwcltac=" ++ pr_econstr c_ty));
+ match EConstr.kind_of_type sigma (Reductionops.whd_all env sigma c_ty) with
+ | AtomicType(e, a) when is_ind_ref sigma e c_eq ->
+ let new_rdx = if dir = L2R then a.(2) else a.(1) in
+ pirrel_rewrite cl rdx rdxt new_rdx dir (sigma,c) c_ty, tclIDTAC, gl
+ | _ ->
+ let cl' = EConstr.mkApp (EConstr.mkNamedLambda pattern_id rdxt cl, [|rdx|]) in
+ let sigma, _ = Typing.type_of env sigma cl' in
+ let gl = pf_merge_uc_of sigma gl in
+ Proofview.V82.of_tactic (convert_concl cl'), rewritetac dir r', gl
+ else
+ let dc, r2 = EConstr.decompose_lam_n_assum (project gl) n r' in
+ let r3, _, r3t =
+ try EConstr.destCast (project gl) r2 with _ ->
+ errorstrm Pp.(str "no cast from " ++ pr_constr_pat (EConstr.Unsafe.to_constr (snd sr))
+ ++ str " to " ++ pr_econstr r2) in
+ let cl' = EConstr.mkNamedProd rule_id (EConstr.it_mkProd_or_LetIn r3t dc) (EConstr.Vars.lift 1 cl) in
+ let cl'' = EConstr.mkNamedProd pattern_id rdxt cl' in
+ let itacs = [introid pattern_id; introid rule_id] in
+ let cltac = Proofview.V82.of_tactic (Tactics.clear [pattern_id; rule_id]) in
+ let rwtacs = [rewritetac dir (EConstr.mkVar rule_id); cltac] in
+ apply_type cl'' [rdx; EConstr.it_mkLambda_or_LetIn r3 dc], tclTHENLIST (itacs @ rwtacs), gl
+ in
+ let cvtac' _ =
+ try cvtac gl with
+ | PRtype_error ->
+ if occur_existential (project gl) (Tacmach.pf_concl gl)
+ then errorstrm Pp.(str "Rewriting impacts evars")
+ else errorstrm Pp.(str "Dependent type error in rewrite of "
+ ++ pr_constr_env (pf_env gl) (project gl) (Term.mkNamedLambda pattern_id (EConstr.Unsafe.to_constr rdxt) (EConstr.Unsafe.to_constr cl)))
+ | CErrors.UserError _ as e -> raise e
+ | e -> anomaly ("cvtac's exception: " ^ Printexc.to_string e);
+ in
+ tclTHEN cvtac' rwtac gl
+
+let prof_rwcltac = mk_profiler "rwrxtac.rwcltac";;
+let rwcltac cl rdx dir sr gl =
+ prof_rwcltac.profile (rwcltac cl rdx dir sr) gl
+;;
+
+
+let lz_coq_prod =
+ let prod = lazy (Coqlib.build_prod ()) in fun () -> Lazy.force prod
+
+let lz_setoid_relation =
+ let sdir = ["Classes"; "RelationClasses"] in
+ let last_srel = ref (Environ.empty_env, None) in
+ fun env -> match !last_srel with
+ | env', srel when env' == env -> srel
+ | _ ->
+ let srel =
+ try Some (Universes.constr_of_global @@
+ Coqlib.coq_reference "Class_setoid" sdir "RewriteRelation")
+ with _ -> None in
+ last_srel := (env, srel); srel
+
+let ssr_is_setoid env =
+ match lz_setoid_relation env with
+ | None -> fun _ _ _ -> false
+ | Some srel ->
+ fun sigma r args ->
+ Rewrite.is_applied_rewrite_relation env
+ sigma [] (EConstr.mkApp (r, args)) <> None
+
+let prof_rwxrtac_find_rule = mk_profiler "rwrxtac.find_rule";;
+
+let closed0_check cl p gl =
+ if closed0 cl then
+ errorstrm Pp.(str"No occurrence of redex "++ pr_constr_env (pf_env gl) (project gl) p)
+
+let dir_org = function L2R -> 1 | R2L -> 2
+
+let rwprocess_rule dir rule gl =
+ let env = pf_env gl in
+ let coq_prod = lz_coq_prod () in
+ let is_setoid = ssr_is_setoid env in
+ let r_sigma, rules =
+ let rec loop d sigma r t0 rs red =
+ let t =
+ if red = 1 then Tacred.hnf_constr env sigma t0
+ else Reductionops.whd_betaiotazeta sigma t0 in
+ ppdebug(lazy Pp.(str"rewrule="++pr_constr_pat (EConstr.Unsafe.to_constr t)));
+ match EConstr.kind sigma t with
+ | Prod (_, xt, at) ->
+ let sigma = Evd.create_evar_defs sigma in
+ let (sigma, x) = Evarutil.new_evar env sigma xt in
+ loop d sigma EConstr.(mkApp (r, [|x|])) (EConstr.Vars.subst1 x at) rs 0
+ | App (pr, a) when is_ind_ref sigma pr coq_prod.Coqlib.typ ->
+ let sr sigma = match EConstr.kind sigma (Tacred.hnf_constr env sigma r) with
+ | App (c, ra) when is_construct_ref sigma c coq_prod.Coqlib.intro ->
+ fun i -> ra.(i + 1), sigma
+ | _ -> let ra = Array.append a [|r|] in
+ function 1 ->
+ let sigma, pi1 = Evd.fresh_global env sigma coq_prod.Coqlib.proj1 in
+ EConstr.mkApp (EConstr.of_constr pi1, ra), sigma
+ | _ ->
+ let sigma, pi2 = Evd.fresh_global env sigma coq_prod.Coqlib.proj2 in
+ EConstr.mkApp (EConstr.of_constr pi2, ra), sigma in
+ if EConstr.eq_constr sigma a.(0) (EConstr.of_constr (Universes.constr_of_global @@ Coqlib.build_coq_True ())) then
+ let s, sigma = sr sigma 2 in
+ loop (converse_dir d) sigma s a.(1) rs 0
+ else
+ let s, sigma = sr sigma 2 in
+ let sigma, rs2 = loop d sigma s a.(1) rs 0 in
+ let s, sigma = sr sigma 1 in
+ loop d sigma s a.(0) rs2 0
+ | App (r_eq, a) when Hipattern.match_with_equality_type sigma t != None ->
+ let (ind, u) = EConstr.destInd sigma r_eq and rhs = Array.last a in
+ let np = Inductiveops.inductive_nparamdecls ind in
+ let indu = (ind, EConstr.EInstance.kind sigma u) in
+ let ind_ct = Inductiveops.type_of_constructors env indu in
+ let lhs0 = last_arg sigma (EConstr.of_constr (strip_prod_assum ind_ct.(0))) in
+ let rdesc = match EConstr.kind sigma lhs0 with
+ | Rel i ->
+ let lhs = a.(np - i) in
+ let lhs, rhs = if d = L2R then lhs, rhs else rhs, lhs in
+(* msgnl (str "RW: " ++ pr_rwdir d ++ str " " ++ pr_constr_pat r ++ str " : "
+ ++ pr_constr_pat lhs ++ str " ~> " ++ pr_constr_pat rhs); *)
+ d, r, lhs, rhs
+(*
+ let l_i, r_i = if d = L2R then i, 1 - ndep else 1 - ndep, i in
+ let lhs = a.(np - l_i) and rhs = a.(np - r_i) in
+ let a' = Array.copy a in let _ = a'.(np - l_i) <- mkVar pattern_id in
+ let r' = mkCast (r, DEFAULTcast, mkApp (r_eq, a')) in
+ (d, r', lhs, rhs)
+*)
+ | _ ->
+ let lhs = EConstr.Vars.substl (array_list_of_tl (Array.sub a 0 np)) lhs0 in
+ let lhs, rhs = if d = R2L then lhs, rhs else rhs, lhs in
+ let d' = if Array.length a = 1 then d else converse_dir d in
+ d', r, lhs, rhs in
+ sigma, rdesc :: rs
+ | App (s_eq, a) when is_setoid sigma s_eq a ->
+ let np = Array.length a and i = 3 - dir_org d in
+ let lhs = a.(np - i) and rhs = a.(np + i - 3) in
+ let a' = Array.copy a in let _ = a'.(np - i) <- EConstr.mkVar pattern_id in
+ let r' = EConstr.mkCast (r, DEFAULTcast, EConstr.mkApp (s_eq, a')) in
+ sigma, (d, r', lhs, rhs) :: rs
+ | _ ->
+ if red = 0 then loop d sigma r t rs 1
+ else errorstrm Pp.(str "not a rewritable relation: " ++ pr_constr_pat (EConstr.Unsafe.to_constr t)
+ ++ spc() ++ str "in rule " ++ pr_constr_pat (EConstr.Unsafe.to_constr (snd rule)))
+ in
+ let sigma, r = rule in
+ let t = Retyping.get_type_of env sigma r in
+ loop dir sigma r t [] 0
+ in
+ r_sigma, rules
+
+let rwrxtac occ rdx_pat dir rule gl =
+ let env = pf_env gl in
+ let r_sigma, rules = rwprocess_rule dir rule gl in
+ let find_rule rdx =
+ let rec rwtac = function
+ | [] ->
+ errorstrm Pp.(str "pattern " ++ pr_constr_pat (EConstr.Unsafe.to_constr rdx) ++
+ str " does not match " ++ pr_dir_side dir ++
+ str " of " ++ pr_constr_pat (EConstr.Unsafe.to_constr (snd rule)))
+ | (d, r, lhs, rhs) :: rs ->
+ try
+ let ise = unify_HO env (Evd.create_evar_defs r_sigma) lhs rdx in
+ if not (rw_progress rhs rdx ise) then raise NoMatch else
+ d, (ise, Evd.evar_universe_context ise, Reductionops.nf_evar ise r)
+ with _ -> rwtac rs in
+ rwtac rules in
+ let find_rule rdx = prof_rwxrtac_find_rule.profile find_rule rdx in
+ let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in
+ let find_R, conclude = match rdx_pat with
+ | Some (_, (In_T _ | In_X_In_T _)) | None ->
+ let upats_origin = dir, EConstr.Unsafe.to_constr (snd rule) in
+ let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) =
+ let sigma, pat =
+ let rw_progress rhs t evd = rw_progress rhs (EConstr.of_constr t) evd in
+ mk_tpattern env sigma0 (sigma,EConstr.to_constr sigma r) (rw_progress rhs) d (EConstr.to_constr sigma lhs) in
+ sigma, pats @ [pat] in
+ let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in
+ let find_R, end_R = mk_tpattern_matcher sigma0 occ ~upats_origin rpats in
+ (fun e c _ i -> find_R ~k:(fun _ _ _ h -> mkRel h) e c i),
+ fun cl -> let rdx,d,r = end_R () in closed0_check cl rdx gl; (d,r),rdx
+ | Some(_, (T e | X_In_T (_,e) | E_As_X_In_T (e,_,_) | E_In_X_In_T (e,_,_))) ->
+ let r = ref None in
+ (fun env c _ h -> do_once r (fun () -> find_rule (EConstr.of_constr c), c); mkRel h),
+ (fun concl -> closed0_check concl e gl;
+ let (d,(ev,ctx,c)) , x = assert_done r in (d,(ev,ctx, EConstr.to_constr ev c)) , x) in
+ let concl0 = EConstr.Unsafe.to_constr concl0 in
+ let concl = eval_pattern env0 sigma0 concl0 rdx_pat occ find_R in
+ let (d, r), rdx = conclude concl in
+ let r = Evd.merge_universe_context (pi1 r) (pi2 r), EConstr.of_constr (pi3 r) in
+ rwcltac (EConstr.of_constr concl) (EConstr.of_constr rdx) d r gl
+;;
+
+let prof_rwxrtac = mk_profiler "rwrxtac";;
+let rwrxtac occ rdx_pat dir rule gl =
+ prof_rwxrtac.profile (rwrxtac occ rdx_pat dir rule) gl
+;;
+
+let ssrinstancesofrule ist dir arg gl =
+ let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in
+ let rule = interp_term ist gl arg in
+ let r_sigma, rules = rwprocess_rule dir rule gl in
+ let find, conclude =
+ let upats_origin = dir, EConstr.Unsafe.to_constr (snd rule) in
+ let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) =
+ let sigma, pat =
+ let rw_progress rhs t evd = rw_progress rhs (EConstr.of_constr t) evd in
+ mk_tpattern env sigma0 (sigma,EConstr.to_constr sigma r) (rw_progress rhs) d (EConstr.to_constr sigma lhs) in
+ sigma, pats @ [pat] in
+ let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in
+ mk_tpattern_matcher ~all_instances:true ~raise_NoMatch:true sigma0 None ~upats_origin rpats in
+ let print env p c _ = Feedback.msg_info Pp.(hov 1 (str"instance:" ++ spc() ++ pr_constr p ++ spc() ++ str "matches:" ++ spc() ++ pr_constr c)); c in
+ Feedback.msg_info Pp.(str"BEGIN INSTANCES");
+ try
+ while true do
+ ignore(find env0 (EConstr.Unsafe.to_constr concl0) 1 ~k:print)
+ done; raise NoMatch
+ with NoMatch -> Feedback.msg_info Pp.(str"END INSTANCES"); tclIDTAC gl
+
+let ipat_rewrite occ dir c gl = rwrxtac occ None dir (project gl, c) gl
+
+let rwargtac ist ((dir, mult), (((oclr, occ), grx), (kind, gt))) gl =
+ let fail = ref false in
+ let interp_rpattern ist gl gc =
+ try interp_rpattern ist gl gc
+ with _ when snd mult = May -> fail := true; project gl, T mkProp in
+ let interp gc gl =
+ try interp_term ist gl gc
+ with _ when snd mult = May -> fail := true; (project gl, EConstr.mkProp) in
+ let rwtac gl =
+ let rx = Option.map (interp_rpattern ist gl) grx in
+ let t = interp gt gl in
+ (match kind with
+ | RWred sim -> simplintac occ rx sim
+ | RWdef -> if dir = R2L then foldtac occ rx t else unfoldintac occ rx t gt
+ | RWeq -> rwrxtac occ rx dir t) gl in
+ let ctac = cleartac (interp_clr (project gl) (oclr, (fst gt, snd (interp gt gl)))) in
+ if !fail then ctac gl else tclTHEN (tclMULT mult rwtac) ctac gl
+
+(** Rewrite argument sequence *)
+
+(* type ssrrwargs = ssrrwarg list *)
+
+(** The "rewrite" tactic *)
+
+let ssrrewritetac ist rwargs =
+ tclTHENLIST (List.map (rwargtac ist) rwargs)
+
+(** The "unlock" tactic *)
+
+let unfoldtac occ ko t kt gl =
+ let env = pf_env gl in
+ let cl, c = pf_fill_occ_term gl occ (fst (strip_unfold_term env t kt)) in
+ let cl' = EConstr.Vars.subst1 (pf_unfoldn [OnlyOccurrences [1], get_evalref (project gl) c] gl c) cl in
+ let f = if ko = None then CClosure.betaiotazeta else CClosure.betaiota in
+ Proofview.V82.of_tactic
+ (convert_concl (pf_reduce (Reductionops.clos_norm_flags f) gl cl')) gl
+
+let unlocktac ist args gl =
+ let utac (occ, gt) gl =
+ unfoldtac occ occ (interp_term ist gl gt) (fst gt) gl in
+ let locked, gl = pf_mkSsrConst "locked" gl in
+ let key, gl = pf_mkSsrConst "master_key" gl in
+ let ktacs = [
+ (fun gl -> unfoldtac None None (project gl,locked) xInParens gl);
+ Ssrelim.casetac key ] in
+ tclTHENLIST (List.map utac args @ ktacs) gl
+
diff --git a/plugins/ssr/ssrequality.mli b/plugins/ssr/ssrequality.mli
new file mode 100644
index 000000000..9c5fd4983
--- /dev/null
+++ b/plugins/ssr/ssrequality.mli
@@ -0,0 +1,62 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Ssrmatching_plugin
+open Ssrast
+
+type ssrwkind = RWred of ssrsimpl | RWdef | RWeq
+type ssrrule = ssrwkind * ssrterm
+type ssrrwarg = (ssrdir * ssrmult) * ((ssrdocc * Ssrmatching.rpattern option) * ssrrule)
+
+val dir_org : ssrdir -> int
+
+val notimes : int
+val nomult : ssrmult
+val mkocc : ssrocc -> ssrdocc
+val mkclr : ssrclear -> ssrdocc
+val nodocc : ssrdocc
+val noclr : ssrdocc
+
+val simpltac : Ssrast.ssrsimpl -> Proof_type.tactic
+
+val newssrcongrtac :
+ int * Ssrast.ssrterm ->
+ Ltac_plugin.Tacinterp.interp_sign ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+
+val mk_rwarg :
+ Ssrast.ssrdir * (int * Ssrast.ssrmmod) ->
+ (Ssrast.ssrclear option * Ssrast.ssrocc) * Ssrmatching.rpattern option ->
+ ssrwkind * Ssrast.ssrterm -> ssrrwarg
+
+val norwmult : ssrdir * ssrmult
+val norwocc : (Ssrast.ssrclear option * Ssrast.ssrocc) * Ssrmatching.rpattern option
+
+val ssrinstancesofrule :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ Ssrast.ssrdir ->
+ Ssrast.ssrterm ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val ssrrewritetac :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ ((Ssrast.ssrdir * (int * Ssrast.ssrmmod)) *
+ (((Ssrast.ssrhyps option * Ssrmatching.occ) *
+ Ssrmatching.rpattern option) *
+ (ssrwkind * Ssrast.ssrterm)))
+ list -> Proof_type.tactic
+
+val ipat_rewrite : ssrocc -> ssrdir -> EConstr.t -> Proof_type.tactic
+
+val unlocktac :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ (Ssrmatching.occ * Ssrast.ssrterm) list ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
diff --git a/plugins/ssr/ssrfun.v b/plugins/ssr/ssrfun.v
new file mode 100644
index 000000000..1f3a9c124
--- /dev/null
+++ b/plugins/ssr/ssrfun.v
@@ -0,0 +1,791 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+Require Import ssreflect.
+
+(******************************************************************************)
+(* This file contains the basic definitions and notations for working with *)
+(* functions. The definitions provide for: *)
+(* *)
+(* - Pair projections: *)
+(* p.1 == first element of a pair *)
+(* p.2 == second element of a pair *)
+(* These notations also apply to p : P /\ Q, via an and >-> pair coercion. *)
+(* *)
+(* - Simplifying functions, beta-reduced by /= and simpl: *)
+(* [fun : T => E] == constant function from type T that returns E *)
+(* [fun x => E] == unary function *)
+(* [fun x : T => E] == unary function with explicit domain type *)
+(* [fun x y => E] == binary function *)
+(* [fun x y : T => E] == binary function with common domain type *)
+(* [fun (x : T) y => E] \ *)
+(* [fun (x : xT) (y : yT) => E] | == binary function with (some) explicit, *)
+(* [fun x (y : T) => E] / independent domain types for each argument *)
+(* *)
+(* - Partial functions using option type: *)
+(* oapp f d ox == if ox is Some x returns f x, d otherwise *)
+(* odflt d ox == if ox is Some x returns x, d otherwise *)
+(* obind f ox == if ox is Some x returns f x, None otherwise *)
+(* omap f ox == if ox is Some x returns Some (f x), None otherwise *)
+(* *)
+(* - Singleton types: *)
+(* all_equal_to x0 == x0 is the only value in its type, so any such value *)
+(* can be rewritten to x0. *)
+(* *)
+(* - A generic wrapper type: *)
+(* wrapped T == the inductive type with values Wrap x for x : T. *)
+(* unwrap w == the projection of w : wrapped T on T. *)
+(* wrap x == the canonical injection of x : T into wrapped T; it is *)
+(* equivalent to Wrap x, but is declared as a (default) *)
+(* Canonical Structure, which lets the Coq HO unification *)
+(* automatically expand x into unwrap (wrap x). The delta *)
+(* reduction of wrap x to Wrap can be exploited to *)
+(* introduce controlled nondeterminism in Canonical *)
+(* Structure inference, as in the implementation of *)
+(* the mxdirect predicate in matrix.v. *)
+(* *)
+(* - Sigma types: *)
+(* tag w == the i of w : {i : I & T i}. *)
+(* tagged w == the T i component of w : {i : I & T i}. *)
+(* Tagged T x == the {i : I & T i} with component x : T i. *)
+(* tag2 w == the i of w : {i : I & T i & U i}. *)
+(* tagged2 w == the T i component of w : {i : I & T i & U i}. *)
+(* tagged2' w == the U i component of w : {i : I & T i & U i}. *)
+(* Tagged2 T U x y == the {i : I & T i} with components x : T i and y : U i. *)
+(* sval u == the x of u : {x : T | P x}. *)
+(* s2val u == the x of u : {x : T | P x & Q x}. *)
+(* The properties of sval u, s2val u are given by lemmas svalP, s2valP, and *)
+(* s2valP'. We provide coercions sigT2 >-> sigT and sig2 >-> sig >-> sigT. *)
+(* A suite of lemmas (all_sig, ...) let us skolemize sig, sig2, sigT, sigT2 *)
+(* and pair, e.g., *)
+(* have /all_sig[f fP] (x : T): {y : U | P y} by ... *)
+(* yields an f : T -> U such that fP : forall x, P (f x). *)
+(* - Identity functions: *)
+(* id == NOTATION for the explicit identity function fun x => x. *)
+(* @id T == notation for the explicit identity at type T. *)
+(* idfun == an expression with a head constant, convertible to id; *)
+(* idfun x simplifies to x. *)
+(* @idfun T == the expression above, specialized to type T. *)
+(* phant_id x y == the function type phantom _ x -> phantom _ y. *)
+(* *** In addition to their casual use in functional programming, identity *)
+(* functions are often used to trigger static unification as part of the *)
+(* construction of dependent Records and Structures. For example, if we need *)
+(* a structure sT over a type T, we take as arguments T, sT, and a "dummy" *)
+(* function T -> sort sT: *)
+(* Definition foo T sT & T -> sort sT := ... *)
+(* We can avoid specifying sT directly by calling foo (@id T), or specify *)
+(* the call completely while still ensuring the consistency of T and sT, by *)
+(* calling @foo T sT idfun. The phant_id type allows us to extend this trick *)
+(* to non-Type canonical projections. It also allows us to sidestep *)
+(* dependent type constraints when building explicit records, e.g., given *)
+(* Record r := R {x; y : T(x)}. *)
+(* if we need to build an r from a given y0 while inferring some x0, such *)
+(* that y0 : T(x0), we pose *)
+(* Definition mk_r .. y .. (x := ...) y' & phant_id y y' := R x y'. *)
+(* Calling @mk_r .. y0 .. id will cause Coq to use y' := y0, while checking *)
+(* the dependent type constraint y0 : T(x0). *)
+(* *)
+(* - Extensional equality for functions and relations (i.e. functions of two *)
+(* arguments): *)
+(* f1 =1 f2 == f1 x is equal to f2 x for all x. *)
+(* f1 =1 f2 :> A == ... and f2 is explicitly typed. *)
+(* f1 =2 f2 == f1 x y is equal to f2 x y for all x y. *)
+(* f1 =2 f2 :> A == ... and f2 is explicitly typed. *)
+(* *)
+(* - Composition for total and partial functions: *)
+(* f^~ y == function f with second argument specialised to y, *)
+(* i.e., fun x => f x y *)
+(* CAVEAT: conditional (non-maximal) implicit arguments *)
+(* of f are NOT inserted in this context *)
+(* @^~ x == application at x, i.e., fun f => f x *)
+(* [eta f] == the explicit eta-expansion of f, i.e., fun x => f x *)
+(* CAVEAT: conditional (non-maximal) implicit arguments *)
+(* of f are NOT inserted in this context. *)
+(* fun=> v := the constant function fun _ => v. *)
+(* f1 \o f2 == composition of f1 and f2. *)
+(* Note: (f1 \o f2) x simplifies to f1 (f2 x). *)
+(* f1 \; f2 == categorical composition of f1 and f2. This expands to *)
+(* to f2 \o f1 and (f1 \; f2) x simplifies to f2 (f1 x). *)
+(* pcomp f1 f2 == composition of partial functions f1 and f2. *)
+(* *)
+(* *)
+(* - Properties of functions: *)
+(* injective f <-> f is injective. *)
+(* cancel f g <-> g is a left inverse of f / f is a right inverse of g. *)
+(* pcancel f g <-> g is a left inverse of f where g is partial. *)
+(* ocancel f g <-> g is a left inverse of f where f is partial. *)
+(* bijective f <-> f is bijective (has a left and right inverse). *)
+(* involutive f <-> f is involutive. *)
+(* *)
+(* - Properties for operations. *)
+(* left_id e op <-> e is a left identity for op (e op x = x). *)
+(* right_id e op <-> e is a right identity for op (x op e = x). *)
+(* left_inverse e inv op <-> inv is a left inverse for op wrt identity e, *)
+(* i.e., (inv x) op x = e. *)
+(* right_inverse e inv op <-> inv is a right inverse for op wrt identity e *)
+(* i.e., x op (i x) = e. *)
+(* self_inverse e op <-> each x is its own op-inverse (x op x = e). *)
+(* idempotent op <-> op is idempotent for op (x op x = x). *)
+(* associative op <-> op is associative, i.e., *)
+(* x op (y op z) = (x op y) op z. *)
+(* commutative op <-> op is commutative (x op y = y op x). *)
+(* left_commutative op <-> op is left commutative, i.e., *)
+(* x op (y op z) = y op (x op z). *)
+(* right_commutative op <-> op is right commutative, i.e., *)
+(* (x op y) op z = (x op z) op y. *)
+(* left_zero z op <-> z is a left zero for op (z op x = z). *)
+(* right_zero z op <-> z is a right zero for op (x op z = z). *)
+(* left_distributive op1 op2 <-> op1 distributes over op2 to the left: *)
+(* (x op2 y) op1 z = (x op1 z) op2 (y op1 z). *)
+(* right_distributive op1 op2 <-> op distributes over add to the right: *)
+(* x op1 (y op2 z) = (x op1 z) op2 (x op1 z). *)
+(* interchange op1 op2 <-> op1 and op2 satisfy an interchange law: *)
+(* (x op2 y) op1 (z op2 t) = (x op1 z) op2 (y op1 t). *)
+(* Note that interchange op op is a commutativity property. *)
+(* left_injective op <-> op is injective in its left argument: *)
+(* x op y = z op y -> x = z. *)
+(* right_injective op <-> op is injective in its right argument: *)
+(* x op y = x op z -> y = z. *)
+(* left_loop inv op <-> op, inv obey the inverse loop left axiom: *)
+(* (inv x) op (x op y) = y for all x, y, i.e., *)
+(* op (inv x) is always a left inverse of op x *)
+(* rev_left_loop inv op <-> op, inv obey the inverse loop reverse left *)
+(* axiom: x op ((inv x) op y) = y, for all x, y. *)
+(* right_loop inv op <-> op, inv obey the inverse loop right axiom: *)
+(* (x op y) op (inv y) = x for all x, y. *)
+(* rev_right_loop inv op <-> op, inv obey the inverse loop reverse right *)
+(* axiom: (x op y) op (inv y) = x for all x, y. *)
+(* Note that familiar "cancellation" identities like x + y - y = x or *)
+(* x - y + x = x are respectively instances of right_loop and rev_right_loop *)
+(* The corresponding lemmas will use the K and NK/VK suffixes, respectively. *)
+(* *)
+(* - Morphisms for functions and relations: *)
+(* {morph f : x / a >-> r} <-> f is a morphism with respect to functions *)
+(* (fun x => a) and (fun x => r); if r == R[x], *)
+(* this states that f a = R[f x] for all x. *)
+(* {morph f : x / a} <-> f is a morphism with respect to the *)
+(* function expression (fun x => a). This is *)
+(* shorthand for {morph f : x / a >-> a}; note *)
+(* that the two instances of a are often *)
+(* interpreted at different types. *)
+(* {morph f : x y / a >-> r} <-> f is a morphism with respect to functions *)
+(* (fun x y => a) and (fun x y => r). *)
+(* {morph f : x y / a} <-> f is a morphism with respect to the *)
+(* function expression (fun x y => a). *)
+(* {homo f : x / a >-> r} <-> f is a homomorphism with respect to the *)
+(* predicates (fun x => a) and (fun x => r); *)
+(* if r == R[x], this states that a -> R[f x] *)
+(* for all x. *)
+(* {homo f : x / a} <-> f is a homomorphism with respect to the *)
+(* predicate expression (fun x => a). *)
+(* {homo f : x y / a >-> r} <-> f is a homomorphism with respect to the *)
+(* relations (fun x y => a) and (fun x y => r). *)
+(* {homo f : x y / a} <-> f is a homomorphism with respect to the *)
+(* relation expression (fun x y => a). *)
+(* {mono f : x / a >-> r} <-> f is monotone with respect to projectors *)
+(* (fun x => a) and (fun x => r); if r == R[x], *)
+(* this states that R[f x] = a for all x. *)
+(* {mono f : x / a} <-> f is monotone with respect to the projector *)
+(* expression (fun x => a). *)
+(* {mono f : x y / a >-> r} <-> f is monotone with respect to relators *)
+(* (fun x y => a) and (fun x y => r). *)
+(* {mono f : x y / a} <-> f is monotone with respect to the relator *)
+(* expression (fun x y => a). *)
+(* *)
+(* The file also contains some basic lemmas for the above concepts. *)
+(* Lemmas relative to cancellation laws use some abbreviated suffixes: *)
+(* K - a cancellation rule like esymK : cancel (@esym T x y) (@esym T y x). *)
+(* LR - a lemma moving an operation from the left hand side of a relation to *)
+(* the right hand side, like canLR: cancel g f -> x = g y -> f x = y. *)
+(* RL - a lemma moving an operation from the right to the left, e.g., canRL. *)
+(* Beware that the LR and RL orientations refer to an "apply" (back chaining) *)
+(* usage; when using the same lemmas with "have" or "move" (forward chaining) *)
+(* the directions will be reversed!. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Delimit Scope fun_scope with FUN.
+Open Scope fun_scope.
+
+(* Notations for argument transpose *)
+Notation "f ^~ y" := (fun x => f x y)
+ (at level 10, y at level 8, no associativity, format "f ^~ y") : fun_scope.
+Notation "@^~ x" := (fun f => f x)
+ (at level 10, x at level 8, no associativity, format "@^~ x") : fun_scope.
+
+Delimit Scope pair_scope with PAIR.
+Open Scope pair_scope.
+
+(* Notations for pair/conjunction projections *)
+Notation "p .1" := (fst p)
+ (at level 2, left associativity, format "p .1") : pair_scope.
+Notation "p .2" := (snd p)
+ (at level 2, left associativity, format "p .2") : pair_scope.
+
+Coercion pair_of_and P Q (PandQ : P /\ Q) := (proj1 PandQ, proj2 PandQ).
+
+Definition all_pair I T U (w : forall i : I, T i * U i) :=
+ (fun i => (w i).1, fun i => (w i).2).
+
+(* Complements on the option type constructor, used below to *)
+(* encode partial functions. *)
+
+Module Option.
+
+Definition apply aT rT (f : aT -> rT) x u := if u is Some y then f y else x.
+
+Definition default T := apply (fun x : T => x).
+
+Definition bind aT rT (f : aT -> option rT) := apply f None.
+
+Definition map aT rT (f : aT -> rT) := bind (fun x => Some (f x)).
+
+End Option.
+
+Notation oapp := Option.apply.
+Notation odflt := Option.default.
+Notation obind := Option.bind.
+Notation omap := Option.map.
+Notation some := (@Some _) (only parsing).
+
+(* Shorthand for some basic equality lemmas. *)
+
+Notation erefl := refl_equal.
+Notation ecast i T e x := (let: erefl in _ = i := e return T in x).
+Definition esym := sym_eq.
+Definition nesym := sym_not_eq.
+Definition etrans := trans_eq.
+Definition congr1 := f_equal.
+Definition congr2 := f_equal2.
+(* Force at least one implicit when used as a view. *)
+Prenex Implicits esym nesym.
+
+(* A predicate for singleton types. *)
+Definition all_equal_to T (x0 : T) := forall x, unkeyed x = x0.
+
+Lemma unitE : all_equal_to tt. Proof. by case. Qed.
+
+(* A generic wrapper type *)
+
+Structure wrapped T := Wrap {unwrap : T}.
+Canonical wrap T x := @Wrap T x.
+
+Prenex Implicits unwrap wrap Wrap.
+
+(* Syntax for defining auxiliary recursive function. *)
+(* Usage: *)
+(* Section FooDefinition. *)
+(* Variables (g1 : T1) (g2 : T2). (globals) *)
+(* Fixoint foo_auxiliary (a3 : T3) ... := *)
+(* body, using [rec e3, ...] for recursive calls *)
+(* where "[ 'rec' a3 , a4 , ... ]" := foo_auxiliary. *)
+(* Definition foo x y .. := [rec e1, ...]. *)
+(* + proofs about foo *)
+(* End FooDefinition. *)
+
+Reserved Notation "[ 'rec' a0 ]"
+ (at level 0, format "[ 'rec' a0 ]").
+Reserved Notation "[ 'rec' a0 , a1 ]"
+ (at level 0, format "[ 'rec' a0 , a1 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 ]"
+ (at level 0, format "[ 'rec' a0 , a1 , a2 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 ]"
+ (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 ]"
+ (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 ]"
+ (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 ]"
+ (at level 0, format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 ]"
+ (at level 0,
+ format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 ]"
+ (at level 0,
+ format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 ]").
+Reserved Notation "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 , a9 ]"
+ (at level 0,
+ format "[ 'rec' a0 , a1 , a2 , a3 , a4 , a5 , a6 , a7 , a8 , a9 ]").
+
+(* Definitions and notation for explicit functions with simplification, *)
+(* i.e., which simpl and /= beta expand (this is complementary to nosimpl). *)
+
+Section SimplFun.
+
+Variables aT rT : Type.
+
+CoInductive simpl_fun := SimplFun of aT -> rT.
+
+Definition fun_of_simpl f := fun x => let: SimplFun lam := f in lam x.
+
+Coercion fun_of_simpl : simpl_fun >-> Funclass.
+
+End SimplFun.
+
+Notation "[ 'fun' : T => E ]" := (SimplFun (fun _ : T => E))
+ (at level 0,
+ format "'[hv' [ 'fun' : T => '/ ' E ] ']'") : fun_scope.
+
+Notation "[ 'fun' x => E ]" := (SimplFun (fun x => E))
+ (at level 0, x ident,
+ format "'[hv' [ 'fun' x => '/ ' E ] ']'") : fun_scope.
+
+Notation "[ 'fun' x : T => E ]" := (SimplFun (fun x : T => E))
+ (at level 0, x ident, only parsing) : fun_scope.
+
+Notation "[ 'fun' x y => E ]" := (fun x => [fun y => E])
+ (at level 0, x ident, y ident,
+ format "'[hv' [ 'fun' x y => '/ ' E ] ']'") : fun_scope.
+
+Notation "[ 'fun' x y : T => E ]" := (fun x : T => [fun y : T => E])
+ (at level 0, x ident, y ident, only parsing) : fun_scope.
+
+Notation "[ 'fun' ( x : T ) y => E ]" := (fun x : T => [fun y => E])
+ (at level 0, x ident, y ident, only parsing) : fun_scope.
+
+Notation "[ 'fun' x ( y : T ) => E ]" := (fun x => [fun y : T => E])
+ (at level 0, x ident, y ident, only parsing) : fun_scope.
+
+Notation "[ 'fun' ( x : xT ) ( y : yT ) => E ]" :=
+ (fun x : xT => [fun y : yT => E])
+ (at level 0, x ident, y ident, only parsing) : fun_scope.
+
+(* For delta functions in eqtype.v. *)
+Definition SimplFunDelta aT rT (f : aT -> aT -> rT) := [fun z => f z z].
+
+(* Extensional equality, for unary and binary functions, including syntactic *)
+(* sugar. *)
+
+Section ExtensionalEquality.
+
+Variables A B C : Type.
+
+Definition eqfun (f g : B -> A) : Prop := forall x, f x = g x.
+
+Definition eqrel (r s : C -> B -> A) : Prop := forall x y, r x y = s x y.
+
+Lemma frefl f : eqfun f f. Proof. by []. Qed.
+Lemma fsym f g : eqfun f g -> eqfun g f. Proof. by move=> eq_fg x. Qed.
+
+Lemma ftrans f g h : eqfun f g -> eqfun g h -> eqfun f h.
+Proof. by move=> eq_fg eq_gh x; rewrite eq_fg. Qed.
+
+Lemma rrefl r : eqrel r r. Proof. by []. Qed.
+
+End ExtensionalEquality.
+
+Typeclasses Opaque eqfun.
+Typeclasses Opaque eqrel.
+
+Hint Resolve frefl rrefl.
+
+Notation "f1 =1 f2" := (eqfun f1 f2)
+ (at level 70, no associativity) : fun_scope.
+Notation "f1 =1 f2 :> A" := (f1 =1 (f2 : A))
+ (at level 70, f2 at next level, A at level 90) : fun_scope.
+Notation "f1 =2 f2" := (eqrel f1 f2)
+ (at level 70, no associativity) : fun_scope.
+Notation "f1 =2 f2 :> A" := (f1 =2 (f2 : A))
+ (at level 70, f2 at next level, A at level 90) : fun_scope.
+
+Section Composition.
+
+Variables A B C : Type.
+
+Definition funcomp u (f : B -> A) (g : C -> B) x := let: tt := u in f (g x).
+Definition catcomp u g f := funcomp u f g.
+Local Notation comp := (funcomp tt).
+
+Definition pcomp (f : B -> option A) (g : C -> option B) x := obind f (g x).
+
+Lemma eq_comp f f' g g' : f =1 f' -> g =1 g' -> comp f g =1 comp f' g'.
+Proof. by move=> eq_ff' eq_gg' x; rewrite /= eq_gg' eq_ff'. Qed.
+
+End Composition.
+
+Notation comp := (funcomp tt).
+Notation "@ 'comp'" := (fun A B C => @funcomp A B C tt).
+Notation "f1 \o f2" := (comp f1 f2)
+ (at level 50, format "f1 \o '/ ' f2") : fun_scope.
+Notation "f1 \; f2" := (catcomp tt f1 f2)
+ (at level 60, right associativity, format "f1 \; '/ ' f2") : fun_scope.
+
+Notation "[ 'eta' f ]" := (fun x => f x)
+ (at level 0, format "[ 'eta' f ]") : fun_scope.
+
+Notation "'fun' => E" := (fun _ => E) (at level 200, only parsing) : fun_scope.
+
+Notation id := (fun x => x).
+Notation "@ 'id' T" := (fun x : T => x)
+ (at level 10, T at level 8, only parsing) : fun_scope.
+
+Definition id_head T u x : T := let: tt := u in x.
+Definition explicit_id_key := tt.
+Notation idfun := (id_head tt).
+Notation "@ 'idfun' T " := (@id_head T explicit_id_key)
+ (at level 10, T at level 8, format "@ 'idfun' T") : fun_scope.
+
+Definition phant_id T1 T2 v1 v2 := phantom T1 v1 -> phantom T2 v2.
+
+(* Strong sigma types. *)
+
+Section Tag.
+
+Variables (I : Type) (i : I) (T_ U_ : I -> Type).
+
+Definition tag := projS1.
+Definition tagged : forall w, T_(tag w) := @projS2 I [eta T_].
+Definition Tagged x := @existS I [eta T_] i x.
+
+Definition tag2 (w : @sigT2 I T_ U_) := let: existT2 _ _ i _ _ := w in i.
+Definition tagged2 w : T_(tag2 w) := let: existT2 _ _ _ x _ := w in x.
+Definition tagged2' w : U_(tag2 w) := let: existT2 _ _ _ _ y := w in y.
+Definition Tagged2 x y := @existS2 I [eta T_] [eta U_] i x y.
+
+End Tag.
+
+Arguments Tagged [I i].
+Arguments Tagged2 [I i].
+Prenex Implicits tag tagged Tagged tag2 tagged2 tagged2' Tagged2.
+
+Coercion tag_of_tag2 I T_ U_ (w : @sigT2 I T_ U_) :=
+ Tagged (fun i => T_ i * U_ i)%type (tagged2 w, tagged2' w).
+
+Lemma all_tag I T U :
+ (forall x : I, {y : T x & U x y}) ->
+ {f : forall x, T x & forall x, U x (f x)}.
+Proof. by move=> fP; exists (fun x => tag (fP x)) => x; case: (fP x). Qed.
+
+Lemma all_tag2 I T U V :
+ (forall i : I, {y : T i & U i y & V i y}) ->
+ {f : forall i, T i & forall i, U i (f i) & forall i, V i (f i)}.
+Proof. by case/all_tag=> f /all_pair[]; exists f. Qed.
+
+(* Refinement types. *)
+
+(* Prenex Implicits and renaming. *)
+Notation sval := (@proj1_sig _ _).
+Notation "@ 'sval'" := (@proj1_sig) (at level 10, format "@ 'sval'").
+
+Section Sig.
+
+Variables (T : Type) (P Q : T -> Prop).
+
+Lemma svalP (u : sig P) : P (sval u). Proof. by case: u. Qed.
+
+Definition s2val (u : sig2 P Q) := let: exist2 _ _ x _ _ := u in x.
+
+Lemma s2valP u : P (s2val u). Proof. by case: u. Qed.
+
+Lemma s2valP' u : Q (s2val u). Proof. by case: u. Qed.
+
+End Sig.
+
+Prenex Implicits svalP s2val s2valP s2valP'.
+
+Coercion tag_of_sig I P (u : @sig I P) := Tagged P (svalP u).
+
+Coercion sig_of_sig2 I P Q (u : @sig2 I P Q) :=
+ exist (fun i => P i /\ Q i) (s2val u) (conj (s2valP u) (s2valP' u)).
+
+Lemma all_sig I T P :
+ (forall x : I, {y : T x | P x y}) ->
+ {f : forall x, T x | forall x, P x (f x)}.
+Proof. by case/all_tag=> f; exists f. Qed.
+
+Lemma all_sig2 I T P Q :
+ (forall x : I, {y : T x | P x y & Q x y}) ->
+ {f : forall x, T x | forall x, P x (f x) & forall x, Q x (f x)}.
+Proof. by case/all_sig=> f /all_pair[]; exists f. Qed.
+
+Section Morphism.
+
+Variables (aT rT sT : Type) (f : aT -> rT).
+
+(* Morphism property for unary and binary functions *)
+Definition morphism_1 aF rF := forall x, f (aF x) = rF (f x).
+Definition morphism_2 aOp rOp := forall x y, f (aOp x y) = rOp (f x) (f y).
+
+(* Homomorphism property for unary and binary relations *)
+Definition homomorphism_1 (aP rP : _ -> Prop) := forall x, aP x -> rP (f x).
+Definition homomorphism_2 (aR rR : _ -> _ -> Prop) :=
+ forall x y, aR x y -> rR (f x) (f y).
+
+(* Stability property for unary and binary relations *)
+Definition monomorphism_1 (aP rP : _ -> sT) := forall x, rP (f x) = aP x.
+Definition monomorphism_2 (aR rR : _ -> _ -> sT) :=
+ forall x y, rR (f x) (f y) = aR x y.
+
+End Morphism.
+
+Notation "{ 'morph' f : x / a >-> r }" :=
+ (morphism_1 f (fun x => a) (fun x => r))
+ (at level 0, f at level 99, x ident,
+ format "{ 'morph' f : x / a >-> r }") : type_scope.
+
+Notation "{ 'morph' f : x / a }" :=
+ (morphism_1 f (fun x => a) (fun x => a))
+ (at level 0, f at level 99, x ident,
+ format "{ 'morph' f : x / a }") : type_scope.
+
+Notation "{ 'morph' f : x y / a >-> r }" :=
+ (morphism_2 f (fun x y => a) (fun x y => r))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'morph' f : x y / a >-> r }") : type_scope.
+
+Notation "{ 'morph' f : x y / a }" :=
+ (morphism_2 f (fun x y => a) (fun x y => a))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'morph' f : x y / a }") : type_scope.
+
+Notation "{ 'homo' f : x / a >-> r }" :=
+ (homomorphism_1 f (fun x => a) (fun x => r))
+ (at level 0, f at level 99, x ident,
+ format "{ 'homo' f : x / a >-> r }") : type_scope.
+
+Notation "{ 'homo' f : x / a }" :=
+ (homomorphism_1 f (fun x => a) (fun x => a))
+ (at level 0, f at level 99, x ident,
+ format "{ 'homo' f : x / a }") : type_scope.
+
+Notation "{ 'homo' f : x y / a >-> r }" :=
+ (homomorphism_2 f (fun x y => a) (fun x y => r))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'homo' f : x y / a >-> r }") : type_scope.
+
+Notation "{ 'homo' f : x y / a }" :=
+ (homomorphism_2 f (fun x y => a) (fun x y => a))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'homo' f : x y / a }") : type_scope.
+
+Notation "{ 'homo' f : x y /~ a }" :=
+ (homomorphism_2 f (fun y x => a) (fun x y => a))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'homo' f : x y /~ a }") : type_scope.
+
+Notation "{ 'mono' f : x / a >-> r }" :=
+ (monomorphism_1 f (fun x => a) (fun x => r))
+ (at level 0, f at level 99, x ident,
+ format "{ 'mono' f : x / a >-> r }") : type_scope.
+
+Notation "{ 'mono' f : x / a }" :=
+ (monomorphism_1 f (fun x => a) (fun x => a))
+ (at level 0, f at level 99, x ident,
+ format "{ 'mono' f : x / a }") : type_scope.
+
+Notation "{ 'mono' f : x y / a >-> r }" :=
+ (monomorphism_2 f (fun x y => a) (fun x y => r))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'mono' f : x y / a >-> r }") : type_scope.
+
+Notation "{ 'mono' f : x y / a }" :=
+ (monomorphism_2 f (fun x y => a) (fun x y => a))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'mono' f : x y / a }") : type_scope.
+
+Notation "{ 'mono' f : x y /~ a }" :=
+ (monomorphism_2 f (fun y x => a) (fun x y => a))
+ (at level 0, f at level 99, x ident, y ident,
+ format "{ 'mono' f : x y /~ a }") : type_scope.
+
+(* In an intuitionistic setting, we have two degrees of injectivity. The *)
+(* weaker one gives only simplification, and the strong one provides a left *)
+(* inverse (we show in `fintype' that they coincide for finite types). *)
+(* We also define an intermediate version where the left inverse is only a *)
+(* partial function. *)
+
+Section Injections.
+
+(* rT must come first so we can use @ to mitigate the Coq 1st order *)
+(* unification bug (e..g., Coq can't infer rT from a "cancel" lemma). *)
+Variables (rT aT : Type) (f : aT -> rT).
+
+Definition injective := forall x1 x2, f x1 = f x2 -> x1 = x2.
+
+Definition cancel g := forall x, g (f x) = x.
+
+Definition pcancel g := forall x, g (f x) = Some x.
+
+Definition ocancel (g : aT -> option rT) h := forall x, oapp h x (g x) = x.
+
+Lemma can_pcan g : cancel g -> pcancel (fun y => Some (g y)).
+Proof. by move=> fK x; congr (Some _). Qed.
+
+Lemma pcan_inj g : pcancel g -> injective.
+Proof. by move=> fK x y /(congr1 g); rewrite !fK => [[]]. Qed.
+
+Lemma can_inj g : cancel g -> injective.
+Proof. by move/can_pcan; apply: pcan_inj. Qed.
+
+Lemma canLR g x y : cancel g -> x = f y -> g x = y.
+Proof. by move=> fK ->. Qed.
+
+Lemma canRL g x y : cancel g -> f x = y -> x = g y.
+Proof. by move=> fK <-. Qed.
+
+End Injections.
+
+Lemma Some_inj {T} : injective (@Some T). Proof. by move=> x y []. Qed.
+
+(* cancellation lemmas for dependent type casts. *)
+Lemma esymK T x y : cancel (@esym T x y) (@esym T y x).
+Proof. by case: y /. Qed.
+
+Lemma etrans_id T x y (eqxy : x = y :> T) : etrans (erefl x) eqxy = eqxy.
+Proof. by case: y / eqxy. Qed.
+
+Section InjectionsTheory.
+
+Variables (A B C : Type) (f g : B -> A) (h : C -> B).
+
+Lemma inj_id : injective (@id A).
+Proof. by []. Qed.
+
+Lemma inj_can_sym f' : cancel f f' -> injective f' -> cancel f' f.
+Proof. by move=> fK injf' x; apply: injf'. Qed.
+
+Lemma inj_comp : injective f -> injective h -> injective (f \o h).
+Proof. by move=> injf injh x y /injf; apply: injh. Qed.
+
+Lemma can_comp f' h' : cancel f f' -> cancel h h' -> cancel (f \o h) (h' \o f').
+Proof. by move=> fK hK x; rewrite /= fK hK. Qed.
+
+Lemma pcan_pcomp f' h' :
+ pcancel f f' -> pcancel h h' -> pcancel (f \o h) (pcomp h' f').
+Proof. by move=> fK hK x; rewrite /pcomp fK /= hK. Qed.
+
+Lemma eq_inj : injective f -> f =1 g -> injective g.
+Proof. by move=> injf eqfg x y; rewrite -2!eqfg; apply: injf. Qed.
+
+Lemma eq_can f' g' : cancel f f' -> f =1 g -> f' =1 g' -> cancel g g'.
+Proof. by move=> fK eqfg eqfg' x; rewrite -eqfg -eqfg'. Qed.
+
+Lemma inj_can_eq f' : cancel f f' -> injective f' -> cancel g f' -> f =1 g.
+Proof. by move=> fK injf' gK x; apply: injf'; rewrite fK. Qed.
+
+End InjectionsTheory.
+
+Section Bijections.
+
+Variables (A B : Type) (f : B -> A).
+
+CoInductive bijective : Prop := Bijective g of cancel f g & cancel g f.
+
+Hypothesis bijf : bijective.
+
+Lemma bij_inj : injective f.
+Proof. by case: bijf => g fK _; apply: can_inj fK. Qed.
+
+Lemma bij_can_sym f' : cancel f' f <-> cancel f f'.
+Proof.
+split=> fK; first exact: inj_can_sym fK bij_inj.
+by case: bijf => h _ hK x; rewrite -[x]hK fK.
+Qed.
+
+Lemma bij_can_eq f' f'' : cancel f f' -> cancel f f'' -> f' =1 f''.
+Proof.
+by move=> fK fK'; apply: (inj_can_eq _ bij_inj); apply/bij_can_sym.
+Qed.
+
+End Bijections.
+
+Section BijectionsTheory.
+
+Variables (A B C : Type) (f : B -> A) (h : C -> B).
+
+Lemma eq_bij : bijective f -> forall g, f =1 g -> bijective g.
+Proof. by case=> f' fK f'K g eqfg; exists f'; eapply eq_can; eauto. Qed.
+
+Lemma bij_comp : bijective f -> bijective h -> bijective (f \o h).
+Proof.
+by move=> [f' fK f'K] [h' hK h'K]; exists (h' \o f'); apply: can_comp; auto.
+Qed.
+
+Lemma bij_can_bij : bijective f -> forall f', cancel f f' -> bijective f'.
+Proof. by move=> bijf; exists f; first by apply/(bij_can_sym bijf). Qed.
+
+End BijectionsTheory.
+
+Section Involutions.
+
+Variables (A : Type) (f : A -> A).
+
+Definition involutive := cancel f f.
+
+Hypothesis Hf : involutive.
+
+Lemma inv_inj : injective f. Proof. exact: can_inj Hf. Qed.
+Lemma inv_bij : bijective f. Proof. by exists f. Qed.
+
+End Involutions.
+
+Section OperationProperties.
+
+Variables S T R : Type.
+
+Section SopTisR.
+Implicit Type op : S -> T -> R.
+Definition left_inverse e inv op := forall x, op (inv x) x = e.
+Definition right_inverse e inv op := forall x, op x (inv x) = e.
+Definition left_injective op := forall x, injective (op^~ x).
+Definition right_injective op := forall y, injective (op y).
+End SopTisR.
+
+
+Section SopTisS.
+Implicit Type op : S -> T -> S.
+Definition right_id e op := forall x, op x e = x.
+Definition left_zero z op := forall x, op z x = z.
+Definition right_commutative op := forall x y z, op (op x y) z = op (op x z) y.
+Definition left_distributive op add :=
+ forall x y z, op (add x y) z = add (op x z) (op y z).
+Definition right_loop inv op := forall y, cancel (op^~ y) (op^~ (inv y)).
+Definition rev_right_loop inv op := forall y, cancel (op^~ (inv y)) (op^~ y).
+End SopTisS.
+
+Section SopTisT.
+Implicit Type op : S -> T -> T.
+Definition left_id e op := forall x, op e x = x.
+Definition right_zero z op := forall x, op x z = z.
+Definition left_commutative op := forall x y z, op x (op y z) = op y (op x z).
+Definition right_distributive op add :=
+ forall x y z, op x (add y z) = add (op x y) (op x z).
+Definition left_loop inv op := forall x, cancel (op x) (op (inv x)).
+Definition rev_left_loop inv op := forall x, cancel (op (inv x)) (op x).
+End SopTisT.
+
+Section SopSisT.
+Implicit Type op : S -> S -> T.
+Definition self_inverse e op := forall x, op x x = e.
+Definition commutative op := forall x y, op x y = op y x.
+End SopSisT.
+
+Section SopSisS.
+Implicit Type op : S -> S -> S.
+Definition idempotent op := forall x, op x x = x.
+Definition associative op := forall x y z, op x (op y z) = op (op x y) z.
+Definition interchange op1 op2 :=
+ forall x y z t, op1 (op2 x y) (op2 z t) = op2 (op1 x z) (op1 y t).
+End SopSisS.
+
+End OperationProperties.
+
+
+
+
+
+
+
+
+
+
diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml
new file mode 100644
index 000000000..663bca15e
--- /dev/null
+++ b/plugins/ssr/ssrfwd.ml
@@ -0,0 +1,409 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Tacmach
+
+open Ssrmatching_plugin.Ssrmatching
+
+open Ssrprinters
+open Ssrcommon
+open Ssrtacticals
+
+module RelDecl = Context.Rel.Declaration
+
+(** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *)
+(** Defined identifier *)
+
+
+let settac id c = Tactics.letin_tac None (Name id) c None
+let posetac id cl = Proofview.V82.of_tactic (settac id cl Locusops.nowhere)
+
+let ssrposetac ist (id, (_, t)) gl =
+ let sigma, t, ucst, _ = pf_abs_ssrterm ist gl t in
+ posetac id t (pf_merge_uc ucst gl)
+
+open Pp
+open Term
+
+let ssrsettac ist id ((_, (pat, pty)), (_, occ)) gl =
+ let pat = interp_cpattern ist gl pat (Option.map snd pty) in
+ let cl, sigma, env = pf_concl gl, project gl, pf_env gl in
+ let (c, ucst), cl =
+ let cl = EConstr.Unsafe.to_constr cl in
+ try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1
+ with NoMatch -> redex_of_pattern ~resolve_typeclasses:true env pat, cl in
+ let c = EConstr.of_constr c in
+ let cl = EConstr.of_constr cl in
+ if Termops.occur_existential sigma c then errorstrm(str"The pattern"++spc()++
+ pr_constr_pat (EConstr.Unsafe.to_constr c)++spc()++str"did not match and has holes."++spc()++
+ str"Did you mean pose?") else
+ let c, (gl, cty) = match EConstr.kind sigma c with
+ | Cast(t, DEFAULTcast, ty) -> t, (gl, ty)
+ | _ -> c, pfe_type_of gl c in
+ let cl' = EConstr.mkLetIn (Name id, c, cty, cl) in
+ let gl = pf_merge_uc ucst gl in
+ Tacticals.tclTHEN (Proofview.V82.of_tactic (convert_concl cl')) (introid id) gl
+
+open Util
+
+let rec is_Evar_or_CastedMeta sigma x =
+ EConstr.isEvar sigma x || EConstr.isMeta sigma x ||
+ (EConstr.isCast sigma x && is_Evar_or_CastedMeta sigma (pi1 (EConstr.destCast sigma x)))
+
+let occur_existential_or_casted_meta c =
+ let rec occrec c = match kind_of_term c with
+ | Evar _ -> raise Not_found
+ | Cast (m,_,_) when isMeta m -> raise Not_found
+ | _ -> iter_constr occrec c
+ in try occrec c; false with Not_found -> true
+
+open Printer
+
+let examine_abstract id gl =
+ let gl, tid = pfe_type_of gl id in
+ let abstract, gl = pf_mkSsrConst "abstract" gl in
+ let sigma = project gl in
+ if not (EConstr.isApp sigma tid) || not (EConstr.eq_constr sigma (fst(EConstr.destApp sigma tid)) abstract) then
+ errorstrm(strbrk"not an abstract constant: "++pr_econstr id);
+ let _, args_id = EConstr.destApp sigma tid in
+ if Array.length args_id <> 3 then
+ errorstrm(strbrk"not a proper abstract constant: "++pr_econstr id);
+ if not (is_Evar_or_CastedMeta sigma args_id.(2)) then
+ errorstrm(strbrk"abstract constant "++pr_econstr id++str" already used");
+ tid, args_id
+
+let pf_find_abstract_proof check_lock gl abstract_n =
+ let fire gl t = EConstr.Unsafe.to_constr (Reductionops.nf_evar (project gl) (EConstr.of_constr t)) in
+ let abstract, gl = pf_mkSsrConst "abstract" gl in
+ let l = Evd.fold_undefined (fun e ei l ->
+ match kind_of_term ei.Evd.evar_concl with
+ | App(hd, [|ty; n; lock|])
+ when (not check_lock ||
+ (occur_existential_or_casted_meta (fire gl ty) &&
+ is_Evar_or_CastedMeta (project gl) (EConstr.of_constr @@ fire gl lock))) &&
+ Term.eq_constr hd (EConstr.Unsafe.to_constr abstract) && Term.eq_constr n abstract_n -> e::l
+ | _ -> l) (project gl) [] in
+ match l with
+ | [e] -> e
+ | _ -> errorstrm(strbrk"abstract constant "++pr_constr abstract_n++
+ strbrk" not found in the evar map exactly once. "++
+ strbrk"Did you tamper with it?")
+
+let reduct_in_concl t = Tactics.reduct_in_concl (t, DEFAULTcast)
+let unfold cl =
+ let module R = Reductionops in let module F = CClosure.RedFlags in
+ reduct_in_concl (R.clos_norm_flags (F.mkflags
+ (List.map (fun c -> F.fCONST (fst (destConst (EConstr.Unsafe.to_constr c)))) cl @
+ [F.fBETA; F.fMATCH; F.fFIX; F.fCOFIX])))
+
+open Ssrast
+open Ssripats
+
+let ssrhaveNOtcresolution = Summary.ref ~name:"SSR:havenotcresolution" false
+
+let inHaveTCResolution = Libobject.declare_object {
+ (Libobject.default_object "SSRHAVETCRESOLUTION") with
+ Libobject.cache_function = (fun (_,v) -> ssrhaveNOtcresolution := v);
+ Libobject.load_function = (fun _ (_,v) -> ssrhaveNOtcresolution := v);
+ Libobject.classify_function = (fun v -> Libobject.Keep v);
+}
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "have type classes";
+ Goptions.optkey = ["SsrHave";"NoTCResolution"];
+ Goptions.optread = (fun _ -> !ssrhaveNOtcresolution);
+ Goptions.optdepr = false;
+ Goptions.optwrite = (fun b ->
+ Lib.add_anonymous_leaf (inHaveTCResolution b)) }
+
+
+open Constrexpr
+open Glob_term
+open Misctypes
+
+let combineCG t1 t2 f g = match t1, t2 with
+ | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None)
+ | (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2))
+ | _, (_, (_, None)) -> anomaly "have: mixed C-G constr"
+ | _ -> anomaly "have: mixed G-C constr"
+
+let basecuttac name c gl =
+ let hd, gl = pf_mkSsrConst name gl in
+ let t = EConstr.mkApp (hd, [|c|]) in
+ let gl, _ = pf_e_type_of gl t in
+ Proofview.V82.of_tactic (Tactics.apply t) gl
+
+let havetac ist
+ (transp,((((clr, pats), binders), simpl), (((fk, _), t), hint)))
+ suff namefst gl
+=
+ let concl = pf_concl gl in
+ let skols, pats =
+ List.partition (function IPatNewHidden _ -> true | _ -> false) pats in
+ let itac_mkabs = introstac ~ist skols in
+ let itac_c = introstac ~ist (IPatClear clr :: pats) in
+ let itac, id, clr = introstac ~ist pats, Tacticals.tclIDTAC, cleartac clr in
+ let binderstac n =
+ let rec aux = function 0 -> [] | n -> IPatAnon One :: aux (n-1) in
+ Tacticals.tclTHEN (if binders <> [] then introstac ~ist (aux n) else Tacticals.tclIDTAC)
+ (introstac ~ist binders) in
+ let simpltac = introstac ~ist simpl in
+ let fixtc =
+ not !ssrhaveNOtcresolution &&
+ match fk with FwdHint(_,true) -> false | _ -> true in
+ let hint = hinttac ist true hint in
+ let cuttac t gl =
+ if transp then
+ let have_let, gl = pf_mkSsrConst "ssr_have_let" gl in
+ let step = EConstr.mkApp (have_let, [|concl;t|]) in
+ let gl, _ = pf_e_type_of gl step in
+ applyn ~with_evars:true ~with_shelve:false 2 step gl
+ else basecuttac "ssr_have" t gl in
+ (* Introduce now abstract constants, so that everything sees them *)
+ let abstract_key, gl = pf_mkSsrConst "abstract_key" gl in
+ let unlock_abs (idty,args_id) gl =
+ let gl, _ = pf_e_type_of gl idty in
+ pf_unify_HO gl args_id.(2) abstract_key in
+ Tacticals.tclTHENFIRST itac_mkabs (fun gl ->
+ let mkt t = mk_term xNoFlag t in
+ let mkl t = (xNoFlag, (t, None)) in
+ let interp gl rtc t = pf_abs_ssrterm ~resolve_typeclasses:rtc ist gl t in
+ let interp_ty gl rtc t =
+ let a,b,_,u = pf_interp_ty ~resolve_typeclasses:rtc ist gl t in a,b,u in
+ let open CAst in
+ let ct, cty, hole, loc = match t with
+ | _, (_, Some { loc; v = CCast (ct, CastConv cty)}) ->
+ mkt ct, mkt cty, mkt (mkCHole None), loc
+ | _, (_, Some ct) ->
+ mkt ct, mkt (mkCHole None), mkt (mkCHole None), None
+ | _, ({ loc; v = GCast (ct, CastConv cty) }, None) ->
+ mkl ct, mkl cty, mkl mkRHole, loc
+ | _, (t, None) -> mkl t, mkl mkRHole, mkl mkRHole, None in
+ let gl, cut, sol, itac1, itac2 =
+ match fk, namefst, suff with
+ | FwdHave, true, true ->
+ errorstrm (str"Suff have does not accept a proof term")
+ | FwdHave, false, true ->
+ let cty = combineCG cty hole (mkCArrow ?loc) mkRArrow in
+ let _,t,uc,_ = interp gl false (combineCG ct cty (mkCCast ?loc) mkRCast) in
+ let gl = pf_merge_uc uc gl in
+ let gl, ty = pfe_type_of gl t in
+ let ctx, _ = EConstr.decompose_prod_n_assum (project gl) 1 ty in
+ 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
+ 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
+ | IPatNewHidden ids -> ids
+ | _ -> assert false) skols) in
+ let skols_args =
+ List.map (fun id -> examine_abstract (EConstr.mkVar id) gl) skols in
+ let gl = List.fold_right unlock_abs skols_args gl in
+ let sigma, t, uc, n_evars =
+ interp gl false (combineCG ct cty (mkCCast ?loc) mkRCast) in
+ if skols <> [] && n_evars <> 0 then
+ CErrors.user_err (Pp.strbrk @@ "Automatic generalization of unresolved implicit "^
+ "arguments together with abstract variables is "^
+ "not supported");
+ let gl = re_sig (sig_it gl) (Evd.merge_universe_context sigma uc) in
+ let gs =
+ List.map (fun (_,a) ->
+ pf_find_abstract_proof false gl (EConstr.Unsafe.to_constr a.(1))) skols_args in
+ let tacopen_skols gl =
+ let stuff, g = Refiner.unpackage gl in
+ Refiner.repackage stuff (gs @ [g]) in
+ let gl, ty = pf_e_type_of gl t in
+ gl, ty, Proofview.V82.of_tactic (Tactics.apply t), id,
+ Tacticals.tclTHEN (Tacticals.tclTHEN itac_c simpltac)
+ (Tacticals.tclTHEN tacopen_skols (fun gl ->
+ let abstract, gl = pf_mkSsrConst "abstract" gl in
+ Proofview.V82.of_tactic (unfold [abstract; abstract_key]) gl))
+ | _,true,true ->
+ let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in
+ gl, EConstr.mkArrow ty concl, hint, itac, clr
+ | _,false,true ->
+ let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in
+ gl, EConstr.mkArrow ty concl, hint, id, itac_c
+ | _, false, false ->
+ let n, cty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in
+ gl, cty, Tacticals.tclTHEN (binderstac n) hint, id, Tacticals.tclTHEN itac_c simpltac
+ | _, true, false -> assert false in
+ Tacticals.tclTHENS (cuttac cut) [ Tacticals.tclTHEN sol itac1; itac2 ] gl)
+ gl
+;;
+
+(* to extend the abstract value one needs:
+ Utility lemma to partially instantiate an abstract constant type.
+ Lemma use_abstract T n l (x : abstract T n l) : T.
+ Proof. by case: l x. Qed.
+*)
+let ssrabstract ist gens (*last*) gl =
+ let main _ (_,cid) ist gl =
+(*
+ let proj1, proj2, prod =
+ let pdata = build_prod () in
+ pdata.Coqlib.proj1, pdata.Coqlib.proj2, pdata.Coqlib.typ in
+*)
+ let concl, env = pf_concl gl, pf_env gl in
+ let fire gl t = Reductionops.nf_evar (project gl) t in
+ let abstract, gl = pf_mkSsrConst "abstract" gl in
+ let abstract_key, gl = pf_mkSsrConst "abstract_key" gl in
+ let cid_interpreted = interp_cpattern ist gl cid None in
+ let id = EConstr.mkVar (Option.get (id_of_pattern cid_interpreted)) in
+ let idty, args_id = examine_abstract id gl in
+ let abstract_n = args_id.(1) in
+ let abstract_proof = pf_find_abstract_proof true gl (EConstr.Unsafe.to_constr abstract_n) in
+ let gl, proof =
+ let pf_unify_HO gl a b =
+ try pf_unify_HO gl a b
+ with _ -> errorstrm(strbrk"The abstract variable "++pr_econstr id++
+ strbrk" cannot abstract this goal. Did you generalize it?") in
+ let find_hole p t =
+ match EConstr.kind (project gl) t with
+ | Evar _ (*when last*) -> pf_unify_HO gl concl t, p
+ | Meta _ (*when last*) -> pf_unify_HO gl concl t, p
+ | Cast(m,_,_) when EConstr.isEvar (project gl) m || EConstr.isMeta
+ (project gl) m (*when last*) -> pf_unify_HO gl concl t, p
+(*
+ | Evar _ ->
+ let sigma, it = project gl, sig_it gl in
+ let sigma, ty = Evarutil.new_type_evar sigma env in
+ let gl = re_sig it sigma in
+ let p = mkApp (proj2,[|ty;concl;p|]) in
+ let concl = mkApp(prod,[|ty; concl|]) in
+ pf_unify_HO gl concl t, p
+ | App(hd, [|left; right|]) when Term.eq_constr hd prod ->
+ find_hole (mkApp (proj1,[|left;right;p|])) left
+*)
+ | _ -> errorstrm(strbrk"abstract constant "++pr_econstr abstract_n++
+ strbrk" has an unexpected shape. Did you tamper with it?")
+ in
+ find_hole
+ ((*if last then*) id
+ (*else mkApp(mkSsrConst "use_abstract",Array.append args_id [|id|])*))
+ (fire gl args_id.(0)) in
+ let gl = (*if last then*) pf_unify_HO gl abstract_key args_id.(2) (*else gl*) in
+ let gl, _ = pf_e_type_of gl idty in
+ let proof = fire gl proof in
+(* if last then *)
+ let tacopen gl =
+ let stuff, g = Refiner.unpackage gl in
+ Refiner.repackage stuff [ g; abstract_proof ] in
+ Tacticals.tclTHENS tacopen [Tacticals.tclSOLVE [Proofview.V82.of_tactic (Tactics.apply proof)]; Proofview.V82.of_tactic (unfold[abstract;abstract_key])] gl
+(* else apply proof gl *)
+ in
+ let introback ist (gens, _) =
+ introstac ~ist
+ (List.map (fun (_,cp) -> match id_of_pattern (interp_cpattern ist gl cp None) with
+ | None -> IPatAnon One
+ | Some id -> IPatId id)
+ (List.tl (List.hd gens))) in
+ Tacticals.tclTHEN (with_dgens gens main ist) (introback ist gens) gl
+
+
+let destProd_or_LetIn sigma c =
+ match EConstr.kind sigma c with
+ | Prod (n,ty,c) -> RelDecl.LocalAssum (n, ty), c
+ | LetIn (n,bo,ty,c) -> RelDecl.LocalDef (n, bo, ty), c
+ | _ -> raise DestKO
+
+let wlogtac ist (((clr0, pats),_),_) (gens, ((_, ct))) hint suff ghave gl =
+ let mkabs gen = abs_wgen false ist (fun x -> x) gen in
+ let mkclr gen clrs = clr_of_wgen gen clrs in
+ let mkpats = function
+ | _, Some ((x, _), _) -> fun pats -> IPatId (hoi_id x) :: pats
+ | _ -> fun x -> x in
+ let open CAst in
+ let ct = match ct with
+ | (a, (b, Some { v = CCast (_, CastConv cty)})) -> a, (b, Some cty)
+ | (a, ({ v = GCast (_, CastConv cty) }, None)) -> a, (cty, None)
+ | _ -> anomaly "wlog: ssr cast hole deleted by typecheck" in
+ let cut_implies_goal = not (suff || ghave <> `NoGen) in
+ let c, args, ct, gl =
+ let gens = List.filter (function _, Some _ -> true | _ -> false) gens in
+ let concl = pf_concl gl in
+ let c = EConstr.mkProp in
+ let c = if cut_implies_goal then EConstr.mkArrow c concl else c in
+ let gl, args, c = List.fold_right mkabs gens (gl,[],c) in
+ let env, _ =
+ List.fold_left (fun (env, c) _ ->
+ let rd, c = destProd_or_LetIn (project gl) c in
+ EConstr.push_rel rd env, c) (pf_env gl, c) gens in
+ let sigma = project gl in
+ let (sigma, ev) = Evarutil.new_evar env sigma EConstr.mkProp in
+ let k, _ = EConstr.destEvar sigma ev in
+ let fake_gl = {Evd.it = k; Evd.sigma = sigma} in
+ let _, ct, _, uc = pf_interp_ty ist fake_gl ct in
+ let rec var2rel c g s = match EConstr.kind sigma c, g with
+ | Prod(Anonymous,_,c), [] -> EConstr.mkProd(Anonymous, EConstr.Vars.subst_vars s ct, c)
+ | Sort _, [] -> EConstr.Vars.subst_vars s ct
+ | LetIn(Name id as n,b,ty,c), _::g -> EConstr.mkLetIn (n,b,ty,var2rel c g (id::s))
+ | Prod(Name id as n,ty,c), _::g -> EConstr.mkProd (n,ty,var2rel c g (id::s))
+ | _ -> CErrors.anomaly(str"SSR: wlog: var2rel: " ++ pr_econstr c) in
+ let c = var2rel c gens [] in
+ let rec pired c = function
+ | [] -> c
+ | t::ts as args -> match EConstr.kind sigma c with
+ | Prod(_,_,c) -> pired (EConstr.Vars.subst1 t c) ts
+ | LetIn(id,b,ty,c) -> EConstr.mkLetIn (id,b,ty,pired c args)
+ | _ -> CErrors.anomaly(str"SSR: wlog: pired: " ++ pr_econstr c) in
+ c, args, pired c args, pf_merge_uc uc gl in
+ let tacipat pats = introstac ~ist pats in
+ let tacigens =
+ Tacticals.tclTHEN
+ (Tacticals.tclTHENLIST(List.rev(List.fold_right mkclr gens [cleartac clr0])))
+ (introstac ~ist (List.fold_right mkpats gens [])) in
+ let hinttac = hinttac ist true hint in
+ let cut_kind, fst_goal_tac, snd_goal_tac =
+ match suff, ghave with
+ | true, `NoGen -> "ssr_wlog", Tacticals.tclTHEN hinttac (tacipat pats), tacigens
+ | false, `NoGen -> "ssr_wlog", hinttac, Tacticals.tclTHEN tacigens (tacipat pats)
+ | true, `Gen _ -> assert false
+ | false, `Gen id ->
+ if gens = [] then errorstrm(str"gen have requires some generalizations");
+ let clear0 = cleartac clr0 in
+ let id, name_general_hyp, cleanup, pats = match id, pats with
+ | None, (IPatId id as ip)::pats -> Some id, tacipat [ip], clear0, pats
+ | None, _ -> None, Tacticals.tclIDTAC, clear0, pats
+ | Some (Some id),_ -> Some id, introid id, clear0, pats
+ | Some _,_ ->
+ let id = mk_anon_id "tmp" gl in
+ Some id, introid id, Tacticals.tclTHEN clear0 (Proofview.V82.of_tactic (Tactics.clear [id])), pats in
+ let tac_specialize = match id with
+ | None -> Tacticals.tclIDTAC
+ | Some id ->
+ if pats = [] then Tacticals.tclIDTAC else
+ let args = Array.of_list args in
+ ppdebug(lazy(str"specialized="++pr_econstr EConstr.(mkApp (mkVar id,args))));
+ ppdebug(lazy(str"specialized_ty="++pr_econstr ct));
+ Tacticals.tclTHENS (basecuttac "ssr_have" ct)
+ [Proofview.V82.of_tactic (Tactics.apply EConstr.(mkApp (mkVar id,args))); Tacticals.tclIDTAC] in
+ "ssr_have",
+ (if hint = nohint then tacigens else hinttac),
+ Tacticals.tclTHENLIST [name_general_hyp; tac_specialize; tacipat pats; cleanup]
+ in
+ Tacticals.tclTHENS (basecuttac cut_kind c) [fst_goal_tac; snd_goal_tac] gl
+
+(** The "suffice" tactic *)
+
+let sufftac ist ((((clr, pats),binders),simpl), ((_, c), hint)) =
+ let htac = Tacticals.tclTHEN (introstac ~ist pats) (hinttac ist true hint) in
+ let open CAst in
+ let c = match c with
+ | (a, (b, Some { v = CCast (_, CastConv cty)})) -> a, (b, Some cty)
+ | (a, ({ v = GCast (_, CastConv cty) }, None)) -> a, (cty, None)
+ | _ -> anomaly "suff: ssr cast hole deleted by typecheck" in
+ let ctac gl =
+ let _,ty,_,uc = pf_interp_ty ist gl c in let gl = pf_merge_uc uc gl in
+ basecuttac "ssr_suff" ty gl in
+ Tacticals.tclTHENS ctac [htac; Tacticals.tclTHEN (cleartac clr) (introstac ~ist (binders@simpl))]
diff --git a/plugins/ssr/ssrfwd.mli b/plugins/ssr/ssrfwd.mli
new file mode 100644
index 000000000..6fb97d524
--- /dev/null
+++ b/plugins/ssr/ssrfwd.mli
@@ -0,0 +1,65 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+
+open Ltac_plugin
+
+open Ssrast
+
+val ssrsettac : ist -> Id.t -> ((ssrfwdfmt * (Ssrmatching_plugin.Ssrmatching.cpattern * ssrterm option)) * ssrdocc) -> v82tac
+
+val ssrposetac : ist -> (Id.t * (ssrfwdfmt * ssrterm)) -> v82tac
+
+val havetac :
+ Ssrast.ist ->
+ bool *
+ ((((Ssrast.ssrclear * Ssrast.ssripat list) * Ssrast.ssripats) *
+ Ssrast.ssripats) *
+ (((Ssrast.ssrfwdkind * 'a) *
+ ('b * (Glob_term.glob_constr * Constrexpr.constr_expr option))) *
+ (bool * Tacinterp.Value.t option list))) ->
+ bool ->
+ bool -> v82tac
+val ssrabstract :
+ Tacinterp.interp_sign ->
+ (Ssrast.ssrdocc * Ssrmatching_plugin.Ssrmatching.cpattern) list
+ list * Ssrast.ssrclear -> v82tac
+
+val basecuttac :
+ string ->
+ EConstr.t -> Proof_type.goal Evd.sigma -> Evar.t list Evd.sigma
+
+val wlogtac :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ ((Ssrast.ssrhyps * Ssrast.ssripats) * 'a) * 'b ->
+ (Ssrast.ssrhyps *
+ ((Ssrast.ssrhyp_or_id * string) *
+ Ssrmatching_plugin.Ssrmatching.cpattern option)
+ option)
+ list *
+ ('c *
+ (Ssrast.ssrtermkind *
+ (Glob_term.glob_constr * Constrexpr.constr_expr option))) ->
+ Ltac_plugin.Tacinterp.Value.t Ssrast.ssrhint ->
+ bool ->
+ [< `Gen of Names.Id.t option option | `NoGen > `NoGen ] ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val sufftac :
+ Ssrast.ist ->
+ (((Ssrast.ssrhyps * Ssrast.ssripats) * Ssrast.ssripat list) *
+ Ssrast.ssripat list) *
+ (('a *
+ (Ssrast.ssrtermkind *
+ (Glob_term.glob_constr * Constrexpr.constr_expr option))) *
+ (bool * Tacinterp.Value.t option list)) ->
+ Proof_type.tactic
+
diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml
new file mode 100644
index 000000000..b850b0e95
--- /dev/null
+++ b/plugins/ssr/ssripats.ml
@@ -0,0 +1,400 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Pp
+open Term
+open Tactics
+open Tacticals
+open Tacmach
+open Coqlib
+open Util
+open Evd
+open Printer
+
+open Ssrmatching_plugin
+open Ssrmatching
+open Ssrast
+open Ssrprinters
+open Ssrcommon
+open Ssrequality
+open Ssrview
+open Ssrelim
+open Ssrbwd
+
+module RelDecl = Context.Rel.Declaration
+(** Extended intro patterns {{{ ***********************************************)
+
+
+(* There are two ways of "applying" a view to term: *)
+(* 1- using a view hint if the view is an instance of some *)
+(* (reflection) inductive predicate. *)
+(* 2- applying the view if it coerces to a function, adding *)
+(* implicit arguments. *)
+(* They require guessing the view hints and the number of *)
+(* implicits, respectively, which we do by brute force. *)
+
+let apply_type x xs = Proofview.V82.of_tactic (apply_type x xs)
+
+let new_tac = Proofview.V82.of_tactic
+
+let with_top tac gl =
+ tac_ctx
+ (tclTHENLIST [ introid top_id; tac (EConstr.mkVar top_id); new_tac (clear [top_id])])
+ gl
+
+let tclTHENS_nonstrict tac tacl taclname gl =
+ let tacres = tac gl in
+ let n_gls = List.length (sig_it tacres) in
+ let n_tac = List.length tacl in
+ if n_gls = n_tac then tclTHENS_a (fun _ -> tacres) tacl gl else
+ if n_gls = 0 then tacres else
+ let pr_only n1 n2 = if n1 < n2 then str "only " else mt () in
+ let pr_nb n1 n2 name =
+ pr_only n1 n2 ++ int n1 ++ str (" " ^ String.plural n1 name) in
+ errorstrm (pr_nb n_tac n_gls taclname ++ spc ()
+ ++ str "for " ++ pr_nb n_gls n_tac "subgoal")
+
+let rec nat_of_n n =
+ if n = 0 then EConstr.mkConstruct path_of_O
+ else EConstr.mkApp (EConstr.mkConstruct path_of_S, [|nat_of_n (n-1)|])
+
+let ssr_abstract_id = Summary.ref ~name:"SSR:abstractid" 0
+
+let mk_abstract_id () = incr ssr_abstract_id; nat_of_n !ssr_abstract_id
+
+let ssrmkabs id gl =
+ let env, concl = pf_env gl, Tacmach.pf_concl gl in
+ let step = begin fun sigma ->
+ let (sigma, (abstract_proof, abstract_ty)) =
+ let (sigma, (ty, _)) =
+ Evarutil.new_type_evar env sigma Evd.univ_flexible_alg in
+ let (sigma, ablock) = mkSsrConst "abstract_lock" env sigma in
+ let (sigma, lock) = Evarutil.new_evar env sigma ablock in
+ let (sigma, abstract) = mkSsrConst "abstract" env sigma in
+ let abstract_ty = EConstr.mkApp(abstract, [|ty;mk_abstract_id ();lock|]) in
+ let (sigma, m) = Evarutil.new_evar env sigma abstract_ty in
+ (sigma, (m, abstract_ty)) in
+ let sigma, kont =
+ let rd = RelDecl.LocalAssum (Name id, abstract_ty) in
+ let (sigma, ev) = Evarutil.new_evar (EConstr.push_rel rd env) sigma concl in
+ (sigma, ev)
+ in
+(* pp(lazy(pr_econstr concl)); *)
+ let term = EConstr.(mkApp (mkLambda(Name id,abstract_ty,kont) ,[|abstract_proof|])) in
+ let sigma, _ = Typing.type_of env sigma term in
+ (sigma, term)
+ end in
+ Proofview.V82.of_tactic
+ (Proofview.tclTHEN
+ (Tactics.New.refine step)
+ (Proofview.tclFOCUS 1 3 Proofview.shelve)) gl
+
+let ssrmkabstac ids =
+ List.fold_right (fun id tac -> tclTHENFIRST (ssrmkabs id) tac) ids tclIDTAC
+
+(* introstac: for "move" and "clear", tclEQINTROS: for "case" and "elim" *)
+(* This block hides the spaghetti-code needed to implement the only two *)
+(* tactics that should be used to process intro patters. *)
+(* The difficulty is that we don't want to always rename, but we can *)
+(* compute needeed renamings only at runtime, so we theread a tree like *)
+(* imperativestructure so that outer renamigs are inherited by inner *)
+(* ipts and that the cler performed at the end of ipatstac clears hyps *)
+(* eventually renamed at runtime. *)
+let delayed_clear force rest clr gl =
+ let gl, ctx = pull_ctx gl in
+ let hyps = pf_hyps gl in
+ let () = if not force then List.iter (check_hyp_exists hyps) clr in
+ if List.exists (fun x -> force || is_name_in_ipats (hyp_id x) rest) clr then
+ 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
+ x', (x, x')) clr) in
+ let ctx = { ctx with delayed_clears = ren_clr @ ctx.delayed_clears } in
+ let gl = push_ctx ctx gl in
+ tac_ctx (Proofview.V82.of_tactic (rename_hyp ren)) gl
+ else
+ let ctx = { ctx with delayed_clears = hyps_ids clr @ ctx.delayed_clears } in
+ let gl = push_ctx ctx gl in
+ tac_ctx tclIDTAC gl
+
+(* Common code to handle generalization lists along with the defective case *)
+
+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
+ | _ -> top_id in
+ let top_gen = mkclr clr, cpattern_of_id top_id in
+ tclTHEN (introid top_id) (maintac deps top_gen ist) gl
+
+let with_defective_a maintac deps clr ist gl =
+ let sigma = sig_sig gl in
+ 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
+ | _ -> 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
+
+let with_dgens (gensl, clr) maintac ist = match gensl with
+ | [deps; []] -> with_defective maintac deps clr ist
+ | [deps; gen :: gens] ->
+ tclTHEN (genstac (gens, clr) ist) (maintac deps gen ist)
+ | [gen :: gens] -> tclTHEN (genstac (gens, clr) ist) (maintac [] gen ist)
+ | _ -> with_defective maintac [] clr ist
+
+let viewmovetac_aux ?(next=ref []) clear name_ref (_, vl as v) _ gen ist gl =
+ let cl, c, clr, gl, gen_pat =
+ let gl, ctx = pull_ctx gl in
+ let _, gen_pat, a, b, c, ucst, gl = pf_interp_gen_aux ist gl false gen in
+ a, b ,c, push_ctx ctx (pf_merge_uc ucst gl), gen_pat in
+ let clr = if clear then clr else [] in
+ name_ref := (match id_of_pattern gen_pat with Some id -> id | _ -> top_id);
+ let clr = if clear then clr else [] in
+ if vl = [] then tac_ctx (genclrtac cl [c] clr) gl
+ else
+ let _, _, gl =
+ pfa_with_view ist ~next v cl c
+ (fun cl c -> tac_ctx (genclrtac cl [c] clr)) clr gl in
+ gl
+
+let move_top_with_view ~next c r v =
+ with_defective_a (viewmovetac_aux ~next c r v) [] []
+
+type block_names = (int * EConstr.types array) option
+
+let (introstac : ?ist:Tacinterp.interp_sign -> ssripats -> Proof_type.tactic),
+ (tclEQINTROS : ?ind:block_names ref -> ?ist:Tacinterp.interp_sign ->
+ Proof_type.tactic -> Proof_type.tactic -> ssripats ->
+ Proof_type.tactic)
+=
+
+ let rec ipattac ?ist ~next p : tac_ctx tac_a = fun gl ->
+(* pp(lazy(str"ipattac: " ++ pr_ipat p)); *)
+ match p with
+ | IPatAnon Drop ->
+ let id, gl = with_ctx new_wild_id gl in
+ tac_ctx (introid id) gl
+ | IPatAnon All -> tac_ctx intro_all gl
+ (* TODO
+ | IPatAnon Temporary ->
+ let (id, orig), gl = with_ctx new_tmp_id gl in
+ introid_a ~orig id gl
+ *)
+ | IPatCase(iorpat) ->
+ tclIORPAT ?ist (with_top (ssrscasetac false)) iorpat gl
+ | IPatInj iorpat ->
+ tclIORPAT ?ist (with_top (ssrscasetac true)) iorpat gl
+ | IPatRewrite (occ, dir) ->
+ with_top (ipat_rewrite occ dir) gl
+ | IPatId id -> tac_ctx (introid id) gl
+ | IPatNewHidden idl -> tac_ctx (ssrmkabstac idl) gl
+ | IPatSimpl sim ->
+ tac_ctx (simpltac sim) gl
+ | IPatClear clr ->
+ delayed_clear false !next clr gl
+ | IPatAnon One -> tac_ctx intro_anon gl
+ | IPatNoop -> tac_ctx tclIDTAC gl
+ | IPatView v ->
+ let ist =
+ match ist with Some x -> x | _ -> anomaly "ipat: view with no ist" in
+ let next_keeps =
+ match !next with (IPatCase _ | IPatRewrite _)::_ -> false | _ -> true in
+ let top_id = ref top_id in
+ tclTHENLIST_a [
+ (move_top_with_view ~next next_keeps top_id (next_keeps,v) ist);
+ (fun gl ->
+ let hyps = without_ctx pf_hyps gl in
+ if not next_keeps && test_hypname_exists hyps !top_id then
+ delayed_clear true !next [SsrHyp (Loc.tag !top_id)] gl
+ else tac_ctx tclIDTAC gl)]
+ gl
+
+ and tclIORPAT ?ist tac = function
+ | [[]] -> tac
+ | orp -> tclTHENS_nonstrict tac (List.map (ipatstac ?ist) orp) "intro pattern"
+
+ and ipatstac ?ist ipats gl =
+ let rec aux ipats gl =
+ match ipats with
+ | [] -> tac_ctx tclIDTAC gl
+ | p :: ps ->
+ let next = ref ps in
+ let gl = ipattac ?ist ~next p gl in
+ tac_on_all gl (aux !next)
+ in
+ aux ipats gl
+ in
+
+ let rec split_itacs ?ist ~ind tac' = function
+ | (IPatSimpl _ | IPatClear _ as spat) :: ipats' ->
+ let tac = ipattac ?ist ~next:(ref ipats') spat in
+ split_itacs ?ist ~ind (tclTHEN_a tac' tac) ipats'
+ | IPatCase iorpat :: ipats' ->
+ tclIORPAT ?ist tac' iorpat, ipats'
+ | ipats' -> tac', ipats' in
+
+ let combine_tacs tac eqtac ipats ?ist ~ind gl =
+ let tac1, ipats' = split_itacs ?ist ~ind tac ipats in
+ let tac2 = ipatstac ?ist ipats' in
+ tclTHENLIST_a [ tac1; eqtac; tac2 ] gl in
+
+ (* Exported code *)
+ let introstac ?ist ipats gl =
+ with_fresh_ctx (tclTHENLIST_a [
+ ipatstac ?ist ipats;
+ gen_tmp_ids ?ist;
+ clear_wilds_and_tmp_and_delayed_ids
+ ]) gl in
+
+ let tclEQINTROS ?(ind=ref None) ?ist tac eqtac ipats gl =
+ with_fresh_ctx (tclTHENLIST_a [
+ combine_tacs (tac_ctx tac) (tac_ctx eqtac) ipats ?ist ~ind;
+ gen_tmp_ids ?ist;
+ clear_wilds_and_tmp_and_delayed_ids;
+ ]) gl in
+
+ introstac, tclEQINTROS
+;;
+
+(* Intro patterns processing for elim tactic*)
+let elim_intro_tac ipats ?ist what eqid ssrelim is_rec clr gl =
+ (* Utils of local interest only *)
+ let iD s ?t gl = let t = match t with None -> pf_concl gl | Some x -> x in
+ ppdebug(lazy Pp.(str s ++ pr_econstr t)); Tacticals.tclIDTAC gl in
+ let protectC, gl = pf_mkSsrConst "protect_term" gl in
+ let eq, gl = pf_fresh_global (Coqlib.build_coq_eq ()) gl in
+ let eq = EConstr.of_constr eq in
+ let fire_subst gl t = Reductionops.nf_evar (project gl) t in
+ let intro_eq =
+ match eqid with
+ | Some (IPatId ipat) when not is_rec ->
+ let rec intro_eq gl = match EConstr.kind_of_type (project gl) (pf_concl gl) with
+ | ProdType (_, src, tgt) ->
+ (match EConstr.kind_of_type (project gl) src with
+ | AtomicType (hd, _) when EConstr.eq_constr (project gl) hd protectC ->
+ Tacticals.tclTHENLIST [unprotecttac; introid ipat] gl
+ | _ -> Tacticals.tclTHENLIST [ iD "IA"; Ssrcommon.intro_anon; intro_eq] gl)
+ |_ -> errorstrm (Pp.str "Too many names in intro pattern") in
+ intro_eq
+ | Some (IPatId ipat) ->
+ let name gl = mk_anon_id "K" gl in
+ let intro_lhs gl =
+ let elim_name = match clr, what with
+ | [SsrHyp(_, x)], _ -> x
+ | _, `EConstr(_,_,t) when EConstr.isVar (project gl) t -> EConstr.destVar (project gl) t
+ | _ -> name gl in
+ if is_name_in_ipats elim_name ipats then introid (name gl) gl
+ else introid elim_name gl
+ in
+ let rec gen_eq_tac gl =
+ let concl = pf_concl gl in
+ let ctx, last = EConstr.decompose_prod_assum (project gl) concl in
+ let args = match EConstr.kind_of_type (project gl) last with
+ | AtomicType (hd, args) -> assert(EConstr.eq_constr (project gl) hd protectC); args
+ | _ -> assert false in
+ let case = args.(Array.length args-1) in
+ if not(EConstr.Vars.closed0 (project gl) case) then Tacticals.tclTHEN Ssrcommon.intro_anon gen_eq_tac gl
+ else
+ let gl, case_ty = pfe_type_of gl case in
+ let refl = EConstr.mkApp (eq, [|EConstr.Vars.lift 1 case_ty; EConstr.mkRel 1; EConstr.Vars.lift 1 case|]) in
+ let new_concl = fire_subst gl
+ EConstr.(mkProd (Name (name gl), case_ty, mkArrow refl (Vars.lift 2 concl))) in
+ let erefl, gl = mkRefl case_ty case gl in
+ let erefl = fire_subst gl erefl in
+ apply_type new_concl [case;erefl] gl in
+ Tacticals.tclTHENLIST [gen_eq_tac; intro_lhs; introid ipat]
+ | _ -> Tacticals.tclIDTAC in
+ let unprot = if eqid <> None && is_rec then unprotecttac else Tacticals.tclIDTAC in
+ tclEQINTROS ?ist ssrelim (Tacticals.tclTHENLIST [intro_eq; unprot]) ipats gl
+
+(* General case *)
+let tclINTROS ist t ip = tclEQINTROS ~ist (t ist) tclIDTAC ip
+
+(* }}} *)
+
+let viewmovetac ?next v deps gen ist gl =
+ with_fresh_ctx
+ (tclTHEN_a
+ (viewmovetac_aux ?next true (ref top_id) v deps gen ist)
+ clear_wilds_and_tmp_and_delayed_ids)
+ gl
+
+let mkCoqEq gl =
+ let sigma = project gl in
+ let (sigma, eq) = EConstr.fresh_global (pf_env gl) sigma (build_coq_eq_data()).eq in
+ let gl = { gl with sigma } in
+ eq, gl
+
+let mkEq dir cl c t n gl =
+ let open EConstr in
+ let eqargs = [|t; c; c|] in eqargs.(dir_org dir) <- mkRel n;
+ let eq, gl = mkCoqEq gl in
+ let refl, gl = mkRefl t c gl in
+ mkArrow (mkApp (eq, eqargs)) (EConstr.Vars.lift 1 cl), refl, gl
+
+let pushmoveeqtac cl c gl =
+ let open EConstr in
+ let x, t, cl1 = destProd (project gl) cl in
+ let cl2, eqc, gl = mkEq R2L cl1 c t 1 gl in
+ apply_type (mkProd (x, t, cl2)) [c; eqc] gl
+
+let eqmovetac _ gen ist gl =
+ let cl, c, _, gl = pf_interp_gen ist gl false gen in pushmoveeqtac cl c gl
+
+let movehnftac gl = match EConstr.kind (project gl) (pf_concl gl) with
+ | Prod _ | LetIn _ -> tclIDTAC gl
+ | _ -> new_tac hnf_in_concl gl
+
+let rec eqmoveipats eqpat = function
+ | (IPatSimpl _ | IPatClear _ as ipat) :: ipats -> ipat :: eqmoveipats eqpat ipats
+ | (IPatAnon All :: _ | []) as ipats -> IPatAnon One :: eqpat :: ipats
+ | ipat :: ipats -> ipat :: eqpat :: ipats
+
+let ssrmovetac ist = function
+ | _::_ as view, (_, (dgens, ipats)) ->
+ let next = ref ipats in
+ let dgentac = with_dgens dgens (viewmovetac ~next (true, view)) ist in
+ tclTHEN dgentac (fun gl -> introstac ~ist !next gl)
+ | _, (Some pat, (dgens, ipats)) ->
+ let dgentac = with_dgens dgens eqmovetac ist in
+ tclTHEN dgentac (introstac ~ist (eqmoveipats pat ipats))
+ | _, (_, (([gens], clr), ipats)) ->
+ let gentac = genstac (gens, clr) ist in
+ tclTHEN gentac (introstac ~ist ipats)
+ | _, (_, ((_, clr), ipats)) ->
+ tclTHENLIST [movehnftac; cleartac clr; introstac ~ist ipats]
+
+let ssrcasetac ist (view, (eqid, (dgens, ipats))) =
+ let ndefectcasetac view eqid ipats deps ((_, occ), _ as gen) ist gl =
+ let simple = (eqid = None && deps = [] && occ = None) in
+ let cl, c, clr, gl = pf_interp_gen ist gl true gen in
+ let _,vc, gl =
+ if view = [] then c,c, gl else pf_with_view_linear ist gl (false, view) cl c in
+ if simple && is_injection_case vc gl then
+ tclTHENLIST [perform_injection vc; cleartac clr; introstac ~ist ipats] gl
+ else
+ (* macro for "case/v E: x" ---> "case E: x / (v x)" *)
+ let deps, clr, occ =
+ if view <> [] && eqid <> None && deps = [] then [gen], [], None
+ else deps, clr, occ in
+ ssrelim ~is_case:true ~ist deps (`EConstr (clr,occ, vc)) eqid (elim_intro_tac ipats) gl
+ in
+ with_dgens dgens (ndefectcasetac view eqid ipats) ist
+
+let ssrapplytac ist (views, (_, ((gens, clr), intros))) =
+ tclINTROS ist (inner_ssrapplytac views gens clr) intros
+
+
+(* vim: set filetype=ocaml foldmethod=marker: *)
diff --git a/plugins/ssr/ssripats.mli b/plugins/ssr/ssripats.mli
new file mode 100644
index 000000000..e90e75552
--- /dev/null
+++ b/plugins/ssr/ssripats.mli
@@ -0,0 +1,82 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Ssrmatching_plugin
+open Ssrast
+open Ssrcommon
+
+type block_names = (int * EConstr.types array) option
+
+(* For case/elim with eq generation: args are elim_tac introeq_tac ipats
+ * elim E : "=> ipats" where E give rise to introeq_tac *)
+val tclEQINTROS :
+ ?ind:block_names ref ->
+ ?ist:ist ->
+ v82tac ->
+ v82tac -> ssripats -> v82tac
+(* special case with no eq and tactic taking ist *)
+val tclINTROS :
+ ist ->
+ (ist -> v82tac) ->
+ ssripats -> v82tac
+
+(* move=> ipats *)
+val introstac : ?ist:ist -> ssripats -> v82tac
+
+val elim_intro_tac :
+ Ssrast.ssripats ->
+ ?ist:Tacinterp.interp_sign ->
+ [> `EConstr of 'a * 'b * EConstr.t ] ->
+ Ssrast.ssripat option ->
+ Proof_type.tactic ->
+ bool ->
+ Ssrast.ssrhyp list ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+(* "move=> top; tac top; clear top" respecting the speed *)
+val with_top : (EConstr.t -> v82tac) -> tac_ctx tac_a
+
+val ssrmovetac :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ Ssrast.ssrterm list *
+ (Ssrast.ssripat option *
+ (((Ssrast.ssrdocc * Ssrmatching.cpattern) list
+ list * Ssrast.ssrclear) *
+ Ssrast.ssripats)) ->
+ Proof_type.tactic
+
+val movehnftac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val with_dgens :
+ (Ssrast.ssrdocc * Ssrmatching.cpattern) list
+ list * Ssrast.ssrclear ->
+ ((Ssrast.ssrdocc * Ssrmatching.cpattern) list ->
+ Ssrast.ssrdocc * Ssrmatching.cpattern ->
+ Ltac_plugin.Tacinterp.interp_sign -> Proof_type.tactic) ->
+ Ltac_plugin.Tacinterp.interp_sign -> Proof_type.tactic
+
+val ssrcasetac :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ Ssrast.ssrterm list *
+ (Ssrast.ssripat option *
+ (((Ssrast.ssrdocc * Ssrmatching.cpattern) list list * Ssrast.ssrclear) *
+ Ssrast.ssripats)) ->
+ Proof_type.tactic
+
+val ssrapplytac :
+ Tacinterp.interp_sign ->
+ Ssrast.ssrterm list *
+ ('a *
+ ((((Ssrast.ssrhyps option * Ssrmatching_plugin.Ssrmatching.occ) *
+ (Ssrast.ssrtermkind * Tacexpr.glob_constr_and_expr))
+ list list * Ssrast.ssrhyps) *
+ Ssrast.ssripats)) ->
+ Proof_type.tactic
+
diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4
new file mode 100644
index 000000000..1fba39150
--- /dev/null
+++ b/plugins/ssr/ssrparser.ml4
@@ -0,0 +1,2349 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Pp
+open Pcoq
+open Ltac_plugin
+open Genarg
+open Stdarg
+open Tacarg
+open Term
+open Libnames
+open Tactics
+open Tacticals
+open Tacmach
+open Glob_term
+open Util
+open Tacexpr
+open Tacinterp
+open Pltac
+open Extraargs
+open Ppconstr
+open Printer
+
+open Misctypes
+open Decl_kinds
+open Constrexpr
+open Constrexpr_ops
+
+open Ssrprinters
+open Ssrcommon
+open Ssrtacticals
+open Ssrbwd
+open Ssrequality
+open Ssrelim
+
+(** Ssreflect load check. *)
+
+(* To allow ssrcoq to be fully compatible with the "plain" Coq, we only *)
+(* turn on its incompatible features (the new rewrite syntax, and the *)
+(* reserved identifiers) when the theory library (ssreflect.v) has *)
+(* has actually been required, or is being defined. Because this check *)
+(* needs to be done often (for each identifier lookup), we implement *)
+(* some caching, repeating the test only when the environment changes. *)
+(* We check for protect_term because it is the first constant loaded; *)
+(* ssr_have would ultimately be a better choice. *)
+let ssr_loaded = Summary.ref ~name:"SSR:loaded" false
+let is_ssr_loaded () =
+ !ssr_loaded ||
+ (if CLexer.is_keyword "SsrSyntax_is_Imported" then ssr_loaded:=true;
+ !ssr_loaded)
+
+DECLARE PLUGIN "ssreflect_plugin"
+(* 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 () ;;
+
+let tacltop = (5,Ppextend.E)
+
+let pr_ssrtacarg _ _ prt = prt tacltop
+ARGUMENT EXTEND ssrtacarg TYPED AS tactic PRINTED BY pr_ssrtacarg
+| [ "YouShouldNotTypeThis" ] -> [ CErrors.anomaly (Pp.str "Grammar placeholder match") ]
+END
+GEXTEND Gram
+ GLOBAL: ssrtacarg;
+ ssrtacarg: [[ tac = tactic_expr LEVEL "5" -> tac ]];
+END
+
+(* Lexically closed tactic for tacticals. *)
+let pr_ssrtclarg _ _ prt tac = prt tacltop tac
+ARGUMENT EXTEND ssrtclarg TYPED AS ssrtacarg
+ PRINTED BY pr_ssrtclarg
+| [ ssrtacarg(tac) ] -> [ tac ]
+END
+
+open Genarg
+
+(** Adding a new uninterpreted generic argument type *)
+let add_genarg tag pr =
+ let wit = Genarg.make0 tag in
+ let tag = Geninterp.Val.create tag in
+ let glob ist x = (ist, x) in
+ let subst _ x = x in
+ let interp ist x = Ftactic.return (Geninterp.Val.Dyn (tag, x)) in
+ let gen_pr _ _ _ = pr in
+ let () = Genintern.register_intern0 wit glob in
+ let () = Genintern.register_subst0 wit subst in
+ let () = Geninterp.register_interp0 wit interp in
+ let () = Geninterp.register_val0 wit (Some (Geninterp.Val.Base tag)) in
+ Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr;
+ wit
+
+(** Primitive parsing to avoid syntax conflicts with basic tactics. *)
+
+let accept_before_syms syms strm =
+ match Util.stream_nth 1 strm with
+ | Tok.KEYWORD sym when List.mem sym syms -> ()
+ | _ -> raise Stream.Failure
+
+let accept_before_syms_or_any_id syms strm =
+ match Util.stream_nth 1 strm with
+ | Tok.KEYWORD sym when List.mem sym syms -> ()
+ | Tok.IDENT _ -> ()
+ | _ -> raise Stream.Failure
+
+let accept_before_syms_or_ids syms ids strm =
+ match Util.stream_nth 1 strm with
+ | Tok.KEYWORD sym when List.mem sym syms -> ()
+ | Tok.IDENT id when List.mem id ids -> ()
+ | _ -> raise Stream.Failure
+
+open Ssrast
+let pr_id = Ppconstr.pr_id
+let pr_name = function Name id -> pr_id id | Anonymous -> str "_"
+let pr_spc () = str " "
+let pr_bar () = Pp.cut() ++ str "|"
+let pr_list = prlist_with_sep
+
+(**************************** ssrhyp **************************************)
+
+let pr_ssrhyp _ _ _ = pr_hyp
+
+let wit_ssrhyprep = add_genarg "ssrhyprep" pr_hyp
+
+let intern_hyp ist (SsrHyp (loc, id) as hyp) =
+ let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_var) (loc, id)) in
+ if not_section_id id then hyp else
+ hyp_err ?loc "Can't clear section hypothesis " id
+
+open Pcoq.Prim
+
+ARGUMENT EXTEND ssrhyp TYPED AS ssrhyprep PRINTED BY pr_ssrhyp
+ INTERPRETED BY interp_hyp
+ GLOBALIZED BY intern_hyp
+ | [ ident(id) ] -> [ SsrHyp (Loc.tag ~loc id) ]
+END
+
+
+let pr_hoi = hoik pr_hyp
+let pr_ssrhoi _ _ _ = pr_hoi
+
+let wit_ssrhoirep = add_genarg "ssrhoirep" pr_hoi
+
+let intern_ssrhoi ist = function
+ | Hyp h -> Hyp (intern_hyp ist h)
+ | Id (SsrHyp (_, id)) as hyp ->
+ let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_ident) id) in
+ hyp
+
+let interp_ssrhoi ist gl = function
+ | Hyp h -> let s, h' = interp_hyp ist gl h in s, Hyp h'
+ | Id (SsrHyp (loc, id)) ->
+ let s, id' = interp_wit wit_ident ist gl id in
+ s, Id (SsrHyp (loc, id'))
+
+ARGUMENT EXTEND ssrhoi_hyp TYPED AS ssrhoirep PRINTED BY pr_ssrhoi
+ INTERPRETED BY interp_ssrhoi
+ GLOBALIZED BY intern_ssrhoi
+ | [ ident(id) ] -> [ Hyp (SsrHyp(Loc.tag ~loc id)) ]
+END
+ARGUMENT EXTEND ssrhoi_id TYPED AS ssrhoirep PRINTED BY pr_ssrhoi
+ INTERPRETED BY interp_ssrhoi
+ GLOBALIZED BY intern_ssrhoi
+ | [ ident(id) ] -> [ Id (SsrHyp(Loc.tag ~loc id)) ]
+END
+
+
+let pr_hyps = pr_list pr_spc pr_hyp
+let pr_ssrhyps _ _ _ = pr_hyps
+
+ARGUMENT EXTEND ssrhyps TYPED AS ssrhyp list PRINTED BY pr_ssrhyps
+ INTERPRETED BY interp_hyps
+ | [ ssrhyp_list(hyps) ] -> [ check_hyps_uniq [] hyps; hyps ]
+END
+
+(** Rewriting direction *)
+
+
+let pr_dir = function L2R -> str "->" | R2L -> str "<-"
+let pr_rwdir = function L2R -> mt() | R2L -> str "-"
+
+let wit_ssrdir = add_genarg "ssrdir" pr_dir
+
+(** Simpl switch *)
+
+
+let pr_simpl = function
+ | Simpl -1 -> str "/="
+ | Cut -1 -> str "//"
+ | Simpl n -> str "/" ++ int n ++ str "="
+ | Cut n -> str "/" ++ int n ++ str"/"
+ | SimplCut (-1,-1) -> str "//="
+ | SimplCut (n,-1) -> str "/" ++ int n ++ str "/="
+ | SimplCut (-1,n) -> str "//" ++ int n ++ str "="
+ | SimplCut (n,m) -> str "/" ++ int n ++ str "/" ++ int m ++ str "="
+ | Nop -> mt ()
+
+let pr_ssrsimpl _ _ _ = pr_simpl
+
+let wit_ssrsimplrep = add_genarg "ssrsimplrep" pr_simpl
+
+let test_ssrslashnum b1 b2 strm =
+ match Util.stream_nth 0 strm with
+ | Tok.KEYWORD "/" ->
+ (match Util.stream_nth 1 strm with
+ | Tok.INT _ when b1 ->
+ (match Util.stream_nth 2 strm with
+ | Tok.KEYWORD "=" | Tok.KEYWORD "/=" when not b2 -> ()
+ | Tok.KEYWORD "/" ->
+ if not b2 then () else begin
+ match Util.stream_nth 3 strm with
+ | Tok.INT _ -> ()
+ | _ -> raise Stream.Failure
+ end
+ | _ -> raise Stream.Failure)
+ | Tok.KEYWORD "/" when not b1 ->
+ (match Util.stream_nth 2 strm with
+ | Tok.KEYWORD "=" when not b2 -> ()
+ | Tok.INT _ when b2 ->
+ (match Util.stream_nth 3 strm with
+ | Tok.KEYWORD "=" -> ()
+ | _ -> raise Stream.Failure)
+ | _ when not b2 -> ()
+ | _ -> raise Stream.Failure)
+ | Tok.KEYWORD "=" when not b1 && not b2 -> ()
+ | _ -> raise Stream.Failure)
+ | Tok.KEYWORD "//" when not b1 ->
+ (match Util.stream_nth 1 strm with
+ | Tok.KEYWORD "=" when not b2 -> ()
+ | Tok.INT _ when b2 ->
+ (match Util.stream_nth 2 strm with
+ | Tok.KEYWORD "=" -> ()
+ | _ -> raise Stream.Failure)
+ | _ when not b2 -> ()
+ | _ -> raise Stream.Failure)
+ | _ -> raise Stream.Failure
+
+let test_ssrslashnum10 = test_ssrslashnum true false
+let test_ssrslashnum11 = test_ssrslashnum true true
+let test_ssrslashnum01 = test_ssrslashnum false true
+let test_ssrslashnum00 = test_ssrslashnum false false
+
+let negate_parser f x =
+ let rc = try Some (f x) with Stream.Failure -> None in
+ match rc with
+ | None -> ()
+ | Some _ -> raise Stream.Failure
+
+let test_not_ssrslashnum =
+ Pcoq.Gram.Entry.of_parser
+ "test_not_ssrslashnum" (negate_parser test_ssrslashnum10)
+let test_ssrslashnum00 =
+ Pcoq.Gram.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum00
+let test_ssrslashnum10 =
+ Pcoq.Gram.Entry.of_parser "test_ssrslashnum10" test_ssrslashnum10
+let test_ssrslashnum11 =
+ Pcoq.Gram.Entry.of_parser "test_ssrslashnum11" test_ssrslashnum11
+let test_ssrslashnum01 =
+ Pcoq.Gram.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum01
+
+
+ARGUMENT EXTEND ssrsimpl_ne TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl
+| [ "//=" ] -> [ SimplCut (~-1,~-1) ]
+| [ "/=" ] -> [ Simpl ~-1 ]
+END
+
+Pcoq.(Prim.(
+GEXTEND Gram
+ GLOBAL: ssrsimpl_ne;
+ ssrsimpl_ne: [
+ [ test_ssrslashnum11; "/"; n = natural; "/"; m = natural; "=" -> SimplCut(n,m)
+ | test_ssrslashnum10; "/"; n = natural; "/" -> Cut n
+ | test_ssrslashnum10; "/"; n = natural; "=" -> Simpl n
+ | test_ssrslashnum10; "/"; n = natural; "/=" -> SimplCut (n,~-1)
+ | test_ssrslashnum10; "/"; n = natural; "/"; "=" -> SimplCut (n,~-1)
+ | test_ssrslashnum01; "//"; m = natural; "=" -> SimplCut (~-1,m)
+ | test_ssrslashnum00; "//" -> Cut ~-1
+ ]];
+
+END
+))
+
+ARGUMENT EXTEND ssrsimpl TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl
+| [ ssrsimpl_ne(sim) ] -> [ sim ]
+| [ ] -> [ Nop ]
+END
+
+let pr_clear_ne clr = str "{" ++ pr_hyps clr ++ str "}"
+let pr_clear sep clr = if clr = [] then mt () else sep () ++ pr_clear_ne clr
+
+let pr_ssrclear _ _ _ = pr_clear mt
+
+ARGUMENT EXTEND ssrclear_ne TYPED AS ssrhyps PRINTED BY pr_ssrclear
+| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ check_hyps_uniq [] clr; clr ]
+END
+
+ARGUMENT EXTEND ssrclear TYPED AS ssrclear_ne PRINTED BY pr_ssrclear
+| [ ssrclear_ne(clr) ] -> [ clr ]
+| [ ] -> [ [] ]
+END
+
+(** Indexes *)
+
+(* Since SSR indexes are always positive numbers, we use the 0 value *)
+(* to encode an omitted index. We reuse the in or_var type, but we *)
+(* supply our own interpretation function, which checks for non *)
+(* positive values, and allows the use of constr numerals, so that *)
+(* e.g., "let n := eval compute in (1 + 3) in (do n!clear)" works. *)
+
+
+let pr_index = function
+ | Misctypes.ArgVar (_, id) -> pr_id id
+ | Misctypes.ArgArg n when n > 0 -> int n
+ | _ -> mt ()
+let pr_ssrindex _ _ _ = pr_index
+
+let noindex = Misctypes.ArgArg 0
+
+let check_index ?loc i =
+ if i > 0 then i else CErrors.user_err ?loc (str"Index not positive")
+let mk_index ?loc = function
+ | Misctypes.ArgArg i -> Misctypes.ArgArg (check_index ?loc i)
+ | iv -> iv
+
+let interp_index ist gl idx =
+ Tacmach.project gl,
+ match idx with
+ | Misctypes.ArgArg _ -> idx
+ | Misctypes.ArgVar (loc, id) ->
+ let i =
+ try
+ let v = Id.Map.find id ist.Tacinterp.lfun in
+ begin match Tacinterp.Value.to_int v with
+ | Some i -> i
+ | None ->
+ begin match Tacinterp.Value.to_constr v with
+ | 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)
+ | _ -> raise Not_found
+ end
+ | None -> raise Not_found
+ end end
+ with _ -> CErrors.user_err ?loc (str"Index not a number") in
+ Misctypes.ArgArg (check_index ?loc i)
+
+open Pltac
+
+ARGUMENT EXTEND ssrindex TYPED AS ssrindex PRINTED BY pr_ssrindex
+ INTERPRETED BY interp_index
+| [ int_or_var(i) ] -> [ mk_index ~loc i ]
+END
+
+
+(** Occurrence switch *)
+
+(* The standard syntax of complemented occurrence lists involves a single *)
+(* initial "-", e.g., {-1 3 5}. An initial *)
+(* "+" may be used to indicate positive occurrences (the default). The *)
+(* "+" is optional, except if the list of occurrences starts with a *)
+(* variable or is empty (to avoid confusion with a clear switch). The *)
+(* empty positive switch "{+}" selects no occurrences, while the empty *)
+(* negative switch "{-}" selects all occurrences explicitly; this is the *)
+(* default, but "{-}" prevents the implicit clear, and can be used to *)
+(* force dependent elimination -- see ndefectelimtac below. *)
+
+
+let pr_ssrocc _ _ _ = pr_occ
+
+open Pcoq.Prim
+
+ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY pr_ssrocc
+| [ natural(n) natural_list(occ) ] -> [
+ Some (false, List.map (check_index ~loc) (n::occ)) ]
+| [ "-" natural_list(occ) ] -> [ Some (true, occ) ]
+| [ "+" natural_list(occ) ] -> [ Some (false, occ) ]
+END
+
+
+(* modality *)
+
+
+let pr_mmod = function May -> str "?" | Must -> str "!" | Once -> mt ()
+
+let wit_ssrmmod = add_genarg "ssrmmod" pr_mmod
+let ssrmmod = Pcoq.create_generic_entry Pcoq.utactic "ssrmmod" (Genarg.rawwit wit_ssrmmod);;
+
+GEXTEND Gram
+ GLOBAL: ssrmmod;
+ ssrmmod: [[ "!" -> Must | LEFTQMARK -> May | "?" -> May]];
+END
+
+(** Rewrite multiplier: !n ?n *)
+
+let pr_mult (n, m) =
+ if n > 0 && m <> Once then int n ++ pr_mmod m else pr_mmod m
+
+let pr_ssrmult _ _ _ = pr_mult
+
+ARGUMENT EXTEND ssrmult_ne TYPED AS int * ssrmmod PRINTED BY pr_ssrmult
+ | [ natural(n) ssrmmod(m) ] -> [ check_index ~loc n, m ]
+ | [ ssrmmod(m) ] -> [ notimes, m ]
+END
+
+ARGUMENT EXTEND ssrmult TYPED AS ssrmult_ne PRINTED BY pr_ssrmult
+ | [ ssrmult_ne(m) ] -> [ m ]
+ | [ ] -> [ nomult ]
+END
+
+(** Discharge occ switch (combined occurrence / clear switch *)
+
+let pr_docc = function
+ | None, occ -> pr_occ occ
+ | Some clr, _ -> pr_clear mt clr
+
+let pr_ssrdocc _ _ _ = pr_docc
+
+ARGUMENT EXTEND ssrdocc TYPED AS ssrclear option * ssrocc PRINTED BY pr_ssrdocc
+| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ mkclr clr ]
+| [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ]
+END
+
+(* kinds of terms *)
+
+let input_ssrtermkind strm = match Util.stream_nth 0 strm with
+ | Tok.KEYWORD "(" -> xInParens
+ | Tok.KEYWORD "@" -> xWithAt
+ | _ -> xNoFlag
+
+let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind
+
+(* terms *)
+
+(** Terms parsing. ********************************************************)
+
+let interp_constr = interp_wit wit_constr
+
+(* Because we allow wildcards in term references, we need to stage the *)
+(* interpretation of terms so that it occurs at the right time during *)
+(* the execution of the tactic (e.g., so that we don't report an error *)
+(* for a term that isn't actually used in the execution). *)
+(* The term representation tracks whether the concrete initial term *)
+(* started with an opening paren, which might avoid a conflict between *)
+(* the ssrreflect term syntax and Gallina notation. *)
+
+(* terms *)
+let pr_ssrterm _ _ _ = pr_term
+let force_term ist gl (_, c) = interp_constr ist gl c
+let glob_ssrterm gs = function
+ | k, (_, Some c) -> k, Tacintern.intern_constr gs c
+ | ct -> ct
+let subst_ssrterm s (k, c) = k, Tacsubst.subst_glob_constr_and_expr s c
+let interp_ssrterm _ gl t = Tacmach.project gl, t
+
+open Pcoq.Constr
+
+ARGUMENT EXTEND ssrterm
+ PRINTED BY pr_ssrterm
+ INTERPRETED BY interp_ssrterm
+ GLOBALIZED BY glob_ssrterm SUBSTITUTED BY subst_ssrterm
+ RAW_PRINTED BY pr_ssrterm
+ GLOB_PRINTED BY pr_ssrterm
+| [ "YouShouldNotTypeThis" constr(c) ] -> [ mk_lterm c ]
+END
+
+
+GEXTEND Gram
+ GLOBAL: ssrterm;
+ ssrterm: [[ k = ssrtermkind; c = Pcoq.Constr.constr -> mk_term k c ]];
+END
+
+(* Views *)
+
+let pr_view = pr_list mt (fun c -> str "/" ++ pr_term c)
+
+let pr_ssrview _ _ _ = pr_view
+
+ARGUMENT EXTEND ssrview TYPED AS ssrterm list
+ PRINTED BY pr_ssrview
+| [ "YouShouldNotTypeThis" ] -> [ [] ]
+END
+
+Pcoq.(
+GEXTEND Gram
+ GLOBAL: ssrview;
+ ssrview: [
+ [ test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr -> [mk_term xNoFlag c]
+ | test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr; w = ssrview ->
+ (mk_term xNoFlag c) :: w ]];
+END
+)
+
+(* }}} *)
+
+(* ipats *)
+
+
+let remove_loc = snd
+
+let ipat_of_intro_pattern p = Misctypes.(
+ let rec ipat_of_intro_pattern = function
+ | IntroNaming (IntroIdentifier id) -> IPatId id
+ | IntroAction IntroWildcard -> IPatAnon Drop
+ | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) ->
+ IPatCase
+ (List.map (List.map ipat_of_intro_pattern)
+ (List.map (List.map remove_loc) iorpat))
+ | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) ->
+ IPatCase
+ [List.map ipat_of_intro_pattern (List.map remove_loc iandpat)]
+ | IntroNaming IntroAnonymous -> IPatAnon One
+ | IntroAction (IntroRewrite b) -> IPatRewrite (allocc, if b then L2R else R2L)
+ | IntroNaming (IntroFresh id) -> IPatAnon One
+ | IntroAction (IntroApplyOn _) -> (* to do *) CErrors.user_err (Pp.str "TO DO")
+ | IntroAction (IntroInjection ips) ->
+ IPatInj [List.map ipat_of_intro_pattern (List.map remove_loc ips)]
+ | IntroForthcoming _ ->
+ (* Unable to determine which kind of ipat interp_introid could
+ * return [HH] *)
+ assert false
+ in
+ ipat_of_intro_pattern p
+)
+
+let rec pr_ipat p =
+ match p with
+ | IPatId id -> pr_id id
+ | IPatSimpl sim -> pr_simpl sim
+ | IPatClear clr -> pr_clear mt clr
+ | IPatCase iorpat -> hov 1 (str "[" ++ pr_iorpat iorpat ++ str "]")
+ | IPatInj iorpat -> hov 1 (str "[=" ++ pr_iorpat iorpat ++ str "]")
+ | IPatRewrite (occ, dir) -> pr_occ occ ++ pr_dir dir
+ | IPatAnon All -> str "*"
+ | IPatAnon Drop -> str "_"
+ | IPatAnon One -> str "?"
+ | IPatView v -> pr_view v
+ | IPatNoop -> str "-"
+ | IPatNewHidden l -> str "[:" ++ pr_list spc pr_id l ++ str "]"
+(* TODO | IPatAnon Temporary -> str "+" *)
+
+and pr_iorpat iorpat = pr_list pr_bar pr_ipats iorpat
+and pr_ipats ipats = pr_list spc pr_ipat ipats
+
+let wit_ssripatrep = add_genarg "ssripatrep" pr_ipat
+
+let pr_ssripat _ _ _ = pr_ipat
+let pr_ssripats _ _ _ = pr_ipats
+let pr_ssriorpat _ _ _ = pr_iorpat
+
+let intern_ipat ist ipat =
+ let rec check_pat = function
+ | IPatClear clr -> ignore (List.map (intern_hyp ist) clr)
+ | IPatCase iorpat -> List.iter (List.iter check_pat) iorpat
+ | IPatInj iorpat -> List.iter (List.iter check_pat) iorpat
+ | _ -> () in
+ check_pat ipat; ipat
+
+let intern_ipats ist = List.map (intern_ipat ist)
+
+let interp_intro_pattern = interp_wit wit_intro_pattern
+
+let interp_introid ist gl id = Misctypes.(
+ try IntroNaming (IntroIdentifier (hyp_id (snd (interp_hyp ist gl (SsrHyp (Loc.tag id))))))
+ with _ -> snd(snd (interp_intro_pattern ist gl (Loc.tag @@ IntroNaming (IntroIdentifier id))))
+)
+
+let rec add_intro_pattern_hyps (loc, ipat) hyps = Misctypes.(
+ match ipat with
+ | IntroNaming (IntroIdentifier id) ->
+ if not_section_id id then SsrHyp (loc, id) :: hyps else
+ hyp_err ?loc "Can't delete section hypothesis " id
+ | IntroAction IntroWildcard -> hyps
+ | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) ->
+ List.fold_right (List.fold_right add_intro_pattern_hyps) iorpat hyps
+ | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) ->
+ List.fold_right add_intro_pattern_hyps iandpat hyps
+ | IntroNaming IntroAnonymous -> []
+ | IntroNaming (IntroFresh _) -> []
+ | IntroAction (IntroRewrite _) -> hyps
+ | IntroAction (IntroInjection ips) -> List.fold_right add_intro_pattern_hyps ips hyps
+ | IntroAction (IntroApplyOn (c,pat)) -> add_intro_pattern_hyps pat hyps
+ | IntroForthcoming _ ->
+ (* As in ipat_of_intro_pattern, was unable to determine which kind
+ of ipat interp_introid could return [HH] *) assert false
+)
+
+(* MD: what does this do? *)
+let interp_ipat ist gl = Misctypes.(
+ let ltacvar id = Id.Map.mem id ist.Tacinterp.lfun in
+ let rec interp = function
+ | IPatId id when ltacvar id ->
+ ipat_of_intro_pattern (interp_introid ist gl id)
+ | IPatClear clr ->
+ let add_hyps (SsrHyp (loc, id) as hyp) hyps =
+ if not (ltacvar id) then hyp :: hyps else
+ add_intro_pattern_hyps (loc, (interp_introid ist gl id)) hyps in
+ let clr' = List.fold_right add_hyps clr [] in
+ check_hyps_uniq [] clr'; IPatClear clr'
+ | IPatCase(iorpat) ->
+ IPatCase(List.map (List.map interp) iorpat)
+ | IPatInj iorpat -> IPatInj (List.map (List.map interp) iorpat)
+ | IPatNewHidden l ->
+ IPatNewHidden
+ (List.map (function
+ | IntroNaming (IntroIdentifier id) -> id
+ | _ -> assert false)
+ (List.map (interp_introid ist gl) l))
+ | ipat -> ipat in
+ interp
+)
+
+let interp_ipats ist gl l = project gl, List.map (interp_ipat ist gl) l
+
+let pushIPatRewrite = function
+ | pats :: orpat -> (IPatRewrite (allocc, L2R) :: pats) :: orpat
+ | [] -> []
+
+let pushIPatNoop = function
+ | pats :: orpat -> (IPatNoop :: pats) :: orpat
+ | [] -> []
+
+ARGUMENT EXTEND ssripat TYPED AS ssripatrep list PRINTED BY pr_ssripats
+ INTERPRETED BY interp_ipats
+ GLOBALIZED BY intern_ipats
+ | [ "_" ] -> [ [IPatAnon Drop] ]
+ | [ "*" ] -> [ [IPatAnon All] ]
+ (*
+ | [ "^" "*" ] -> [ [IPatFastMode] ]
+ | [ "^" "_" ] -> [ [IPatSeed `Wild] ]
+ | [ "^_" ] -> [ [IPatSeed `Wild] ]
+ | [ "^" "?" ] -> [ [IPatSeed `Anon] ]
+ | [ "^?" ] -> [ [IPatSeed `Anon] ]
+ | [ "^" ident(id) ] -> [ [IPatSeed (`Id(id,`Pre))] ]
+ | [ "^" "~" ident(id) ] -> [ [IPatSeed (`Id(id,`Post))] ]
+ | [ "^~" ident(id) ] -> [ [IPatSeed (`Id(id,`Post))] ]
+ *)
+ | [ ident(id) ] -> [ [IPatId id] ]
+ | [ "?" ] -> [ [IPatAnon One] ]
+(* TODO | [ "+" ] -> [ [IPatAnon Temporary] ] *)
+ | [ ssrsimpl_ne(sim) ] -> [ [IPatSimpl sim] ]
+ | [ ssrdocc(occ) "->" ] -> [ match occ with
+ | None, occ -> [IPatRewrite (occ, L2R)]
+ | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, L2R)]]
+ | [ ssrdocc(occ) "<-" ] -> [ match occ with
+ | None, occ -> [IPatRewrite (occ, R2L)]
+ | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, R2L)]]
+ | [ ssrdocc(occ) ] -> [ match occ with
+ | Some cl, _ -> check_hyps_uniq [] cl; [IPatClear cl]
+ | _ -> CErrors.user_err ~loc (str"Only identifiers are allowed here")]
+ | [ "->" ] -> [ [IPatRewrite (allocc, L2R)] ]
+ | [ "<-" ] -> [ [IPatRewrite (allocc, R2L)] ]
+ | [ "-" ] -> [ [IPatNoop] ]
+ | [ "-/" "=" ] -> [ [IPatNoop;IPatSimpl(Simpl ~-1)] ]
+ | [ "-/=" ] -> [ [IPatNoop;IPatSimpl(Simpl ~-1)] ]
+ | [ "-/" "/" ] -> [ [IPatNoop;IPatSimpl(Cut ~-1)] ]
+ | [ "-//" ] -> [ [IPatNoop;IPatSimpl(Cut ~-1)] ]
+ | [ "-/" integer(n) "/" ] -> [ [IPatNoop;IPatSimpl(Cut n)] ]
+ | [ "-/" "/=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] ]
+ | [ "-//" "=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] ]
+ | [ "-//=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] ]
+ | [ "-/" integer(n) "/=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (n,~-1))] ]
+ | [ "-/" integer(n) "/" integer (m) "=" ] ->
+ [ [IPatNoop;IPatSimpl(SimplCut(n,m))] ]
+ | [ ssrview(v) ] -> [ [IPatView v] ]
+ | [ "[" ":" ident_list(idl) "]" ] -> [ [IPatNewHidden idl] ]
+ | [ "[:" ident_list(idl) "]" ] -> [ [IPatNewHidden idl] ]
+END
+
+ARGUMENT EXTEND ssripats TYPED AS ssripat PRINTED BY pr_ssripats
+ | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ]
+ | [ ] -> [ [] ]
+END
+
+ARGUMENT EXTEND ssriorpat TYPED AS ssripat list PRINTED BY pr_ssriorpat
+| [ ssripats(pats) "|" ssriorpat(orpat) ] -> [ pats :: orpat ]
+| [ ssripats(pats) "|-" ">" ssriorpat(orpat) ] -> [ pats :: pushIPatRewrite orpat ]
+| [ ssripats(pats) "|-" ssriorpat(orpat) ] -> [ pats :: pushIPatNoop orpat ]
+| [ ssripats(pats) "|->" ssriorpat(orpat) ] -> [ pats :: pushIPatRewrite orpat ]
+| [ ssripats(pats) "||" ssriorpat(orpat) ] -> [ pats :: [] :: orpat ]
+| [ ssripats(pats) "|||" ssriorpat(orpat) ] -> [ pats :: [] :: [] :: orpat ]
+| [ ssripats(pats) "||||" ssriorpat(orpat) ] -> [ [pats; []; []; []] @ orpat ]
+| [ ssripats(pats) ] -> [ [pats] ]
+END
+
+let reject_ssrhid strm =
+ match Util.stream_nth 0 strm with
+ | Tok.KEYWORD "[" ->
+ (match Util.stream_nth 1 strm with
+ | Tok.KEYWORD ":" -> raise Stream.Failure
+ | _ -> ())
+ | _ -> ()
+
+let test_nohidden = Pcoq.Gram.Entry.of_parser "test_ssrhid" reject_ssrhid
+
+ARGUMENT EXTEND ssrcpat TYPED AS ssripatrep PRINTED BY pr_ssripat
+ | [ "YouShouldNotTypeThis" ssriorpat(x) ] -> [ IPatCase(x) ]
+END
+
+Pcoq.(
+GEXTEND Gram
+ GLOBAL: ssrcpat;
+ ssrcpat: [
+ [ test_nohidden; "["; iorpat = ssriorpat; "]" ->
+(* check_no_inner_seed !@loc false iorpat;
+ IPatCase (understand_case_type iorpat) *)
+ IPatCase iorpat
+ | test_nohidden; "[="; iorpat = ssriorpat; "]" ->
+(* check_no_inner_seed !@loc false iorpat; *)
+ IPatInj iorpat ]];
+END
+);;
+
+Pcoq.(
+GEXTEND Gram
+ GLOBAL: ssripat;
+ ssripat: [[ pat = ssrcpat -> [pat] ]];
+END
+)
+
+ARGUMENT EXTEND ssripats_ne TYPED AS ssripat PRINTED BY pr_ssripats
+ | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ]
+ END
+
+(* subsets of patterns *)
+
+(* TODO: review what this function does, it looks suspicious *)
+let check_ssrhpats loc w_binders ipats =
+ let err_loc s = CErrors.user_err ~loc ~hdr:"ssreflect" s in
+ let clr, ipats =
+ let rec aux clr = function
+ | IPatClear cl :: tl -> aux (clr @ cl) tl
+(* | IPatSimpl (cl, sim) :: tl -> clr @ cl, IPatSimpl ([], sim) :: tl *)
+ | tl -> clr, tl
+ in aux [] ipats in
+ let simpl, ipats =
+ match List.rev ipats with
+ | IPatSimpl _ as s :: tl -> [s], List.rev tl
+ | _ -> [], ipats in
+ if simpl <> [] && not w_binders then
+ err_loc (str "No s-item allowed here: " ++ pr_ipats simpl);
+ let ipat, binders =
+ let rec loop ipat = function
+ | [] -> ipat, []
+ | ( IPatId _| IPatAnon _| IPatCase _| IPatRewrite _ as i) :: tl ->
+ if w_binders then
+ if simpl <> [] && tl <> [] then
+ err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl))
+ else if not (List.for_all (function IPatId _ -> true | _ -> false) tl)
+ then err_loc (str "Only binders allowed here: " ++ pr_ipats tl)
+ else ipat @ [i], tl
+ else
+ if tl = [] then ipat @ [i], []
+ else err_loc (str "No binder or s-item allowed here: " ++ pr_ipats tl)
+ | hd :: tl -> loop (ipat @ [hd]) tl
+ in loop [] ipats in
+ ((clr, ipat), binders), simpl
+
+let pr_hpats (((clr, ipat), binders), simpl) =
+ pr_clear mt clr ++ pr_ipats ipat ++ pr_ipats binders ++ pr_ipats simpl
+let pr_ssrhpats _ _ _ = pr_hpats
+let pr_ssrhpats_wtransp _ _ _ (_, x) = pr_hpats x
+
+ARGUMENT EXTEND ssrhpats TYPED AS ((ssrclear * ssripat) * ssripat) * ssripat
+PRINTED BY pr_ssrhpats
+ | [ ssripats(i) ] -> [ check_ssrhpats loc true i ]
+END
+
+ARGUMENT EXTEND ssrhpats_wtransp
+ TYPED AS bool * (((ssrclear * ssripats) * ssripats) * ssripats)
+ PRINTED BY pr_ssrhpats_wtransp
+ | [ ssripats(i) ] -> [ false,check_ssrhpats loc true i ]
+ | [ ssripats(i) "@" ssripats(j) ] -> [ true,check_ssrhpats loc true (i @ j) ]
+END
+
+ARGUMENT EXTEND ssrhpats_nobs
+TYPED AS ((ssrclear * ssripats) * ssripats) * ssripats PRINTED BY pr_ssrhpats
+ | [ ssripats(i) ] -> [ check_ssrhpats loc false i ]
+END
+
+ARGUMENT EXTEND ssrrpat TYPED AS ssripatrep PRINTED BY pr_ssripat
+ | [ "->" ] -> [ IPatRewrite (allocc, L2R) ]
+ | [ "<-" ] -> [ IPatRewrite (allocc, R2L) ]
+END
+
+let pr_intros sep intrs =
+ if intrs = [] then mt() else sep () ++ str "=>" ++ pr_ipats intrs
+let pr_ssrintros _ _ _ = pr_intros mt
+
+ARGUMENT EXTEND ssrintros_ne TYPED AS ssripat
+ PRINTED BY pr_ssrintros
+ | [ "=>" ssripats_ne(pats) ] -> [ pats ]
+(* TODO | [ "=>" ">" ssripats_ne(pats) ] -> [ IPatFastMode :: pats ]
+ | [ "=>>" ssripats_ne(pats) ] -> [ IPatFastMode :: pats ] *)
+END
+
+ARGUMENT EXTEND ssrintros TYPED AS ssrintros_ne PRINTED BY pr_ssrintros
+ | [ ssrintros_ne(intrs) ] -> [ intrs ]
+ | [ ] -> [ [] ]
+END
+
+let pr_ssrintrosarg _ _ prt (tac, ipats) =
+ prt tacltop tac ++ pr_intros spc ipats
+
+ARGUMENT EXTEND ssrintrosarg TYPED AS tactic * ssrintros
+ PRINTED BY pr_ssrintrosarg
+| [ "YouShouldNotTypeThis" ssrtacarg(arg) ssrintros_ne(ipats) ] -> [ arg, ipats ]
+END
+
+TACTIC EXTEND ssrtclintros
+| [ "YouShouldNotTypeThis" ssrintrosarg(arg) ] ->
+ [ let tac, intros = arg in
+ Proofview.V82.tactic (Ssripats.tclINTROS ist (fun ist -> ssrevaltac ist tac) intros) ]
+ END
+
+(** Defined identifier *)
+let pr_ssrfwdid id = pr_spc () ++ pr_id id
+
+let pr_ssrfwdidx _ _ _ = pr_ssrfwdid
+
+(* We use a primitive parser for the head identifier of forward *)
+(* tactis to avoid syntactic conflicts with basic Coq tactics. *)
+ARGUMENT EXTEND ssrfwdid TYPED AS ident PRINTED BY pr_ssrfwdidx
+ | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+let accept_ssrfwdid strm =
+ match stream_nth 0 strm with
+ | Tok.IDENT id -> accept_before_syms_or_any_id [":"; ":="; "("] strm
+ | _ -> raise Stream.Failure
+
+
+let test_ssrfwdid = Gram.Entry.of_parser "test_ssrfwdid" accept_ssrfwdid
+
+GEXTEND Gram
+ GLOBAL: ssrfwdid;
+ ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> id ]];
+ END
+
+
+(* by *)
+(** Tactical arguments. *)
+
+(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *)
+(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *)
+(* and subgoal reordering tacticals (; first & ; last), respectively. *)
+
+
+let pr_ortacs prt =
+ let rec pr_rec = function
+ | [None] -> spc() ++ str "|" ++ spc()
+ | None :: tacs -> spc() ++ str "|" ++ pr_rec tacs
+ | Some tac :: tacs -> spc() ++ str "| " ++ prt tacltop tac ++ pr_rec tacs
+ | [] -> mt() in
+ function
+ | [None] -> spc()
+ | None :: tacs -> pr_rec tacs
+ | Some tac :: tacs -> prt tacltop tac ++ pr_rec tacs
+ | [] -> mt()
+let pr_ssrortacs _ _ = pr_ortacs
+
+ARGUMENT EXTEND ssrortacs TYPED AS tactic option list PRINTED BY pr_ssrortacs
+| [ ssrtacarg(tac) "|" ssrortacs(tacs) ] -> [ Some tac :: tacs ]
+| [ ssrtacarg(tac) "|" ] -> [ [Some tac; None] ]
+| [ ssrtacarg(tac) ] -> [ [Some tac] ]
+| [ "|" ssrortacs(tacs) ] -> [ None :: tacs ]
+| [ "|" ] -> [ [None; None] ]
+END
+
+let pr_hintarg prt = function
+ | true, tacs -> hv 0 (str "[ " ++ pr_ortacs prt tacs ++ str " ]")
+ | false, [Some tac] -> prt tacltop tac
+ | _, _ -> mt()
+
+let pr_ssrhintarg _ _ = pr_hintarg
+
+
+ARGUMENT EXTEND ssrhintarg TYPED AS bool * ssrortacs PRINTED BY pr_ssrhintarg
+| [ "[" "]" ] -> [ nullhint ]
+| [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ]
+| [ ssrtacarg(arg) ] -> [ mk_hint arg ]
+END
+
+ARGUMENT EXTEND ssrortacarg TYPED AS ssrhintarg PRINTED BY pr_ssrhintarg
+| [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ]
+END
+
+
+let pr_hint prt arg =
+ if arg = nohint then mt() else str "by " ++ pr_hintarg prt arg
+let pr_ssrhint _ _ = pr_hint
+
+ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY pr_ssrhint
+| [ ] -> [ nohint ]
+END
+(** The "in" pseudo-tactical {{{ **********************************************)
+
+(* We can't make "in" into a general tactical because this would create a *)
+(* crippling conflict with the ltac let .. in construct. Hence, we add *)
+(* explicitly an "in" suffix to all the extended tactics for which it is *)
+(* relevant (including move, case, elim) and to the extended do tactical *)
+(* below, which yields a general-purpose "in" of the form do [...] in ... *)
+
+(* This tactical needs to come before the intro tactics because the latter *)
+(* must take precautions in order not to interfere with the discharged *)
+(* assumptions. This is especially difficult for discharged "let"s, which *)
+(* the default simpl and unfold tactics would erase blindly. *)
+
+open Ssrmatching_plugin.Ssrmatching
+
+let pr_wgen = function
+ | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id
+ | (clr, Some((id,k),Some p)) ->
+ spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hoi id ++ str ":=" ++
+ pr_cpattern p ++ str ")"
+ | (clr, None) -> spc () ++ pr_clear mt clr
+let pr_ssrwgen _ _ _ = pr_wgen
+
+(* no globwith for char *)
+ARGUMENT EXTEND ssrwgen
+ TYPED AS ssrclear * ((ssrhoi_hyp * string) * cpattern option) option
+ PRINTED BY pr_ssrwgen
+| [ ssrclear_ne(clr) ] -> [ clr, None ]
+| [ ssrhoi_hyp(hyp) ] -> [ [], Some((hyp, " "), None) ]
+| [ "@" ssrhoi_hyp(hyp) ] -> [ [], Some((hyp, "@"), None) ]
+| [ "(" ssrhoi_id(id) ":=" lcpattern(p) ")" ] ->
+ [ [], Some ((id," "),Some p) ]
+| [ "(" ssrhoi_id(id) ")" ] -> [ [], Some ((id,"("), None) ]
+| [ "(@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] ->
+ [ [], Some ((id,"@"),Some p) ]
+| [ "(" "@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] ->
+ [ [], Some ((id,"@"),Some p) ]
+END
+
+let pr_clseq = function
+ | InGoal | InHyps -> mt ()
+ | InSeqGoal -> str "|- *"
+ | InHypsSeqGoal -> str " |- *"
+ | InHypsGoal -> str " *"
+ | InAll -> str "*"
+ | InHypsSeq -> str " |-"
+ | InAllHyps -> str "* |-"
+
+let wit_ssrclseq = add_genarg "ssrclseq" pr_clseq
+let pr_clausehyps = pr_list pr_spc pr_wgen
+let pr_ssrclausehyps _ _ _ = pr_clausehyps
+
+ARGUMENT EXTEND ssrclausehyps
+TYPED AS ssrwgen list PRINTED BY pr_ssrclausehyps
+| [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> [ hyp :: hyps ]
+| [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> [ hyp :: hyps ]
+| [ ssrwgen(hyp) ] -> [ [hyp] ]
+END
+
+(* type ssrclauses = ssrahyps * ssrclseq *)
+
+let pr_clauses (hyps, clseq) =
+ if clseq = InGoal then mt ()
+ else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq
+let pr_ssrclauses _ _ _ = pr_clauses
+
+ARGUMENT EXTEND ssrclauses TYPED AS ssrwgen list * ssrclseq
+ PRINTED BY pr_ssrclauses
+ | [ "in" ssrclausehyps(hyps) "|-" "*" ] -> [ hyps, InHypsSeqGoal ]
+ | [ "in" ssrclausehyps(hyps) "|-" ] -> [ hyps, InHypsSeq ]
+ | [ "in" ssrclausehyps(hyps) "*" ] -> [ hyps, InHypsGoal ]
+ | [ "in" ssrclausehyps(hyps) ] -> [ hyps, InHyps ]
+ | [ "in" "|-" "*" ] -> [ [], InSeqGoal ]
+ | [ "in" "*" ] -> [ [], InAll ]
+ | [ "in" "*" "|-" ] -> [ [], InAllHyps ]
+ | [ ] -> [ [], InGoal ]
+END
+
+
+
+
+(** Definition value formatting *)
+
+(* We use an intermediate structure to correctly render the binder list *)
+(* abbreviations. We use a list of hints to extract the binders and *)
+(* base term from a term, for the two first levels of representation of *)
+(* of constr terms. *)
+
+let pr_binder prl = function
+ | Bvar x ->
+ pr_name x
+ | Bdecl (xs, t) ->
+ str "(" ++ pr_list pr_spc pr_name xs ++ str " : " ++ prl t ++ str ")"
+ | Bdef (x, None, v) ->
+ str "(" ++ pr_name x ++ str " := " ++ prl v ++ str ")"
+ | Bdef (x, Some t, v) ->
+ str "(" ++ pr_name x ++ str " : " ++ prl t ++
+ str " := " ++ prl v ++ str ")"
+ | Bstruct x ->
+ str "{struct " ++ pr_name x ++ str "}"
+ | Bcast t ->
+ str ": " ++ prl t
+
+let rec mkBstruct i = function
+ | Bvar x :: b ->
+ if i = 0 then [Bstruct x] else mkBstruct (i - 1) b
+ | Bdecl (xs, _) :: b ->
+ let i' = i - List.length xs in
+ if i' < 0 then [Bstruct (List.nth xs i)] else mkBstruct i' b
+ | _ :: b -> mkBstruct i b
+ | [] -> []
+
+let rec format_local_binders h0 bl0 = match h0, bl0 with
+ | BFvar :: h, CLocalAssum ([_, x], _, _) :: bl ->
+ Bvar x :: format_local_binders h bl
+ | BFdecl _ :: h, CLocalAssum (lxs, _, t) :: bl ->
+ Bdecl (List.map snd lxs, t) :: format_local_binders h bl
+ | BFdef :: h, CLocalDef ((_, x), v, oty) :: bl ->
+ Bdef (x, oty, v) :: format_local_binders h bl
+ | _ -> []
+
+let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with
+ | BFvar :: h, { v = CLambdaN ([[_, x], _, _], c) } ->
+ let bs, c' = format_constr_expr h c in
+ Bvar x :: bs, c'
+ | BFdecl _:: h, { v = CLambdaN ([lxs, _, t], c) } ->
+ let bs, c' = format_constr_expr h c in
+ Bdecl (List.map snd lxs, t) :: bs, c'
+ | BFdef :: h, { v = CLetIn((_, x), v, oty, c) } ->
+ let bs, c' = format_constr_expr h c in
+ Bdef (x, oty, v) :: bs, c'
+ | [BFcast], { v = CCast (c, CastConv t) } ->
+ [Bcast t], c
+ | BFrec (has_str, has_cast) :: h,
+ { v = CFix ( _, [_, (Some locn, CStructRec), bl, t, c]) } ->
+ let bs = format_local_binders h bl in
+ let bstr = if has_str then [Bstruct (Name (snd locn))] else [] in
+ bs @ bstr @ (if has_cast then [Bcast t] else []), c
+ | BFrec (_, has_cast) :: h, { v = CCoFix ( _, [_, bl, t, c]) } ->
+ format_local_binders h bl @ (if has_cast then [Bcast t] else []), c
+ | _, c ->
+ [], c
+
+let rec format_glob_decl h0 d0 = match h0, d0 with
+ | BFvar :: h, (x, _, None, _) :: d ->
+ Bvar x :: format_glob_decl h d
+ | BFdecl 1 :: h, (x, _, None, t) :: d ->
+ Bdecl ([x], t) :: format_glob_decl h d
+ | BFdecl n :: h, (x, _, None, t) :: d when n > 1 ->
+ begin match format_glob_decl (BFdecl (n - 1) :: h) d with
+ | Bdecl (xs, _) :: bs -> Bdecl (x :: xs, t) :: bs
+ | bs -> Bdecl ([x], t) :: bs
+ end
+ | BFdef :: h, (x, _, Some v, _) :: d ->
+ Bdef (x, None, v) :: format_glob_decl h d
+ | _, (x, _, None, t) :: d ->
+ Bdecl ([x], t) :: format_glob_decl [] d
+ | _, (x, _, Some v, _) :: d ->
+ Bdef (x, None, v) :: format_glob_decl [] d
+ | _, [] -> []
+
+let rec format_glob_constr h0 c0 = let open CAst in match h0, c0 with
+ | BFvar :: h, { v = GLambda (x, _, _, c) } ->
+ let bs, c' = format_glob_constr h c in
+ Bvar x :: bs, c'
+ | BFdecl 1 :: h, { v = GLambda (x, _, t, c) } ->
+ let bs, c' = format_glob_constr h c in
+ Bdecl ([x], t) :: bs, c'
+ | BFdecl n :: h, { v = GLambda (x, _, t, c) } when n > 1 ->
+ begin match format_glob_constr (BFdecl (n - 1) :: h) c with
+ | Bdecl (xs, _) :: bs, c' -> Bdecl (x :: xs, t) :: bs, c'
+ | _ -> [Bdecl ([x], t)], c
+ end
+ | BFdef :: h, { v = GLetIn(x, v, oty, c) } ->
+ let bs, c' = format_glob_constr h c in
+ Bdef (x, oty, v) :: bs, c'
+ | [BFcast], { v = GCast (c, CastConv t) } ->
+ [Bcast t], c
+ | BFrec (has_str, has_cast) :: h, { v = GRec (f, _, bl, t, c) }
+ when Array.length c = 1 ->
+ let bs = format_glob_decl h bl.(0) in
+ let bstr = match has_str, f with
+ | true, GFix ([|Some i, GStructRec|], _) -> mkBstruct i bs
+ | _ -> [] in
+ bs @ bstr @ (if has_cast then [Bcast t.(0)] else []), c.(0)
+ | _, c ->
+ [], c
+
+(** Forward chaining argument *)
+
+(* There are three kinds of forward definitions: *)
+(* - Hint: type only, cast to Type, may have proof hint. *)
+(* - Have: type option + value, no space before type *)
+(* - Pose: binders + value, space before binders. *)
+
+let pr_fwdkind = function
+ | FwdHint (s,_) -> str (s ^ " ") | _ -> str " :=" ++ spc ()
+let pr_fwdfmt (fk, _ : ssrfwdfmt) = pr_fwdkind fk
+
+let wit_ssrfwdfmt = add_genarg "ssrfwdfmt" pr_fwdfmt
+
+(* type ssrfwd = ssrfwdfmt * ssrterm *)
+
+let mkFwdVal fk c = ((fk, []), mk_term xNoFlag c)
+let mkssrFwdVal fk c = ((fk, []), (c,None))
+let dC t = CastConv t
+
+let mkFwdCast fk ?loc t c = ((fk, [BFcast]), mk_term ' ' (CAst.make ?loc @@ CCast (c, dC t)))
+let mkssrFwdCast fk loc t c = ((fk, [BFcast]), (c, Some t))
+
+let mkFwdHint s t =
+ let loc = Constrexpr_ops.constr_loc t in
+ mkFwdCast (FwdHint (s,false)) ?loc t (mkCHole loc)
+let mkFwdHintNoTC s t =
+ let loc = Constrexpr_ops.constr_loc t in
+ mkFwdCast (FwdHint (s,true)) ?loc t (mkCHole loc)
+
+let pr_gen_fwd prval prc prlc fk (bs, c) =
+ let prc s = str s ++ spc () ++ prval prc prlc c in
+ match fk, bs with
+ | FwdHint (s,_), [Bcast t] -> str s ++ spc () ++ prlc t
+ | FwdHint (s,_), _ -> prc (s ^ "(* typeof *)")
+ | FwdHave, [Bcast t] -> str ":" ++ spc () ++ prlc t ++ prc " :="
+ | _, [] -> prc " :="
+ | _, _ -> spc () ++ pr_list spc (pr_binder prlc) bs ++ prc " :="
+
+let pr_fwd_guarded prval prval' = function
+| (fk, h), (_, (_, Some c)) ->
+ pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c)
+| (fk, h), (_, (c, None)) ->
+ pr_gen_fwd prval' pr_glob_constr prl_glob_constr fk (format_glob_constr h c)
+
+let pr_unguarded prc prlc = prlc
+
+let pr_fwd = pr_fwd_guarded pr_unguarded pr_unguarded
+let pr_ssrfwd _ _ _ = pr_fwd
+
+ARGUMENT EXTEND ssrfwd TYPED AS (ssrfwdfmt * ssrterm) PRINTED BY pr_ssrfwd
+ | [ ":=" lconstr(c) ] -> [ mkFwdVal FwdPose c ]
+ | [ ":" lconstr (t) ":=" lconstr(c) ] -> [ mkFwdCast FwdPose ~loc t c ]
+END
+
+(** Independent parsing for binders *)
+
+(* The pose, pose fix, and pose cofix tactics use these internally to *)
+(* parse argument fragments. *)
+
+let pr_ssrbvar prc _ _ v = prc v
+
+ARGUMENT EXTEND ssrbvar TYPED AS constr PRINTED BY pr_ssrbvar
+| [ ident(id) ] -> [ mkCVar ~loc id ]
+| [ "_" ] -> [ mkCHole (Some loc) ]
+END
+
+let bvar_lname = let open CAst in function
+ | { v = CRef (Ident (loc, id), _) } -> Loc.tag ?loc @@ Name id
+ | { loc = loc } -> Loc.tag ?loc Anonymous
+
+let pr_ssrbinder prc _ _ (_, c) = prc c
+
+ARGUMENT EXTEND ssrbinder TYPED AS ssrfwdfmt * constr PRINTED BY pr_ssrbinder
+ | [ ssrbvar(bv) ] ->
+ [ let xloc, _ as x = bvar_lname bv in
+ (FwdPose, [BFvar]),
+ CAst.make ~loc @@ CLambdaN ([[x],Default Explicit,mkCHole xloc],mkCHole (Some loc)) ]
+ | [ "(" ssrbvar(bv) ")" ] ->
+ [ let xloc, _ as x = bvar_lname bv in
+ (FwdPose, [BFvar]),
+ CAst.make ~loc @@ CLambdaN ([[x],Default Explicit,mkCHole xloc],mkCHole (Some loc)) ]
+ | [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] ->
+ [ let x = bvar_lname bv in
+ (FwdPose, [BFdecl 1]),
+ CAst.make ~loc @@ CLambdaN ([[x], Default Explicit, t], mkCHole (Some loc)) ]
+ | [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] ->
+ [ let xs = List.map bvar_lname (bv :: bvs) in
+ let n = List.length xs in
+ (FwdPose, [BFdecl n]),
+ CAst.make ~loc @@ CLambdaN ([xs, Default Explicit, t], mkCHole (Some loc)) ]
+ | [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] ->
+ [ (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, Some t, mkCHole (Some loc)) ]
+ | [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] ->
+ [ (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, None, mkCHole (Some loc)) ]
+ END
+
+GEXTEND Gram
+ GLOBAL: ssrbinder;
+ ssrbinder: [
+ [ ["of" | "&"]; c = operconstr LEVEL "99" ->
+ let loc = !@loc in
+ (FwdPose, [BFvar]),
+ CAst.make ~loc @@ CLambdaN ([[Loc.tag ~loc Anonymous],Default Explicit,c],mkCHole (Some loc)) ]
+ ];
+END
+
+let rec binders_fmts = function
+ | ((_, h), _) :: bs -> h @ binders_fmts bs
+ | _ -> []
+
+let push_binders c2 bs =
+ let loc2 = constr_loc c2 in let mkloc loc1 = Loc.merge_opt loc1 loc2 in
+ let open CAst in
+ let rec loop ty c = function
+ | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs when ty ->
+ CAst.make ?loc:(mkloc loc1) @@ CProdN (b, loop ty c bs)
+ | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs ->
+ CAst.make ?loc:(mkloc loc1) @@ CLambdaN (b, loop ty c bs)
+ | (_, { loc = loc1; v = CLetIn (x, v, oty, _) } ) :: bs ->
+ CAst.make ?loc:(mkloc loc1) @@ CLetIn (x, v, oty, loop ty c bs)
+ | [] -> c
+ | _ -> anomaly "binder not a lambda nor a let in" in
+ match c2 with
+ | { loc; v = CCast (ct, CastConv cty) } ->
+ CAst.make ?loc @@ (CCast (loop false ct bs, CastConv (loop true cty bs)))
+ | ct -> loop false ct bs
+
+let rec fix_binders = let open CAst in function
+ | (_, { v = CLambdaN ([xs, _, t], _) } ) :: bs ->
+ CLocalAssum (xs, Default Explicit, t) :: fix_binders bs
+ | (_, { v = CLetIn (x, v, oty, _) } ) :: bs ->
+ CLocalDef (x, v, oty) :: fix_binders bs
+ | _ -> []
+
+let pr_ssrstruct _ _ _ = function
+ | Some id -> str "{struct " ++ pr_id id ++ str "}"
+ | None -> mt ()
+
+ARGUMENT EXTEND ssrstruct TYPED AS ident option PRINTED BY pr_ssrstruct
+| [ "{" "struct" ident(id) "}" ] -> [ Some id ]
+| [ ] -> [ None ]
+END
+
+(** The "pose" tactic *)
+
+(* The plain pose form. *)
+
+let bind_fwd bs = function
+ | (fk, h), (ck, (rc, Some c)) ->
+ (fk,binders_fmts bs @ h), (ck,(rc,Some (push_binders c bs)))
+ | fwd -> fwd
+
+ARGUMENT EXTEND ssrposefwd TYPED AS ssrfwd PRINTED BY pr_ssrfwd
+ | [ ssrbinder_list(bs) ssrfwd(fwd) ] -> [ bind_fwd bs fwd ]
+END
+
+(* The pose fix form. *)
+
+let pr_ssrfixfwd _ _ _ (id, fwd) = str " fix " ++ pr_id id ++ pr_fwd fwd
+
+let bvar_locid = function
+ | { CAst.v = CRef (Ident (loc, id), _) } -> loc, id
+ | _ -> CErrors.user_err (Pp.str "Missing identifier after \"(co)fix\"")
+
+
+ARGUMENT EXTEND ssrfixfwd TYPED AS ident * ssrfwd PRINTED BY pr_ssrfixfwd
+ | [ "fix" ssrbvar(bv) ssrbinder_list(bs) ssrstruct(sid) ssrfwd(fwd) ] ->
+ [ let (_, id) as lid = bvar_locid bv in
+ let (fk, h), (ck, (rc, oc)) = fwd in
+ let c = Option.get oc in
+ let has_cast, t', c' = match format_constr_expr h c with
+ | [Bcast t'], c' -> true, t', c'
+ | _ -> false, mkCHole (constr_loc c), c in
+ let lb = fix_binders bs in
+ let has_struct, i =
+ let rec loop = function
+ (l', Name id') :: _ when Option.equal Id.equal sid (Some id') -> true, (l', id')
+ | [l', Name id'] when sid = None -> false, (l', id')
+ | _ :: bn -> loop bn
+ | [] -> CErrors.user_err (Pp.str "Bad structural argument") in
+ loop (names_of_local_assums lb) in
+ let h' = BFrec (has_struct, has_cast) :: binders_fmts bs in
+ let fix = CAst.make ~loc @@ CFix (lid, [lid, (Some i, CStructRec), lb, t', c']) in
+ id, ((fk, h'), (ck, (rc, Some fix))) ]
+END
+
+
+(* The pose cofix form. *)
+
+let pr_ssrcofixfwd _ _ _ (id, fwd) = str " cofix " ++ pr_id id ++ pr_fwd fwd
+
+ARGUMENT EXTEND ssrcofixfwd TYPED AS ssrfixfwd PRINTED BY pr_ssrcofixfwd
+ | [ "cofix" ssrbvar(bv) ssrbinder_list(bs) ssrfwd(fwd) ] ->
+ [ let _, id as lid = bvar_locid bv in
+ let (fk, h), (ck, (rc, oc)) = fwd in
+ let c = Option.get oc in
+ let has_cast, t', c' = match format_constr_expr h c with
+ | [Bcast t'], c' -> true, t', c'
+ | _ -> false, mkCHole (constr_loc c), c in
+ let h' = BFrec (false, has_cast) :: binders_fmts bs in
+ let cofix = CAst.make ~loc @@ CCoFix (lid, [lid, fix_binders bs, t', c']) in
+ id, ((fk, h'), (ck, (rc, Some cofix)))
+ ]
+END
+
+(* This does not print the type, it should be fixed... *)
+let pr_ssrsetfwd _ _ _ (((fk,_),(t,_)), docc) =
+ pr_gen_fwd (fun _ _ -> pr_cpattern)
+ (fun _ -> mt()) (fun _ -> mt()) fk ([Bcast ()],t)
+
+ARGUMENT EXTEND ssrsetfwd
+TYPED AS (ssrfwdfmt * (lcpattern * ssrterm option)) * ssrdocc
+PRINTED BY pr_ssrsetfwd
+| [ ":" lconstr(t) ":=" "{" ssrocc(occ) "}" cpattern(c) ] ->
+ [ mkssrFwdCast FwdPose loc (mk_lterm t) c, mkocc occ ]
+| [ ":" lconstr(t) ":=" lcpattern(c) ] ->
+ [ mkssrFwdCast FwdPose loc (mk_lterm t) c, nodocc ]
+| [ ":=" "{" ssrocc(occ) "}" cpattern(c) ] ->
+ [ mkssrFwdVal FwdPose c, mkocc occ ]
+| [ ":=" lcpattern(c) ] -> [ mkssrFwdVal FwdPose c, nodocc ]
+END
+
+
+let pr_ssrhavefwd _ _ prt (fwd, hint) = pr_fwd fwd ++ pr_hint prt hint
+
+ARGUMENT EXTEND ssrhavefwd TYPED AS ssrfwd * ssrhint PRINTED BY pr_ssrhavefwd
+| [ ":" lconstr(t) ssrhint(hint) ] -> [ mkFwdHint ":" t, hint ]
+| [ ":" lconstr(t) ":=" lconstr(c) ] -> [ mkFwdCast FwdHave ~loc t c, nohint ]
+| [ ":" lconstr(t) ":=" ] -> [ mkFwdHintNoTC ":" t, nohint ]
+| [ ":=" lconstr(c) ] -> [ mkFwdVal FwdHave c, nohint ]
+END
+
+let intro_id_to_binder = List.map (function
+ | IPatId id ->
+ let xloc, _ as x = bvar_lname (mkCVar id) in
+ (FwdPose, [BFvar]),
+ CAst.make @@ CLambdaN ([[x], Default Explicit, mkCHole xloc],
+ mkCHole None)
+ | _ -> anomaly "non-id accepted as binder")
+
+let binder_to_intro_id = CAst.(List.map (function
+ | (FwdPose, [BFvar]), { v = CLambdaN ([ids,_,_],_) }
+ | (FwdPose, [BFdecl _]), { v = CLambdaN ([ids,_,_],_) } ->
+ List.map (function (_, Name id) -> IPatId id | _ -> IPatAnon One) ids
+ | (FwdPose, [BFdef]), { v = CLetIn ((_,Name id),_,_,_) } -> [IPatId id]
+ | (FwdPose, [BFdef]), { v = CLetIn ((_,Anonymous),_,_,_) } -> [IPatAnon One]
+ | _ -> anomaly "ssrbinder is not a binder"))
+
+let pr_ssrhavefwdwbinders _ _ prt (tr,((hpats, (fwd, hint)))) =
+ pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint
+
+ARGUMENT EXTEND ssrhavefwdwbinders
+ TYPED AS bool * (ssrhpats * (ssrfwd * ssrhint))
+ PRINTED BY pr_ssrhavefwdwbinders
+| [ ssrhpats_wtransp(trpats) ssrbinder_list(bs) ssrhavefwd(fwd) ] ->
+ [ let tr, pats = trpats in
+ let ((clr, pats), binders), simpl = pats in
+ let allbs = intro_id_to_binder binders @ bs in
+ let allbinders = binders @ List.flatten (binder_to_intro_id bs) in
+ let hint = bind_fwd allbs (fst fwd), snd fwd in
+ tr, ((((clr, pats), allbinders), simpl), hint) ]
+END
+
+
+let pr_ssrdoarg prc _ prt (((n, m), tac), clauses) =
+ pr_index n ++ pr_mmod m ++ pr_hintarg prt tac ++ pr_clauses clauses
+
+ARGUMENT EXTEND ssrdoarg
+ TYPED AS ((ssrindex * ssrmmod) * ssrhintarg) * ssrclauses
+ PRINTED BY pr_ssrdoarg
+| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+(* type ssrseqarg = ssrindex * (ssrtacarg * ssrtac option) *)
+
+let pr_seqtacarg prt = function
+ | (is_first, []), _ -> str (if is_first then "first" else "last")
+ | tac, Some dtac ->
+ hv 0 (pr_hintarg prt tac ++ spc() ++ str "|| " ++ prt tacltop dtac)
+ | tac, _ -> pr_hintarg prt tac
+
+let pr_ssrseqarg _ _ prt = function
+ | ArgArg 0, tac -> pr_seqtacarg prt tac
+ | i, tac -> pr_index i ++ str " " ++ pr_seqtacarg prt tac
+
+(* We must parse the index separately to resolve the conflict with *)
+(* an unindexed tactic. *)
+ARGUMENT EXTEND ssrseqarg TYPED AS ssrindex * (ssrhintarg * tactic option)
+ PRINTED BY pr_ssrseqarg
+| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+let sq_brace_tacnames =
+ ["first"; "solve"; "do"; "rewrite"; "have"; "suffices"; "wlog"]
+ (* "by" is a keyword *)
+let accept_ssrseqvar strm =
+ match stream_nth 0 strm with
+ | Tok.IDENT id when not (List.mem id sq_brace_tacnames) ->
+ accept_before_syms_or_ids ["["] ["first";"last"] strm
+ | _ -> raise Stream.Failure
+
+let test_ssrseqvar = Gram.Entry.of_parser "test_ssrseqvar" accept_ssrseqvar
+
+let swaptacarg (loc, b) = (b, []), Some (TacId [])
+
+let check_seqtacarg dir arg = match snd arg, dir with
+ | ((true, []), Some (TacAtom (loc, _))), L2R ->
+ CErrors.user_err ?loc (str "expected \"last\"")
+ | ((false, []), Some (TacAtom (loc, _))), R2L ->
+ CErrors.user_err ?loc (str "expected \"first\"")
+ | _, _ -> arg
+
+let ssrorelse = Gram.entry_create "ssrorelse"
+GEXTEND Gram
+ GLOBAL: ssrorelse ssrseqarg;
+ ssrseqidx: [
+ [ test_ssrseqvar; id = Prim.ident -> ArgVar (Loc.tag ~loc:!@loc id)
+ | n = Prim.natural -> ArgArg (check_index ~loc:!@loc n)
+ ] ];
+ ssrswap: [[ IDENT "first" -> !@loc, true | IDENT "last" -> !@loc, false ]];
+ ssrorelse: [[ "||"; tac = tactic_expr LEVEL "2" -> tac ]];
+ ssrseqarg: [
+ [ arg = ssrswap -> noindex, swaptacarg arg
+ | i = ssrseqidx; tac = ssrortacarg; def = OPT ssrorelse -> i, (tac, def)
+ | i = ssrseqidx; arg = ssrswap -> i, swaptacarg arg
+ | tac = tactic_expr LEVEL "3" -> noindex, (mk_hint tac, None)
+ ] ];
+END
+
+let tactic_expr = Pltac.tactic_expr
+
+(** 1. Utilities *)
+
+(** Tactic-level diagnosis *)
+
+(* debug *)
+
+(* Let's play with the new proof engine API *)
+let old_tac = Proofview.V82.tactic
+
+
+(** Name generation {{{ *******************************************************)
+
+(* Since Coq now does repeated internal checks of its external lexical *)
+(* rules, we now need to carve ssreflect reserved identifiers out of *)
+(* out of the user namespace. We use identifiers of the form _id_ for *)
+(* this purpose, e.g., we "anonymize" an identifier id as _id_, adding *)
+(* an extra leading _ if this might clash with an internal identifier. *)
+(* We check for ssreflect identifiers in the ident grammar rule; *)
+(* when the ssreflect Module is present this is normally an error, *)
+(* but we provide a compatibility flag to reduce this to a warning. *)
+
+let ssr_reserved_ids = Summary.ref ~name:"SSR:idents" true
+
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "ssreflect identifiers";
+ Goptions.optkey = ["SsrIdents"];
+ Goptions.optdepr = false;
+ Goptions.optread = (fun _ -> !ssr_reserved_ids);
+ Goptions.optwrite = (fun b -> ssr_reserved_ids := b)
+ }
+
+let is_ssr_reserved s =
+ let n = String.length s in n > 2 && s.[0] = '_' && s.[n - 1] = '_'
+
+let ssr_id_of_string loc s =
+ if is_ssr_reserved s && is_ssr_loaded () then begin
+ if !ssr_reserved_ids then
+ CErrors.user_err ~loc (str ("The identifier " ^ s ^ " is reserved."))
+ else if is_internal_name s then
+ Feedback.msg_warning (str ("Conflict between " ^ s ^ " and ssreflect internal names."))
+ 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
+
+let ssr_null_entry = Gram.Entry.of_parser "ssr_null" (fun _ -> ())
+
+let (!@) = Pcoq.to_coqloc
+
+GEXTEND Gram
+ GLOBAL: Prim.ident;
+ Prim.ident: [[ s = IDENT; ssr_null_entry -> ssr_id_of_string !@loc s ]];
+END
+
+let perm_tag = "_perm_Hyp_"
+let _ = add_internal_name (is_tagged perm_tag)
+
+(* }}} *)
+
+(* We must not anonymize context names discharged by the "in" tactical. *)
+
+(** Tactical extensions. {{{ **************************************************)
+
+(* The TACTIC EXTEND facility can't be used for defining new user *)
+(* tacticals, because: *)
+(* - the concrete syntax must start with a fixed string *)
+(* We use the following workaround: *)
+(* - We use the (unparsable) "YouShouldNotTypeThis" token for tacticals that *)
+(* don't start with a token, then redefine the grammar and *)
+(* printer using GEXTEND and set_pr_ssrtac, respectively. *)
+
+type ssrargfmt = ArgSsr of string | ArgSep of string
+
+let ssrtac_name name = {
+ mltac_plugin = "ssreflect_plugin";
+ mltac_tactic = "ssr" ^ name;
+}
+
+let ssrtac_entry name n = {
+ mltac_name = ssrtac_name name;
+ mltac_index = n;
+}
+
+let set_pr_ssrtac name prec afmt = (* FIXME *) () (*
+ let fmt = List.map (function
+ | ArgSep s -> Egramml.GramTerminal s
+ | ArgSsr s -> Egramml.GramTerminal s
+ | ArgCoq at -> Egramml.GramTerminal "COQ_ARG") afmt in
+ let tacname = ssrtac_name name in () *)
+
+let ssrtac_atom ?loc name args = TacML (Loc.tag ?loc (ssrtac_entry name 0, args))
+let ssrtac_expr ?loc name args = ssrtac_atom ?loc name args
+
+let tclintros_expr ?loc tac ipats =
+ let args = [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrintrosarg) (tac, ipats))] in
+ ssrtac_expr ?loc "tclintros" args
+
+GEXTEND Gram
+ GLOBAL: tactic_expr;
+ tactic_expr: LEVEL "1" [ RIGHTA
+ [ tac = tactic_expr; intros = ssrintros_ne -> tclintros_expr ~loc:!@loc tac intros
+ ] ];
+END
+
+(* }}} *)
+
+
+(** Bracketing tactical *)
+
+(* The tactic pretty-printer doesn't know that some extended tactics *)
+(* are actually tacticals. To prevent it from improperly removing *)
+(* parentheses we override the parsing rule for bracketed tactic *)
+(* expressions so that the pretty-print always reflects the input. *)
+(* (Removing user-specified parentheses is dubious anyway). *)
+
+GEXTEND Gram
+ GLOBAL: tactic_expr;
+ ssrparentacarg: [[ "("; tac = tactic_expr; ")" -> Loc.tag ~loc:!@loc (Tacexp tac) ]];
+ tactic_expr: LEVEL "0" [[ arg = ssrparentacarg -> TacArg arg ]];
+END
+
+(** The internal "done" and "ssrautoprop" tactics. *)
+
+(* For additional flexibility, "done" and "ssrautoprop" are *)
+(* defined in Ltac. *)
+(* Although we provide a default definition in ssreflect, *)
+(* we look up the definition dynamically at each call point, *)
+(* to allow for user extensions. "ssrautoprop" defaults to *)
+(* trivial. *)
+
+let ssrautoprop gl =
+ try
+ let tacname =
+ 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
+ with Not_found -> Proofview.V82.of_tactic (Auto.full_trivial []) gl
+
+let () = ssrautoprop_tac := ssrautoprop
+
+let tclBY tac = tclTHEN tac (donetac ~-1)
+
+(** Tactical arguments. *)
+
+(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *)
+(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *)
+(* and subgoal reordering tacticals (; first & ; last), respectively. *)
+
+(* Force use of the tactic_expr parsing entry, to rule out tick marks. *)
+
+(** The "by" tactical. *)
+
+
+open Ssrfwd
+
+TACTIC EXTEND ssrtclby
+| [ "by" ssrhintarg(tac) ] -> [ Proofview.V82.tactic (hinttac ist true tac) ]
+END
+
+(* }}} *)
+(* We can't parse "by" in ARGUMENT EXTEND because it will only be made *)
+(* into a keyword in ssreflect.v; so we anticipate this in GEXTEND. *)
+
+GEXTEND Gram
+ GLOBAL: ssrhint simple_tactic;
+ ssrhint: [[ "by"; arg = ssrhintarg -> arg ]];
+END
+
+open Ssripats
+
+(** The "do" tactical. ********************************************************)
+
+(*
+type ssrdoarg = ((ssrindex * ssrmmod) * ssrhint) * ssrclauses
+*)
+TACTIC EXTEND ssrtcldo
+| [ "YouShouldNotTypeThis" "do" ssrdoarg(arg) ] -> [ Proofview.V82.tactic (ssrdotac ist arg) ]
+END
+set_pr_ssrtac "tcldo" 3 [ArgSep "do "; ArgSsr "doarg"]
+
+let ssrdotac_expr ?loc n m tac clauses =
+ let arg = ((n, m), tac), clauses in
+ ssrtac_expr ?loc "tcldo" [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrdoarg) arg)]
+
+GEXTEND Gram
+ GLOBAL: tactic_expr;
+ ssrdotac: [
+ [ tac = tactic_expr LEVEL "3" -> mk_hint tac
+ | tacs = ssrortacarg -> tacs
+ ] ];
+ tactic_expr: LEVEL "3" [ RIGHTA
+ [ IDENT "do"; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses ->
+ ssrdotac_expr ~loc:!@loc noindex m tac clauses
+ | IDENT "do"; tac = ssrortacarg; clauses = ssrclauses ->
+ ssrdotac_expr ~loc:!@loc noindex Once tac clauses
+ | IDENT "do"; n = int_or_var; m = ssrmmod;
+ tac = ssrdotac; clauses = ssrclauses ->
+ ssrdotac_expr ~loc:!@loc (mk_index ~loc:!@loc n) m tac clauses
+ ] ];
+END
+(* }}} *)
+
+
+(* We can't actually parse the direction separately because this *)
+(* would introduce conflicts with the basic ltac syntax. *)
+let pr_ssrseqdir _ _ _ = function
+ | L2R -> str ";" ++ spc () ++ str "first "
+ | R2L -> str ";" ++ spc () ++ str "last "
+
+ARGUMENT EXTEND ssrseqdir TYPED AS ssrdir PRINTED BY pr_ssrseqdir
+| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+TACTIC EXTEND ssrtclseq
+| [ "YouShouldNotTypeThis" ssrtclarg(tac) ssrseqdir(dir) ssrseqarg(arg) ] ->
+ [ Proofview.V82.tactic (tclSEQAT ist tac dir arg) ]
+END
+set_pr_ssrtac "tclseq" 5 [ArgSsr "tclarg"; ArgSsr "seqdir"; ArgSsr "seqarg"]
+
+let tclseq_expr ?loc tac dir arg =
+ let arg1 = in_gen (rawwit wit_ssrtclarg) tac in
+ let arg2 = in_gen (rawwit wit_ssrseqdir) dir in
+ let arg3 = in_gen (rawwit wit_ssrseqarg) (check_seqtacarg dir arg) in
+ ssrtac_expr ?loc "tclseq" (List.map (fun x -> Tacexpr.TacGeneric x) [arg1; arg2; arg3])
+
+GEXTEND Gram
+ GLOBAL: tactic_expr;
+ ssr_first: [
+ [ tac = ssr_first; ipats = ssrintros_ne -> tclintros_expr ~loc:!@loc tac ipats
+ | "["; tacl = LIST0 tactic_expr SEP "|"; "]" -> TacFirst tacl
+ ] ];
+ ssr_first_else: [
+ [ tac1 = ssr_first; tac2 = ssrorelse -> TacOrelse (tac1, tac2)
+ | tac = ssr_first -> tac ]];
+ tactic_expr: LEVEL "4" [ LEFTA
+ [ tac1 = tactic_expr; ";"; IDENT "first"; tac2 = ssr_first_else ->
+ TacThen (tac1, tac2)
+ | tac = tactic_expr; ";"; IDENT "first"; arg = ssrseqarg ->
+ tclseq_expr ~loc:!@loc tac L2R arg
+ | tac = tactic_expr; ";"; IDENT "last"; arg = ssrseqarg ->
+ tclseq_expr ~loc:!@loc tac R2L arg
+ ] ];
+END
+(* }}} *)
+
+(** 5. Bookkeeping tactics (clear, move, case, elim) *)
+
+(** Generalization (discharge) item *)
+
+(* An item is a switch + term pair. *)
+
+(* type ssrgen = ssrdocc * ssrterm *)
+
+let pr_gen (docc, dt) = pr_docc docc ++ pr_cpattern dt
+
+let pr_ssrgen _ _ _ = pr_gen
+
+ARGUMENT EXTEND ssrgen TYPED AS ssrdocc * cpattern PRINTED BY pr_ssrgen
+| [ ssrdocc(docc) cpattern(dt) ] -> [ docc, dt ]
+| [ cpattern(dt) ] -> [ nodocc, dt ]
+END
+
+let has_occ ((_, occ), _) = occ <> None
+
+(** Generalization (discharge) sequence *)
+
+(* A discharge sequence is represented as a list of up to two *)
+(* lists of d-items, plus an ident list set (the possibly empty *)
+(* final clear switch). The main list is empty iff the command *)
+(* is defective, and has length two if there is a sequence of *)
+(* dependent terms (and in that case it is the first of the two *)
+(* lists). Thus, the first of the two lists is never empty. *)
+
+(* type ssrgens = ssrgen list *)
+(* type ssrdgens = ssrgens list * ssrclear *)
+
+let gens_sep = function [], [] -> mt | _ -> spc
+
+let pr_dgens pr_gen (gensl, clr) =
+ let prgens s gens = str s ++ pr_list spc pr_gen gens in
+ let prdeps deps = prgens ": " deps ++ spc () ++ str "/" in
+ match gensl with
+ | [deps; []] -> prdeps deps ++ pr_clear pr_spc clr
+ | [deps; gens] -> prdeps deps ++ prgens " " gens ++ pr_clear spc clr
+ | [gens] -> prgens ": " gens ++ pr_clear spc clr
+ | _ -> pr_clear pr_spc clr
+
+let pr_ssrdgens _ _ _ = pr_dgens pr_gen
+
+let cons_gen gen = function
+ | gens :: gensl, clr -> (gen :: gens) :: gensl, clr
+ | _ -> anomaly "missing gen list"
+
+let cons_dep (gensl, clr) =
+ if List.length gensl = 1 then ([] :: gensl, clr) else
+ CErrors.user_err (Pp.str "multiple dependents switches '/'")
+
+ARGUMENT EXTEND ssrdgens_tl TYPED AS ssrgen list list * ssrclear
+ PRINTED BY pr_ssrdgens
+| [ "{" ne_ssrhyp_list(clr) "}" cpattern(dt) ssrdgens_tl(dgens) ] ->
+ [ cons_gen (mkclr clr, dt) dgens ]
+| [ "{" ne_ssrhyp_list(clr) "}" ] ->
+ [ [[]], clr ]
+| [ "{" ssrocc(occ) "}" cpattern(dt) ssrdgens_tl(dgens) ] ->
+ [ cons_gen (mkocc occ, dt) dgens ]
+| [ "/" ssrdgens_tl(dgens) ] ->
+ [ cons_dep dgens ]
+| [ cpattern(dt) ssrdgens_tl(dgens) ] ->
+ [ cons_gen (nodocc, dt) dgens ]
+| [ ] ->
+ [ [[]], [] ]
+END
+
+ARGUMENT EXTEND ssrdgens TYPED AS ssrdgens_tl PRINTED BY pr_ssrdgens
+| [ ":" ssrgen(gen) ssrdgens_tl(dgens) ] -> [ cons_gen gen dgens ]
+END
+
+(** Equations *)
+
+(* argument *)
+
+let pr_eqid = function Some pat -> str " " ++ pr_ipat pat | None -> mt ()
+let pr_ssreqid _ _ _ = pr_eqid
+
+(* We must use primitive parsing here to avoid conflicts with the *)
+(* basic move, case, and elim tactics. *)
+ARGUMENT EXTEND ssreqid TYPED AS ssripatrep option PRINTED BY pr_ssreqid
+| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+let accept_ssreqid strm =
+ match Util.stream_nth 0 strm with
+ | Tok.IDENT _ -> accept_before_syms [":"] strm
+ | Tok.KEYWORD ":" -> ()
+ | Tok.KEYWORD pat when List.mem pat ["_"; "?"; "->"; "<-"] ->
+ accept_before_syms [":"] strm
+ | _ -> raise Stream.Failure
+
+let test_ssreqid = Gram.Entry.of_parser "test_ssreqid" accept_ssreqid
+
+GEXTEND Gram
+ GLOBAL: ssreqid;
+ ssreqpat: [
+ [ id = Prim.ident -> IPatId id
+ | "_" -> IPatAnon Drop
+ | "?" -> IPatAnon One
+ | occ = ssrdocc; "->" -> (match occ with
+ | None, occ -> IPatRewrite (occ, L2R)
+ | _ -> CErrors.user_err ~loc:!@loc (str"Only occurrences are allowed here"))
+ | occ = ssrdocc; "<-" -> (match occ with
+ | None, occ -> IPatRewrite (occ, R2L)
+ | _ -> CErrors.user_err ~loc:!@loc (str "Only occurrences are allowed here"))
+ | "->" -> IPatRewrite (allocc, L2R)
+ | "<-" -> IPatRewrite (allocc, R2L)
+ ]];
+ ssreqid: [
+ [ test_ssreqid; pat = ssreqpat -> Some pat
+ | test_ssreqid -> None
+ ]];
+END
+
+(** Bookkeeping (discharge-intro) argument *)
+
+(* Since all bookkeeping ssr commands have the same discharge-intro *)
+(* argument format we use a single grammar entry point to parse them. *)
+(* the entry point parses only non-empty arguments to avoid conflicts *)
+(* with the basic Coq tactics. *)
+
+(* type ssrarg = ssrview * (ssreqid * (ssrdgens * ssripats)) *)
+
+let pr_ssrarg _ _ _ (view, (eqid, (dgens, ipats))) =
+ let pri = pr_intros (gens_sep dgens) in
+ pr_view view ++ pr_eqid eqid ++ pr_dgens pr_gen dgens ++ pri ipats
+
+ARGUMENT EXTEND ssrarg TYPED AS ssrview * (ssreqid * (ssrdgens * ssrintros))
+ PRINTED BY pr_ssrarg
+| [ ssrview(view) ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] ->
+ [ view, (eqid, (dgens, ipats)) ]
+| [ ssrview(view) ssrclear(clr) ssrintros(ipats) ] ->
+ [ view, (None, (([], clr), ipats)) ]
+| [ ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] ->
+ [ [], (eqid, (dgens, ipats)) ]
+| [ ssrclear_ne(clr) ssrintros(ipats) ] ->
+ [ [], (None, (([], clr), ipats)) ]
+| [ ssrintros_ne(ipats) ] ->
+ [ [], (None, (([], []), ipats)) ]
+END
+
+(** The "clear" tactic *)
+
+(* We just add a numeric version that clears the n top assumptions. *)
+
+let poptac ist n = introstac ~ist (List.init n (fun _ -> IPatAnon Drop))
+
+TACTIC EXTEND ssrclear
+ | [ "clear" natural(n) ] -> [ Proofview.V82.tactic (poptac ist n) ]
+END
+
+(** The "move" tactic *)
+
+(* TODO: review this, in particular the => _ and => [] cases *)
+let rec improper_intros = function
+ | IPatSimpl _ :: ipats -> improper_intros ipats
+ | (IPatId _ | IPatAnon _ | IPatCase _) :: _ -> false
+ | _ -> true (* FIXME *)
+
+let check_movearg = function
+ | view, (eqid, _) when view <> [] && eqid <> None ->
+ CErrors.user_err (Pp.str "incompatible view and equation in move tactic")
+ | view, (_, (([gen :: _], _), _)) when view <> [] && has_occ gen ->
+ CErrors.user_err (Pp.str "incompatible view and occurrence switch in move tactic")
+ | _, (_, ((dgens, _), _)) when List.length dgens > 1 ->
+ CErrors.user_err (Pp.str "dependents switch `/' in move tactic")
+ | _, (eqid, (_, ipats)) when eqid <> None && improper_intros ipats ->
+ CErrors.user_err (Pp.str "no proper intro pattern for equation in move tactic")
+ | arg -> arg
+
+ARGUMENT EXTEND ssrmovearg TYPED AS ssrarg PRINTED BY pr_ssrarg
+| [ ssrarg(arg) ] -> [ check_movearg arg ]
+END
+
+
+
+TACTIC EXTEND ssrmove
+| [ "move" ssrmovearg(arg) ssrrpat(pat) ] ->
+ [ Proofview.V82.tactic (tclTHEN (ssrmovetac ist arg) (introstac ~ist [pat])) ]
+| [ "move" ssrmovearg(arg) ssrclauses(clauses) ] ->
+ [ Proofview.V82.tactic (tclCLAUSES ist (ssrmovetac ist arg) clauses) ]
+| [ "move" ssrrpat(pat) ] -> [ Proofview.V82.tactic (introstac ~ist [pat]) ]
+| [ "move" ] -> [ Proofview.V82.tactic (movehnftac) ]
+END
+
+let check_casearg = function
+| view, (_, (([_; gen :: _], _), _)) when view <> [] && has_occ gen ->
+ CErrors.user_err (Pp.str "incompatible view and occurrence switch in dependent case tactic")
+| arg -> arg
+
+ARGUMENT EXTEND ssrcasearg TYPED AS ssrarg PRINTED BY pr_ssrarg
+| [ ssrarg(arg) ] -> [ check_casearg arg ]
+END
+
+
+TACTIC EXTEND ssrcase
+| [ "case" ssrcasearg(arg) ssrclauses(clauses) ] ->
+ [ old_tac (tclCLAUSES ist (ssrcasetac ist arg) clauses) ]
+| [ "case" ] -> [ old_tac (with_fresh_ctx (with_top (ssrscasetac false))) ]
+END
+
+(** The "elim" tactic *)
+
+(* Elim views are elimination lemmas, so the eliminated term is not addded *)
+(* to the dependent terms as for "case", unless it actually occurs in the *)
+(* goal, the "all occurrences" {+} switch is used, or the equation switch *)
+(* is used and there are no dependents. *)
+
+let ssrelimtac ist (view, (eqid, (dgens, ipats))) =
+ let ndefectelimtac view eqid ipats deps gen ist gl =
+ let elim = match view with [v] -> Some (snd(force_term ist gl v)) | _ -> None in
+ ssrelim ~ist deps (`EGen gen) ?elim eqid (elim_intro_tac ipats) gl
+ in
+ with_dgens dgens (ndefectelimtac view eqid ipats) ist
+
+TACTIC EXTEND ssrelim
+| [ "elim" ssrarg(arg) ssrclauses(clauses) ] ->
+ [ old_tac (tclCLAUSES ist (ssrelimtac ist arg) clauses) ]
+| [ "elim" ] -> [ old_tac (with_fresh_ctx (with_top elimtac)) ]
+END
+
+(** 6. Backward chaining tactics: apply, exact, congr. *)
+
+(** The "apply" tactic *)
+
+let pr_agen (docc, dt) = pr_docc docc ++ pr_term dt
+let pr_ssragen _ _ _ = pr_agen
+let pr_ssragens _ _ _ = pr_dgens pr_agen
+
+ARGUMENT EXTEND ssragen TYPED AS ssrdocc * ssrterm PRINTED BY pr_ssragen
+| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ] -> [ mkclr clr, dt ]
+| [ ssrterm(dt) ] -> [ nodocc, dt ]
+END
+
+ARGUMENT EXTEND ssragens TYPED AS ssragen list list * ssrclear
+PRINTED BY pr_ssragens
+| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ssragens(agens) ] ->
+ [ cons_gen (mkclr clr, dt) agens ]
+| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ [[]], clr]
+| [ ssrterm(dt) ssragens(agens) ] ->
+ [ cons_gen (nodocc, dt) agens ]
+| [ ] -> [ [[]], [] ]
+END
+
+let mk_applyarg views agens intros = views, (None, (agens, intros))
+
+let pr_ssraarg _ _ _ (view, (eqid, (dgens, ipats))) =
+ let pri = pr_intros (gens_sep dgens) in
+ pr_view view ++ pr_eqid eqid ++ pr_dgens pr_agen dgens ++ pri ipats
+
+ARGUMENT EXTEND ssrapplyarg
+TYPED AS ssrview * (ssreqid * (ssragens * ssrintros))
+PRINTED BY pr_ssraarg
+| [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] ->
+ [ mk_applyarg [] (cons_gen gen dgens) intros ]
+| [ ssrclear_ne(clr) ssrintros(intros) ] ->
+ [ mk_applyarg [] ([], clr) intros ]
+| [ ssrintros_ne(intros) ] ->
+ [ mk_applyarg [] ([], []) intros ]
+| [ ssrview(view) ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] ->
+ [ mk_applyarg view (cons_gen gen dgens) intros ]
+| [ ssrview(view) ssrclear(clr) ssrintros(intros) ] ->
+ [ mk_applyarg view ([], clr) intros ]
+ END
+
+TACTIC EXTEND ssrapply
+| [ "apply" ssrapplyarg(arg) ] -> [ Proofview.V82.tactic (ssrapplytac ist arg) ]
+| [ "apply" ] -> [ Proofview.V82.tactic apply_top_tac ]
+END
+
+(** The "exact" tactic *)
+
+let mk_exactarg views dgens = mk_applyarg views dgens []
+
+ARGUMENT EXTEND ssrexactarg TYPED AS ssrapplyarg PRINTED BY pr_ssraarg
+| [ ":" ssragen(gen) ssragens(dgens) ] ->
+ [ mk_exactarg [] (cons_gen gen dgens) ]
+| [ ssrview(view) ssrclear(clr) ] ->
+ [ mk_exactarg view ([], clr) ]
+| [ ssrclear_ne(clr) ] ->
+ [ mk_exactarg [] ([], clr) ]
+END
+
+let vmexacttac pf =
+ Proofview.Goal.nf_enter begin fun gl ->
+ exact_no_check (EConstr.mkCast (pf, VMcast, Tacmach.New.pf_concl gl))
+ end
+
+TACTIC EXTEND ssrexact
+| [ "exact" ssrexactarg(arg) ] -> [ Proofview.V82.tactic (tclBY (ssrapplytac ist arg)) ]
+| [ "exact" ] -> [ Proofview.V82.tactic (tclORELSE (donetac ~-1) (tclBY apply_top_tac)) ]
+| [ "exact" "<:" lconstr(pf) ] -> [ vmexacttac pf ]
+END
+
+(** The "congr" tactic *)
+
+(* type ssrcongrarg = open_constr * (int * constr) *)
+
+let pr_ssrcongrarg _ _ _ ((n, f), dgens) =
+ (if n <= 0 then mt () else str " " ++ int n) ++
+ str " " ++ pr_term f ++ pr_dgens pr_gen dgens
+
+ARGUMENT EXTEND ssrcongrarg TYPED AS (int * ssrterm) * ssrdgens
+ PRINTED BY pr_ssrcongrarg
+| [ natural(n) constr(c) ssrdgens(dgens) ] -> [ (n, mk_term xNoFlag c), dgens ]
+| [ natural(n) constr(c) ] -> [ (n, mk_term xNoFlag c),([[]],[]) ]
+| [ constr(c) ssrdgens(dgens) ] -> [ (0, mk_term xNoFlag c), dgens ]
+| [ constr(c) ] -> [ (0, mk_term xNoFlag c), ([[]],[]) ]
+END
+
+
+
+TACTIC EXTEND ssrcongr
+| [ "congr" ssrcongrarg(arg) ] ->
+[ let arg, dgens = arg in
+ Proofview.V82.tactic begin
+ match dgens with
+ | [gens], clr -> tclTHEN (genstac (gens,clr) ist) (newssrcongrtac arg ist)
+ | _ -> errorstrm (str"Dependent family abstractions not allowed in congr")
+ end]
+END
+
+(** 7. Rewriting tactics (rewrite, unlock) *)
+
+(** Coq rewrite compatibility flag *)
+
+(** Rewrite clear/occ switches *)
+
+let pr_rwocc = function
+ | None, None -> mt ()
+ | None, occ -> pr_occ occ
+ | Some clr, _ -> pr_clear_ne clr
+
+let pr_ssrrwocc _ _ _ = pr_rwocc
+
+ARGUMENT EXTEND ssrrwocc TYPED AS ssrdocc PRINTED BY pr_ssrrwocc
+| [ "{" ssrhyp_list(clr) "}" ] -> [ mkclr clr ]
+| [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ]
+| [ ] -> [ noclr ]
+END
+
+(** Rewrite rules *)
+
+let pr_rwkind = function
+ | RWred s -> pr_simpl s
+ | RWdef -> str "/"
+ | RWeq -> mt ()
+
+let wit_ssrrwkind = add_genarg "ssrrwkind" pr_rwkind
+
+let pr_rule = function
+ | RWred s, _ -> pr_simpl s
+ | RWdef, r-> str "/" ++ pr_term r
+ | RWeq, r -> pr_term r
+
+let pr_ssrrule _ _ _ = pr_rule
+
+let noruleterm loc = mk_term xNoFlag (mkCProp loc)
+
+ARGUMENT EXTEND ssrrule_ne TYPED AS ssrrwkind * ssrterm PRINTED BY pr_ssrrule
+ | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+GEXTEND Gram
+ GLOBAL: ssrrule_ne;
+ ssrrule_ne : [
+ [ test_not_ssrslashnum; x =
+ [ "/"; t = ssrterm -> RWdef, t
+ | t = ssrterm -> RWeq, t
+ | s = ssrsimpl_ne -> RWred s, noruleterm (Some !@loc)
+ ] -> x
+ | s = ssrsimpl_ne -> RWred s, noruleterm (Some !@loc)
+ ]];
+END
+
+ARGUMENT EXTEND ssrrule TYPED AS ssrrule_ne PRINTED BY pr_ssrrule
+ | [ ssrrule_ne(r) ] -> [ r ]
+ | [ ] -> [ RWred Nop, noruleterm (Some loc) ]
+END
+
+(** Rewrite arguments *)
+
+let pr_option f = function None -> mt() | Some x -> f x
+let pr_pattern_squarep= pr_option (fun r -> str "[" ++ pr_rpattern r ++ str "]")
+let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep
+let pr_rwarg ((d, m), ((docc, rx), r)) =
+ pr_rwdir d ++ pr_mult m ++ pr_rwocc docc ++ pr_pattern_squarep rx ++ pr_rule r
+
+let pr_ssrrwarg _ _ _ = pr_rwarg
+
+ARGUMENT EXTEND ssrpattern_squarep
+TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep
+ | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ]
+ | [ ] -> [ None ]
+END
+
+ARGUMENT EXTEND ssrpattern_ne_squarep
+TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep
+ | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ]
+END
+
+
+ARGUMENT EXTEND ssrrwarg
+ TYPED AS (ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule)
+ PRINTED BY pr_ssrrwarg
+ | [ "-" ssrmult(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] ->
+ [ mk_rwarg (R2L, m) (docc, rx) r ]
+ | [ "-/" ssrterm(t) ] -> (* just in case '-/' should become a token *)
+ [ mk_rwarg (R2L, nomult) norwocc (RWdef, t) ]
+ | [ ssrmult_ne(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] ->
+ [ mk_rwarg (L2R, m) (docc, rx) r ]
+ | [ "{" ne_ssrhyp_list(clr) "}" ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] ->
+ [ mk_rwarg norwmult (mkclr clr, rx) r ]
+ | [ "{" ne_ssrhyp_list(clr) "}" ssrrule(r) ] ->
+ [ mk_rwarg norwmult (mkclr clr, None) r ]
+ | [ "{" ssrocc(occ) "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] ->
+ [ mk_rwarg norwmult (mkocc occ, rx) r ]
+ | [ "{" "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] ->
+ [ mk_rwarg norwmult (nodocc, rx) r ]
+ | [ ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] ->
+ [ mk_rwarg norwmult (noclr, rx) r ]
+ | [ ssrrule_ne(r) ] ->
+ [ mk_rwarg norwmult norwocc r ]
+END
+
+TACTIC EXTEND ssrinstofruleL2R
+| [ "ssrinstancesofruleL2R" ssrterm(arg) ] -> [ Proofview.V82.tactic (ssrinstancesofrule ist L2R arg) ]
+END
+TACTIC EXTEND ssrinstofruleR2L
+| [ "ssrinstancesofruleR2L" ssrterm(arg) ] -> [ Proofview.V82.tactic (ssrinstancesofrule ist R2L arg) ]
+END
+
+(** Rewrite argument sequence *)
+
+(* type ssrrwargs = ssrrwarg list *)
+
+let pr_ssrrwargs _ _ _ rwargs = pr_list spc pr_rwarg rwargs
+
+ARGUMENT EXTEND ssrrwargs TYPED AS ssrrwarg list PRINTED BY pr_ssrrwargs
+ | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ]
+END
+
+let ssr_rw_syntax = Summary.ref ~name:"SSR:rewrite" true
+
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "ssreflect rewrite";
+ Goptions.optkey = ["SsrRewrite"];
+ Goptions.optread = (fun _ -> !ssr_rw_syntax);
+ Goptions.optdepr = false;
+ Goptions.optwrite = (fun b -> ssr_rw_syntax := b) }
+
+let test_ssr_rw_syntax =
+ let test strm =
+ if not !ssr_rw_syntax then raise Stream.Failure else
+ if is_ssr_loaded () then () else
+ match Util.stream_nth 0 strm with
+ | Tok.KEYWORD key when List.mem key.[0] ['{'; '['; '/'] -> ()
+ | _ -> raise Stream.Failure in
+ Gram.Entry.of_parser "test_ssr_rw_syntax" test
+
+GEXTEND Gram
+ GLOBAL: ssrrwargs;
+ ssrrwargs: [[ test_ssr_rw_syntax; a = LIST1 ssrrwarg -> a ]];
+END
+
+(** The "rewrite" tactic *)
+
+TACTIC EXTEND ssrrewrite
+ | [ "rewrite" ssrrwargs(args) ssrclauses(clauses) ] ->
+ [ Proofview.V82.tactic (tclCLAUSES ist (ssrrewritetac ist args) clauses) ]
+END
+
+(** The "unlock" tactic *)
+
+let pr_unlockarg (occ, t) = pr_occ occ ++ pr_term t
+let pr_ssrunlockarg _ _ _ = pr_unlockarg
+
+ARGUMENT EXTEND ssrunlockarg TYPED AS ssrocc * ssrterm
+ PRINTED BY pr_ssrunlockarg
+ | [ "{" ssrocc(occ) "}" ssrterm(t) ] -> [ occ, t ]
+ | [ ssrterm(t) ] -> [ None, t ]
+END
+
+let pr_ssrunlockargs _ _ _ args = pr_list spc pr_unlockarg args
+
+ARGUMENT EXTEND ssrunlockargs TYPED AS ssrunlockarg list
+ PRINTED BY pr_ssrunlockargs
+ | [ ssrunlockarg_list(args) ] -> [ args ]
+END
+
+TACTIC EXTEND ssrunlock
+ | [ "unlock" ssrunlockargs(args) ssrclauses(clauses) ] ->
+[ Proofview.V82.tactic (tclCLAUSES ist (unlocktac ist args) clauses) ]
+END
+
+(** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *)
+
+
+TACTIC EXTEND ssrpose
+| [ "pose" ssrfixfwd(ffwd) ] -> [ Proofview.V82.tactic (ssrposetac ist ffwd) ]
+| [ "pose" ssrcofixfwd(ffwd) ] -> [ Proofview.V82.tactic (ssrposetac ist ffwd) ]
+| [ "pose" ssrfwdid(id) ssrposefwd(fwd) ] -> [ Proofview.V82.tactic (ssrposetac ist (id, fwd)) ]
+END
+
+(** The "set" tactic *)
+
+(* type ssrsetfwd = ssrfwd * ssrdocc *)
+
+TACTIC EXTEND ssrset
+| [ "set" ssrfwdid(id) ssrsetfwd(fwd) ssrclauses(clauses) ] ->
+ [ Proofview.V82.tactic (tclCLAUSES ist (ssrsettac ist id fwd) clauses) ]
+END
+
+(** The "have" tactic *)
+
+(* type ssrhavefwd = ssrfwd * ssrhint *)
+
+
+(* Pltac. *)
+
+(* The standard TACTIC EXTEND does not work for abstract *)
+GEXTEND Gram
+ GLOBAL: tactic_expr;
+ tactic_expr: LEVEL "3"
+ [ RIGHTA [ IDENT "abstract"; gens = ssrdgens ->
+ ssrtac_expr ~loc:!@loc "abstract"
+ [Tacexpr.TacGeneric (Genarg.in_gen (Genarg.rawwit wit_ssrdgens) gens)] ]];
+END
+TACTIC EXTEND ssrabstract
+| [ "abstract" ssrdgens(gens) ] -> [
+ if List.length (fst gens) <> 1 then
+ errorstrm (str"dependents switches '/' not allowed here");
+ Proofview.V82.tactic (ssrabstract ist gens) ]
+END
+
+TACTIC EXTEND ssrhave
+| [ "have" ssrhavefwdwbinders(fwd) ] ->
+ [ Proofview.V82.tactic (havetac ist fwd false false) ]
+END
+
+TACTIC EXTEND ssrhavesuff
+| [ "have" "suff" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] ->
+ [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true false) ]
+END
+
+TACTIC EXTEND ssrhavesuffices
+| [ "have" "suffices" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] ->
+ [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true false) ]
+END
+
+TACTIC EXTEND ssrsuffhave
+| [ "suff" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] ->
+ [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true true) ]
+END
+
+TACTIC EXTEND ssrsufficeshave
+| [ "suffices" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] ->
+ [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true true) ]
+END
+
+(** The "suffice" tactic *)
+
+let pr_ssrsufffwdwbinders _ _ prt (hpats, (fwd, hint)) =
+ pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint
+
+ARGUMENT EXTEND ssrsufffwd
+ TYPED AS ssrhpats * (ssrfwd * ssrhint) PRINTED BY pr_ssrsufffwdwbinders
+| [ ssrhpats(pats) ssrbinder_list(bs) ":" lconstr(t) ssrhint(hint) ] ->
+ [ let ((clr, pats), binders), simpl = pats in
+ let allbs = intro_id_to_binder binders @ bs in
+ let allbinders = binders @ List.flatten (binder_to_intro_id bs) in
+ let fwd = mkFwdHint ":" t in
+ (((clr, pats), allbinders), simpl), (bind_fwd allbs fwd, hint) ]
+END
+
+
+TACTIC EXTEND ssrsuff
+| [ "suff" ssrsufffwd(fwd) ] -> [ Proofview.V82.tactic (sufftac ist fwd) ]
+END
+
+TACTIC EXTEND ssrsuffices
+| [ "suffices" ssrsufffwd(fwd) ] -> [ Proofview.V82.tactic (sufftac ist fwd) ]
+END
+
+(** The "wlog" (Without Loss Of Generality) tactic *)
+
+(* type ssrwlogfwd = ssrwgen list * ssrfwd *)
+
+let pr_ssrwlogfwd _ _ _ (gens, t) =
+ str ":" ++ pr_list mt pr_wgen gens ++ spc() ++ pr_fwd t
+
+ARGUMENT EXTEND ssrwlogfwd TYPED AS ssrwgen list * ssrfwd
+ PRINTED BY pr_ssrwlogfwd
+| [ ":" ssrwgen_list(gens) "/" lconstr(t) ] -> [ gens, mkFwdHint "/" t]
+END
+
+
+TACTIC EXTEND ssrwlog
+| [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] ->
+ [ Proofview.V82.tactic (wlogtac ist pats fwd hint false `NoGen) ]
+END
+
+TACTIC EXTEND ssrwlogs
+| [ "wlog" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] ->
+ [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ]
+END
+
+TACTIC EXTEND ssrwlogss
+| [ "wlog" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]->
+ [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ]
+END
+
+TACTIC EXTEND ssrwithoutloss
+| [ "without" "loss" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] ->
+ [ Proofview.V82.tactic (wlogtac ist pats fwd hint false `NoGen) ]
+END
+
+TACTIC EXTEND ssrwithoutlosss
+| [ "without" "loss" "suff"
+ ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] ->
+ [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ]
+END
+
+TACTIC EXTEND ssrwithoutlossss
+| [ "without" "loss" "suffices"
+ ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]->
+ [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ]
+END
+
+(* Generally have *)
+let pr_idcomma _ _ _ = function
+ | None -> mt()
+ | Some None -> str"_, "
+ | Some (Some id) -> pr_id id ++ str", "
+
+ARGUMENT EXTEND ssr_idcomma TYPED AS ident option option PRINTED BY pr_idcomma
+ | [ ] -> [ None ]
+END
+
+let accept_idcomma strm =
+ match stream_nth 0 strm with
+ | Tok.IDENT _ | Tok.KEYWORD "_" -> accept_before_syms [","] strm
+ | _ -> raise Stream.Failure
+
+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 ]; "," ->
+ Some ip
+ ] ];
+END
+
+let augment_preclr clr1 (((clr0, x),y),z) = (((clr1 @ clr0, x),y),z)
+
+TACTIC EXTEND ssrgenhave
+| [ "gen" "have" ssrclear(clr)
+ ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] ->
+ [ let pats = augment_preclr clr pats in
+ Proofview.V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) ]
+END
+
+TACTIC EXTEND ssrgenhave2
+| [ "generally" "have" ssrclear(clr)
+ ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] ->
+ [ let pats = augment_preclr clr pats in
+ Proofview.V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) ]
+END
+
+(* We wipe out all the keywords generated by the grammar rules we defined. *)
+(* The user is supposed to Require Import ssreflect or Require ssreflect *)
+(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *)
+(* consequence the extended ssreflect grammar. *)
+let () = CLexer.set_keyword_state frozen_lexer ;;
+
+
+(* vim: set filetype=ocaml foldmethod=marker: *)
diff --git a/plugins/ssr/ssrparser.mli b/plugins/ssr/ssrparser.mli
new file mode 100644
index 000000000..bf6f44f11
--- /dev/null
+++ b/plugins/ssr/ssrparser.mli
@@ -0,0 +1,20 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+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
+
+val ssrtclarg : Tacexpr.raw_tactic_expr Pcoq.Gram.entry
+val wit_ssrtclarg : (Tacexpr.raw_tactic_expr, Tacexpr.glob_tactic_expr, Geninterp.Val.t) Genarg.genarg_type
+val pr_ssrtclarg : 'a -> 'b -> (int * Ppextend.parenRelation -> 'c -> 'd) -> 'c -> 'd
+
+val add_genarg : string -> ('a -> Pp.std_ppcmds) -> 'a Genarg.uniform_genarg_type
+
diff --git a/plugins/ssr/ssrprinters.ml b/plugins/ssr/ssrprinters.ml
new file mode 100644
index 000000000..e865ef706
--- /dev/null
+++ b/plugins/ssr/ssrprinters.ml
@@ -0,0 +1,85 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Pp
+open Names
+open Printer
+open Tacmach
+
+open Ssrmatching_plugin
+open Ssrast
+
+let pr_spc () = str " "
+let pr_bar () = Pp.cut() ++ str "|"
+let pr_list = prlist_with_sep
+
+let pp_concat hd ?(sep=str", ") = function [] -> hd | x :: xs ->
+ hd ++ List.fold_left (fun acc x -> acc ++ sep ++ x) x xs
+
+let pp_term gl t =
+ let t = Reductionops.nf_evar (project gl) t in pr_econstr t
+
+(* FIXME *)
+(* terms are pre constr, the kind is parsing/printing flag to distinguish
+ * between x, @x and (x). It affects automatic clear and let-in preservation.
+ * Cpattern is a temporary flag that becomes InParens ASAP. *)
+(* type ssrtermkind = InParens | WithAt | NoFlag | Cpattern *)
+let xInParens = '('
+let xWithAt = '@'
+let xNoFlag = ' '
+let xCpattern = 'x'
+
+(* Term printing utilities functions for deciding bracketing. *)
+let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")")
+(* String lexing utilities *)
+let skip_wschars s =
+ let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop
+(* We also guard characters that might interfere with the ssreflect *)
+(* tactic syntax. *)
+let guard_term ch1 s i = match s.[i] with
+ | '(' -> false
+ | '{' | '/' | '=' -> true
+ | _ -> ch1 = xInParens
+
+(* We also guard characters that might interfere with the ssreflect *)
+(* tactic syntax. *)
+let pr_guarded guard prc c =
+ pp_with Format.str_formatter (prc c);
+ let s = Format.flush_str_formatter () ^ "$" in
+ if guard s (skip_wschars s 0) then pr_paren prc c else prc c
+
+let prl_constr_expr = Ppconstr.pr_lconstr_expr
+let pr_glob_constr c = Printer.pr_glob_constr_env (Global.env ()) c
+let prl_glob_constr c = Printer.pr_lglob_constr_env (Global.env ()) c
+let pr_glob_constr_and_expr = function
+ | _, Some c -> Ppconstr.pr_constr_expr c
+ | c, None -> pr_glob_constr c
+let pr_term (k, c) = pr_guarded (guard_term k) pr_glob_constr_and_expr c
+
+let pr_hyp (SsrHyp (_, id)) = Id.print id
+
+let pr_occ = function
+ | Some (true, occ) -> str "{-" ++ pr_list pr_spc int occ ++ str "}"
+ | Some (false, occ) -> str "{+" ++ pr_list pr_spc int occ ++ str "}"
+ | None -> str "{}"
+
+(* 0 cost pp function. Active only if Debug Ssreflect is Set *)
+let ppdebug_ref = ref (fun _ -> ())
+let ssr_pp s = Feedback.msg_debug (str"SSR: "++Lazy.force s)
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optname = "ssreflect debugging";
+ Goptions.optkey = ["Debug";"Ssreflect"];
+ Goptions.optdepr = false;
+ Goptions.optread = (fun _ -> !ppdebug_ref == ssr_pp);
+ Goptions.optwrite = (fun b ->
+ Ssrmatching.debug b;
+ if b then ppdebug_ref := ssr_pp else ppdebug_ref := fun _ -> ()) }
+let ppdebug s = !ppdebug_ref s
diff --git a/plugins/ssr/ssrprinters.mli b/plugins/ssr/ssrprinters.mli
new file mode 100644
index 000000000..56ec145ad
--- /dev/null
+++ b/plugins/ssr/ssrprinters.mli
@@ -0,0 +1,45 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Ssrast
+
+val pp_term :
+ Proof_type.goal Evd.sigma -> EConstr.constr -> Pp.std_ppcmds
+
+val pr_spc : unit -> Pp.std_ppcmds
+val pr_bar : unit -> Pp.std_ppcmds
+val pr_list :
+ (unit -> Pp.std_ppcmds) -> ('a -> Pp.std_ppcmds) -> 'a list -> Pp.std_ppcmds
+
+val pp_concat :
+ Pp.std_ppcmds ->
+ ?sep:Pp.std_ppcmds -> Pp.std_ppcmds list -> Pp.std_ppcmds
+
+val xInParens : ssrtermkind
+val xWithAt : ssrtermkind
+val xNoFlag : ssrtermkind
+val xCpattern : ssrtermkind
+
+val pr_term :
+ ssrtermkind * (Glob_term.glob_constr * Constrexpr.constr_expr option) ->
+ Pp.std_ppcmds
+
+val pr_hyp : ssrhyp -> Pp.std_ppcmds
+
+val prl_constr_expr : Constrexpr.constr_expr -> Pp.std_ppcmds
+val prl_glob_constr : Glob_term.glob_constr -> Pp.std_ppcmds
+
+val pr_guarded :
+ (string -> int -> bool) -> ('a -> Pp.std_ppcmds) -> 'a -> Pp.std_ppcmds
+
+val pr_occ : ssrocc -> Pp.std_ppcmds
+
+val ppdebug : Pp.std_ppcmds Lazy.t -> unit
+
diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml
new file mode 100644
index 000000000..0fe8fdc26
--- /dev/null
+++ b/plugins/ssr/ssrtacticals.ml
@@ -0,0 +1,160 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Constr
+open Termops
+open Tacmach
+open Misctypes
+open Locusops
+
+open Ssrast
+open Ssrcommon
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
+(** Tacticals (+, -, *, done, by, do, =>, first, and last). *)
+
+let get_index = function ArgArg i -> i | _ ->
+ anomaly "Uninterpreted index"
+(* Toplevel constr must be globalized twice ! *)
+
+(** The "first" and "last" tacticals. *)
+
+let tclPERM perm tac gls =
+ let subgls = tac gls in
+ let sigma, subgll = Refiner.unpackage subgls in
+ let subgll' = perm subgll in
+ Refiner.repackage sigma subgll'
+
+let rot_hyps dir i hyps =
+ let n = List.length hyps in
+ if i = 0 then List.rev hyps else
+ if i > n then CErrors.user_err (Pp.str "Not enough subgoals") else
+ let rec rot i l_hyps = function
+ | hyp :: hyps' when i > 0 -> rot (i - 1) (hyp :: l_hyps) hyps'
+ | hyps' -> hyps' @ (List.rev l_hyps) in
+ rot (match dir with L2R -> i | R2L -> n - i) [] hyps
+
+let tclSEQAT ist atac1 dir (ivar, ((_, atacs2), atac3)) =
+ let i = get_index ivar in
+ let evtac = ssrevaltac ist in
+ let tac1 = evtac atac1 in
+ if atacs2 = [] && atac3 <> None then tclPERM (rot_hyps dir i) tac1 else
+ let evotac = function Some atac -> evtac atac | _ -> Tacticals.tclIDTAC in
+ let tac3 = evotac atac3 in
+ let rec mk_pad n = if n > 0 then tac3 :: mk_pad (n - 1) else [] in
+ match dir, mk_pad (i - 1), List.map evotac atacs2 with
+ | L2R, [], [tac2] when atac3 = None -> Tacticals.tclTHENFIRST tac1 tac2
+ | L2R, [], [tac2] when atac3 = None -> Tacticals.tclTHENLAST tac1 tac2
+ | L2R, pad, tacs2 -> Tacticals.tclTHENSFIRSTn tac1 (Array.of_list (pad @ tacs2)) tac3
+ | R2L, pad, tacs2 -> Tacticals.tclTHENSLASTn tac1 tac3 (Array.of_list (tacs2 @ pad))
+
+(** The "in" pseudo-tactical {{{ **********************************************)
+
+let hidden_goal_tag = "the_hidden_goal"
+
+let check_wgen_uniq gens =
+ let clears = List.flatten (List.map fst gens) in
+ check_hyps_uniq [] clears;
+ let ids = CList.map_filter
+ (function (_,Some ((id,_),_)) -> Some (hoi_id id) | _ -> None) gens in
+ let rec check ids = function
+ | id :: _ when List.mem id ids ->
+ errorstrm Pp.(str"Duplicate generalization " ++ Id.print id)
+ | id :: hyps -> check (id :: ids) hyps
+ | [] -> () in
+ check [] ids
+
+let pf_clauseids gl gens clseq =
+ let keep_clears = List.map (fun (x, _) -> x, None) in
+ if gens <> [] then (check_wgen_uniq gens; gens) else
+ if clseq <> InAll && clseq <> InAllHyps then keep_clears gens else
+ CErrors.user_err (Pp.str "assumptions should be named explicitly")
+
+let hidden_clseq = function InHyps | InHypsSeq | InAllHyps -> true | _ -> false
+
+let settac id c = Tactics.letin_tac None (Name id) c None
+let posetac id cl = Proofview.V82.of_tactic (settac id cl nowhere)
+
+let hidetacs clseq idhide cl0 =
+ if not (hidden_clseq clseq) then [] else
+ [posetac idhide cl0;
+ Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkVar idhide))]
+
+let endclausestac id_map clseq gl_id cl0 gl =
+ let not_hyp' id = not (List.mem_assoc id id_map) in
+ let orig_id id = try List.assoc id id_map with _ -> id in
+ let dc, c = EConstr.decompose_prod_assum (project gl) (pf_concl gl) in
+ let hide_goal = hidden_clseq clseq in
+ let c_hidden = hide_goal && EConstr.eq_constr (project gl) c (EConstr.mkVar gl_id) in
+ let rec fits forced = function
+ | (id, _) :: ids, decl :: dc' when RelDecl.get_name decl = Name id ->
+ fits true (ids, dc')
+ | 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 ->
+ EConstr.mkProd (Name (orig_id id), unmark t, unmark c')
+ | 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 =
+ Proofview.V82.of_tactic
+ (Tactics.convert_hyp_no_check (NamedDecl.map_constr unmark hyp)) in
+ let utacs = List.map utac (pf_hyps gl) in
+ let ugtac gl' =
+ Proofview.V82.of_tactic
+ (convert_concl_no_check (unmark (pf_concl gl'))) gl' in
+ let ctacs = if hide_goal then [Proofview.V82.of_tactic (Tactics.clear [gl_id])] else [] in
+ let mktac itacs = Tacticals.tclTHENLIST (itacs @ utacs @ ugtac :: ctacs) in
+ let itac (_, id) = Proofview.V82.of_tactic (Tactics.introduction id) in
+ if fits false (id_map, List.rev dc) then mktac (List.map itac id_map) gl else
+ let all_ids = ids_of_rel_context dc @ pf_ids_of_hyps gl in
+ if List.for_all not_hyp' all_ids && not c_hidden then mktac [] gl else
+ CErrors.user_err (Pp.str "tampering with discharged assumptions of \"in\" tactical")
+
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+
+let tclCLAUSES ist tac (gens, clseq) gl =
+ if clseq = InGoal || clseq = InSeqGoal then tac gl else
+ let clr_gens = pf_clauseids gl gens clseq in
+ let clear = Tacticals.tclTHENLIST (List.rev(List.fold_right clr_of_wgen clr_gens [])) in
+ let gl_id = mk_anon_id hidden_goal_tag gl in
+ let cl0 = pf_concl gl in
+ let dtac gl =
+ let c = pf_concl gl in
+ let gl, args, c =
+ List.fold_right (abs_wgen true ist mk_discharged_id) gens (gl,[], c) in
+ apply_type c args gl in
+ let endtac =
+ let id_map = CList.map_filter (function
+ | _, Some ((x,_),_) -> let id = hoi_id x in Some (mk_discharged_id id, id)
+ | _, None -> None) gens in
+ endclausestac id_map clseq gl_id cl0 in
+ Tacticals.tclTHENLIST (hidetacs clseq gl_id cl0 @ [dtac; clear; tac; endtac]) gl
+
+(** The "do" tactical. ********************************************************)
+
+let hinttac ist is_by (is_or, atacs) =
+ let dtac = if is_by then donetac ~-1 else Tacticals.tclIDTAC in
+ let mktac = function
+ | Some atac -> Tacticals.tclTHEN (ssrevaltac ist atac) dtac
+ | _ -> dtac in
+ match List.map mktac atacs with
+ | [] -> if is_or then dtac else Tacticals.tclIDTAC
+ | [tac] -> tac
+ | tacs -> Tacticals.tclFIRST tacs
+
+let ssrdotac ist (((n, m), tac), clauses) =
+ let mul = get_index n, m in
+ tclCLAUSES ist (tclMULT mul (hinttac ist false tac)) clauses
diff --git a/plugins/ssr/ssrtacticals.mli b/plugins/ssr/ssrtacticals.mli
new file mode 100644
index 000000000..b8e95b2b1
--- /dev/null
+++ b/plugins/ssr/ssrtacticals.mli
@@ -0,0 +1,44 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+val tclSEQAT :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ Ltac_plugin.Tacinterp.Value.t ->
+ Ssrast.ssrdir ->
+ int Misctypes.or_var *
+ (('a * Ltac_plugin.Tacinterp.Value.t option list) *
+ Ltac_plugin.Tacinterp.Value.t option) ->
+ Proof_type.tactic
+
+val tclCLAUSES :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ Proofview.V82.tac ->
+ (Ssrast.ssrhyps *
+ ((Ssrast.ssrhyp_or_id * string) *
+ Ssrmatching_plugin.Ssrmatching.cpattern option)
+ option)
+ list * Ssrast.ssrclseq ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
+val hinttac :
+ Tacinterp.interp_sign ->
+ bool -> bool * Tacinterp.Value.t option list -> Ssrast.v82tac
+
+val ssrdotac :
+ Ltac_plugin.Tacinterp.interp_sign ->
+ ((int Misctypes.or_var * Ssrast.ssrmmod) *
+ (bool * Ltac_plugin.Tacinterp.Value.t option list)) *
+ ((Ssrast.ssrhyps *
+ ((Ssrast.ssrhyp_or_id * string) *
+ Ssrmatching_plugin.Ssrmatching.cpattern option)
+ option)
+ list * Ssrast.ssrclseq) ->
+ Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma
+
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
new file mode 100644
index 000000000..b154cf217
--- /dev/null
+++ b/plugins/ssr/ssrvernac.ml4
@@ -0,0 +1,600 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Names
+open Term
+open Termops
+open Constrexpr
+open Constrexpr_ops
+open Pcoq
+open Pcoq.Prim
+open Pcoq.Constr
+open Pcoq.Vernac_
+open Ltac_plugin
+open Notation_ops
+open Notation_term
+open Glob_term
+open Globnames
+open Stdarg
+open Genarg
+open Misctypes
+open Decl_kinds
+open Libnames
+open Pp
+open Ppconstr
+open Printer
+open Util
+open Extraargs
+open Evar_kinds
+open Ssrprinters
+open Ssrcommon
+open Ssrparser
+DECLARE PLUGIN "ssreflect_plugin"
+
+let (!@) = Pcoq.to_coqloc
+
+(* 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 () ;;
+
+(* global syntactic changes and vernacular commands *)
+
+(** Alternative notations for "match" and anonymous arguments. {{{ ************)
+
+(* Syntax: *)
+(* if <term> is <pattern> then ... else ... *)
+(* if <term> is <pattern> [in ..] return ... then ... else ... *)
+(* let: <pattern> := <term> in ... *)
+(* let: <pattern> [in ...] := <term> return ... in ... *)
+(* The scope of a top-level 'as' in the pattern extends over the *)
+(* 'return' type (dependent if/let). *)
+(* Note that the optional "in ..." appears next to the <pattern> *)
+(* rather than the <term> in then "let:" syntax. The alternative *)
+(* would lead to ambiguities in, e.g., *)
+(* let: p1 := (*v---INNER LET:---v *) *)
+(* let: p2 := let: p3 := e3 in k return t in k2 in k1 return t' *)
+(* in b (*^--ALTERNATIVE INNER LET--------^ *) *)
+
+(* Caveat : There is no pretty-printing support, since this would *)
+(* require a modification to the Coq kernel (adding a new match *)
+(* display style -- why aren't these strings?); also, the v8.1 *)
+(* pretty-printer only allows extension hooks for printing *)
+(* integer or string literals. *)
+(* Also note that in the v8 grammar "is" needs to be a keyword; *)
+(* as this can't be done from an ML extension file, the new *)
+(* syntax will only work when ssreflect.v is imported. *)
+
+let no_ct = None, None and no_rt = None in
+let aliasvar = function
+ | [_, [{ CAst.v = CPatAlias (_, id); loc }]] -> Some (loc,Name id)
+ | _ -> None in
+let mk_cnotype mp = aliasvar mp, None in
+let mk_ctype mp t = aliasvar mp, Some t in
+let mk_rtype t = Some t in
+let mk_dthen ?loc (mp, ct, rt) c = (Loc.tag ?loc (mp, c)), ct, rt in
+let mk_let ?loc rt ct mp c1 =
+ CAst.make ?loc @@ CCases (LetPatternStyle, rt, ct, [Loc.tag ?loc (mp, c1)]) in
+let mk_pat c (na, t) = (c, na, t) in
+GEXTEND Gram
+ GLOBAL: binder_constr;
+ ssr_rtype: [[ "return"; t = operconstr LEVEL "100" -> mk_rtype t ]];
+ ssr_mpat: [[ p = pattern -> [Loc.tag ~loc:!@loc [p]] ]];
+ ssr_dpat: [
+ [ mp = ssr_mpat; "in"; t = pattern; rt = ssr_rtype -> mp, mk_ctype mp t, rt
+ | mp = ssr_mpat; rt = ssr_rtype -> mp, mk_cnotype mp, rt
+ | mp = ssr_mpat -> mp, no_ct, no_rt
+ ] ];
+ ssr_dthen: [[ dp = ssr_dpat; "then"; c = lconstr -> mk_dthen ~loc:!@loc dp c ]];
+ ssr_elsepat: [[ "else" -> [Loc.tag ~loc:!@loc [CAst.make ~loc:!@loc @@ CPatAtom None]] ]];
+ ssr_else: [[ mp = ssr_elsepat; c = lconstr -> Loc.tag ~loc:!@loc (mp, c) ]];
+ binder_constr: [
+ [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else ->
+ let b1, ct, rt = db1 in CAst.make ~loc:!@loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2])
+ | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else ->
+ let b1, ct, rt = db1 in
+ let b1, b2 =
+ let (l1, (p1, r1)), (l2, (p2, r2)) = b1, b2 in (l1, (p1, r2)), (l2, (p2, r1)) in
+ CAst.make ~loc:!@loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2])
+ | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; "in"; c1 = lconstr ->
+ mk_let ~loc:!@loc no_rt [mk_pat c no_ct] mp c1
+ | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr;
+ rt = ssr_rtype; "in"; c1 = lconstr ->
+ mk_let ~loc:!@loc rt [mk_pat c (mk_cnotype mp)] mp c1
+ | "let"; ":"; mp = ssr_mpat; "in"; t = pattern; ":="; c = lconstr;
+ rt = ssr_rtype; "in"; c1 = lconstr ->
+ mk_let ~loc:!@loc rt [mk_pat c (mk_ctype mp t)] mp c1
+ ] ];
+END
+
+GEXTEND Gram
+ GLOBAL: closed_binder;
+ closed_binder: [
+ [ ["of" | "&"]; c = operconstr LEVEL "99" ->
+ [CLocalAssum ([Loc.tag ~loc:!@loc Anonymous], Default Explicit, c)]
+ ] ];
+END
+(* }}} *)
+
+(** Vernacular commands: Prenex Implicits and Search {{{ **********************)
+
+(* This should really be implemented as an extension to the implicit *)
+(* arguments feature, but unfortuately that API is sealed. The current *)
+(* workaround uses a combination of notations that works reasonably, *)
+(* with the following caveats: *)
+(* - The pretty-printing always elides prenex implicits, even when *)
+(* they are obviously needed. *)
+(* - Prenex Implicits are NEVER exported from a module, because this *)
+(* would lead to faulty pretty-printing and scoping errors. *)
+(* - The command "Import Prenex Implicits" can be used to reassert *)
+(* Prenex Implicits for all the visible constants that had been *)
+(* declared as Prenex Implicits. *)
+
+let declare_one_prenex_implicit locality f =
+ let fref =
+ try Smartlocate.global_with_alias f
+ with _ -> errorstrm (pr_reference f ++ str " is not declared") in
+ let rec loop = function
+ | a :: args' when Impargs.is_status_implicit a ->
+ (ExplByName (Impargs.name_of_implicit a), (true, true, true)) :: loop args'
+ | args' when List.exists Impargs.is_status_implicit args' ->
+ errorstrm (str "Expected prenex implicits for " ++ pr_reference f)
+ | _ -> [] in
+ let impls =
+ match Impargs.implicits_of_global fref with
+ | [cond,impls] -> impls
+ | [] -> errorstrm (str "Expected some implicits for " ++ pr_reference f)
+ | _ -> errorstrm (str "Multiple implicits not supported") in
+ match loop impls with
+ | [] ->
+ errorstrm (str "Expected some implicits for " ++ pr_reference f)
+ | impls ->
+ Impargs.declare_manual_implicits locality fref ~enriching:false [impls]
+
+VERNAC COMMAND EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF
+ | [ "Prenex" "Implicits" ne_global_list(fl) ]
+ -> [ let locality =
+ Locality.make_section_locality (Locality.LocalityFixme.consume ()) in
+ List.iter (declare_one_prenex_implicit locality) fl ]
+END
+
+(* Vernac grammar visibility patch *)
+
+GEXTEND Gram
+ GLOBAL: gallina_ext;
+ gallina_ext:
+ [ [ IDENT "Import"; IDENT "Prenex"; IDENT "Implicits" ->
+ Vernacexpr.VernacUnsetOption (["Printing"; "Implicit"; "Defensive"])
+ ] ]
+ ;
+END
+
+(** Extend Search to subsume SearchAbout, also adding hidden Type coercions. *)
+
+(* Main prefilter *)
+
+type raw_glob_search_about_item =
+ | RGlobSearchSubPattern of constr_expr
+ | RGlobSearchString of Loc.t * string * string option
+
+let pr_search_item = function
+ | RGlobSearchString (_,s,_) -> str s
+ | RGlobSearchSubPattern p -> pr_constr_expr p
+
+let wit_ssr_searchitem = add_genarg "ssr_searchitem" pr_search_item
+
+let pr_ssr_search_item _ _ _ = pr_search_item
+
+(* Workaround the notation API that can only print notations *)
+
+let is_ident s = try CLexer.check_ident s; true with _ -> false
+
+let is_ident_part s = is_ident ("H" ^ s)
+
+let interp_search_notation ?loc tag okey =
+ let err msg = CErrors.user_err ?loc ~hdr:"interp_search_notation" msg in
+ let mk_pntn s for_key =
+ let n = String.length s in
+ let s' = Bytes.make (n + 2) ' ' in
+ let rec loop i i' =
+ if i >= n then s', i' - 2 else if s.[i] = ' ' then loop (i + 1) i' else
+ let j = try String.index_from s (i + 1) ' ' with _ -> n in
+ let m = j - i in
+ if s.[i] = '\'' && i < j - 2 && s.[j - 1] = '\'' then
+ (String.blit s (i + 1) s' i' (m - 2); loop (j + 1) (i' + m - 1))
+ else if for_key && is_ident (String.sub s i m) then
+ (Bytes.set s' i' '_'; loop (j + 1) (i' + 2))
+ else (String.blit s i s' i' m; loop (j + 1) (i' + m + 1)) in
+ loop 0 1 in
+ let trim_ntn (pntn, m) = Bytes.sub_string pntn 1 (max 0 m) in
+ let pr_ntn ntn = str "(" ++ str ntn ++ str ")" in
+ let pr_and_list pr = function
+ | [x] -> pr x
+ | x :: lx -> pr_list pr_comma pr lx ++ pr_comma () ++ str "and " ++ pr x
+ | [] -> mt () in
+ let pr_sc sc = str (if sc = "" then "independently" else sc) in
+ let pr_scs = function
+ | [""] -> pr_sc ""
+ | scs -> str "in " ++ pr_and_list pr_sc scs in
+ let generator, pr_tag_sc =
+ let ign _ = mt () in match okey with
+ | Some key ->
+ let sc = Notation.find_delimiters_scope ?loc key in
+ let pr_sc s_in = str s_in ++ spc() ++ str sc ++ pr_comma() in
+ Notation.pr_scope ign sc, pr_sc
+ | None -> Notation.pr_scopes ign, ign in
+ let qtag s_in = pr_tag_sc s_in ++ qstring tag ++ spc()in
+ let ptag, ttag =
+ let ptag, m = mk_pntn tag false in
+ if m <= 0 then err (str "empty notation fragment");
+ ptag, trim_ntn (ptag, m) in
+ let last = ref "" and last_sc = ref "" in
+ let scs = ref [] and ntns = ref [] in
+ let push_sc sc = match !scs with
+ | "" :: scs' -> scs := "" :: sc :: scs'
+ | scs' -> scs := sc :: scs' in
+ let get s _ _ = match !last with
+ | "Scope " -> last_sc := s; last := ""
+ | "Lonely notation" -> last_sc := ""; last := ""
+ | "\"" ->
+ let pntn, m = mk_pntn s true in
+ if String.string_contains ~where:(Bytes.to_string pntn) ~what:(Bytes.to_string ptag) then begin
+ let ntn = trim_ntn (pntn, m) in
+ match !ntns with
+ | [] -> ntns := [ntn]; scs := [!last_sc]
+ | ntn' :: _ when ntn' = ntn -> push_sc !last_sc
+ | _ when ntn = ttag -> ntns := ntn :: !ntns; scs := [!last_sc]
+ | _ :: ntns' when List.mem ntn ntns' -> ()
+ | ntn' :: ntns' -> ntns := ntn' :: ntn :: ntns'
+ end;
+ last := ""
+ | _ -> last := s in
+ pp_with (Format.make_formatter get (fun _ -> ())) generator;
+ let ntn = match !ntns with
+ | [] ->
+ err (hov 0 (qtag "in" ++ str "does not occur in any notation"))
+ | ntn :: ntns' when ntn = ttag ->
+ if ntns' <> [] then begin
+ let pr_ntns' = pr_and_list pr_ntn ntns' in
+ Feedback.msg_warning (hov 4 (qtag "In" ++ str "also occurs in " ++ pr_ntns'))
+ end; ntn
+ | [ntn] ->
+ Feedback.msg_info (hov 4 (qtag "In" ++ str "is part of notation " ++ pr_ntn ntn)); ntn
+ | ntns' ->
+ let e = str "occurs in" ++ spc() ++ pr_and_list pr_ntn ntns' in
+ err (hov 4 (str "ambiguous: " ++ qtag "in" ++ e)) in
+ let (nvars, body), ((_, pat), osc) = match !scs with
+ | [sc] -> Notation.interp_notation ?loc ntn (None, [sc])
+ | scs' ->
+ try Notation.interp_notation ?loc ntn (None, []) with _ ->
+ let e = pr_ntn ntn ++ spc() ++ str "is defined " ++ pr_scs scs' in
+ err (hov 4 (str "ambiguous: " ++ pr_tag_sc "in" ++ e)) in
+ let sc = Option.default "" osc in
+ let _ =
+ let m_sc =
+ if osc <> None then str "In " ++ str sc ++ pr_comma() else mt() in
+ let ntn_pat = trim_ntn (mk_pntn pat false) in
+ let rbody = glob_constr_of_notation_constr ?loc body in
+ let m_body = hov 0 (Constrextern.without_symbols prl_glob_constr rbody) in
+ let m = m_sc ++ pr_ntn ntn_pat ++ spc () ++ str "denotes " ++ m_body in
+ Feedback.msg_info (hov 0 m) in
+ if List.length !scs > 1 then
+ let scs' = List.remove (=) sc !scs in
+ let w = pr_ntn ntn ++ str " is also defined " ++ pr_scs scs' in
+ Feedback.msg_warning (hov 4 w)
+ else if String.string_contains ~where:ntn ~what:" .. " then
+ err (pr_ntn ntn ++ str " is an n-ary notation");
+ let nvars = List.filter (fun (_,(_,typ)) -> typ = NtnTypeConstr) nvars in
+ let rec sub () = function
+ | NVar x when List.mem_assoc x nvars -> CAst.make ?loc @@ GPatVar (FirstOrderPatVar x)
+ | c ->
+ glob_constr_of_notation_constr_with_binders ?loc (fun _ x -> (), x) sub () c in
+ let _, npat = Patternops.pattern_of_glob_constr (sub () body) in
+ Search.GlobSearchSubPattern npat
+
+ARGUMENT EXTEND ssr_search_item TYPED AS ssr_searchitem
+ PRINTED BY pr_ssr_search_item
+ | [ string(s) ] -> [ RGlobSearchString (loc,s,None) ]
+ | [ string(s) "%" preident(key) ] -> [ RGlobSearchString (loc,s,Some key) ]
+ | [ constr_pattern(p) ] -> [ RGlobSearchSubPattern p ]
+END
+
+let pr_ssr_search_arg _ _ _ =
+ let pr_item (b, p) = str (if b then "-" else "") ++ pr_search_item p in
+ pr_list spc pr_item
+
+ARGUMENT EXTEND ssr_search_arg TYPED AS (bool * ssr_searchitem) list
+ PRINTED BY pr_ssr_search_arg
+ | [ "-" ssr_search_item(p) ssr_search_arg(a) ] -> [ (false, p) :: a ]
+ | [ ssr_search_item(p) ssr_search_arg(a) ] -> [ (true, p) :: a ]
+ | [ ] -> [ [] ]
+END
+
+(* Main type conclusion pattern filter *)
+
+let rec splay_search_pattern na = function
+ | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp
+ | Pattern.PLetIn (_, _, _, bp) -> splay_search_pattern na bp
+ | Pattern.PRef hr -> hr, na
+ | _ -> CErrors.user_err (Pp.str "no head constant in head search pattern")
+
+let push_rels_assum l e =
+ let l = List.map (fun (n,t) -> n, EConstr.Unsafe.to_constr t) l in
+ push_rels_assum l e
+
+let coerce_search_pattern_to_sort hpat =
+ let env = Global.env () and sigma = Evd.empty in
+ let mkPApp fp n_imps args =
+ let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in
+ Pattern.PApp (fp, args') in
+ let hr, na = splay_search_pattern 0 hpat in
+ let dc, ht =
+ Reductionops.splay_prod env sigma (EConstr.of_constr (Universes.unsafe_type_of_global hr)) in
+ let np = List.length dc in
+ if np < na then CErrors.user_err (Pp.str "too many arguments in head search pattern") else
+ let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in
+ let warn () =
+ Feedback.msg_warning (str "Listing only lemmas with conclusion matching " ++
+ pr_constr_pattern hpat') in
+ if EConstr.isSort sigma ht then begin warn (); true, hpat' end else
+ let filter_head, coe_path =
+ try
+ let _, cp =
+ Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in
+ warn ();
+ true, cp
+ with _ -> false, [] in
+ let coerce hp coe_index =
+ let coe = Classops.get_coercion_value coe_index in
+ try
+ let coe_ref = reference_of_constr coe in
+ let n_imps = Option.get (Classops.hide_coercion coe_ref) in
+ mkPApp (Pattern.PRef coe_ref) n_imps [|hp|]
+ with _ ->
+ errorstrm (str "need explicit coercion " ++ pr_constr coe ++ spc ()
+ ++ str "to interpret head search pattern as type") in
+ filter_head, List.fold_left coerce hpat' coe_path
+
+let interp_head_pat hpat =
+ let filter_head, p = coerce_search_pattern_to_sort hpat in
+ let rec loop c = match kind_of_term c with
+ | Cast (c', _, _) -> loop c'
+ | Prod (_, _, c') -> loop c'
+ | LetIn (_, _, _, c') -> loop c'
+ | _ -> Constr_matching.is_matching (Global.env()) Evd.empty p (EConstr.of_constr c) in
+ filter_head, loop
+
+let all_true _ = true
+
+let rec interp_search_about args accu = match args with
+| [] -> accu
+| (flag, arg) :: rem ->
+ fun gr env typ ->
+ let ans = Search.search_about_filter arg gr env typ in
+ (if flag then ans else not ans) && interp_search_about rem accu gr env typ
+
+let interp_search_arg arg =
+ let arg = List.map (fun (x,arg) -> x, match arg with
+ | RGlobSearchString (loc,s,key) ->
+ if is_ident_part s then Search.GlobSearchString s else
+ interp_search_notation ~loc s key
+ | RGlobSearchSubPattern p ->
+ try
+ let intern = Constrintern.intern_constr_pattern in
+ Search.GlobSearchSubPattern (snd (intern (Global.env()) p))
+ with e -> let e = CErrors.push e in iraise (ExplainErr.process_vernac_interp_error e)) arg in
+ let hpat, a1 = match arg with
+ | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a'
+ | (true, Search.GlobSearchSubPattern p) :: a' ->
+ let filter_head, p = interp_head_pat p in
+ if filter_head then p, a' else all_true, arg
+ | _ -> all_true, arg in
+ let is_string =
+ function (_, Search.GlobSearchString _) -> true | _ -> false in
+ let a2, a3 = List.partition is_string a1 in
+ interp_search_about (a2 @ a3) (fun gr env typ -> hpat typ)
+
+(* Module path postfilter *)
+
+let pr_modloc (b, m) = if b then str "-" ++ pr_reference m else pr_reference m
+
+let wit_ssrmodloc = add_genarg "ssrmodloc" pr_modloc
+
+let pr_ssr_modlocs _ _ _ ml =
+ if ml = [] then str "" else spc () ++ str "in " ++ pr_list spc pr_modloc ml
+
+ARGUMENT EXTEND ssr_modlocs TYPED AS ssrmodloc list PRINTED BY pr_ssr_modlocs
+ | [ ] -> [ [] ]
+END
+
+GEXTEND Gram
+ GLOBAL: ssr_modlocs;
+ modloc: [[ "-"; m = global -> true, m | m = global -> false, m]];
+ ssr_modlocs: [[ "in"; ml = LIST1 modloc -> ml ]];
+END
+
+let interp_modloc mr =
+ let interp_mod (_, mr) =
+ let (loc, qid) = qualid_of_reference mr in
+ try Nametab.full_name_module qid with Not_found ->
+ CErrors.user_err ?loc (str "No Module " ++ pr_qualid qid) in
+ let mr_out, mr_in = List.partition fst mr in
+ let interp_bmod b = function
+ | [] -> fun _ _ _ -> true
+ | rmods -> Search.module_filter (List.map interp_mod rmods, b) in
+ let is_in = interp_bmod false mr_in and is_out = interp_bmod true mr_out in
+ fun gr env typ -> is_in gr env typ && is_out gr env typ
+
+(* The unified, extended vernacular "Search" command *)
+
+let ssrdisplaysearch gr env t =
+ let pr_res = pr_global gr ++ spc () ++ str " " ++ pr_lconstr_env env Evd.empty t in
+ Feedback.msg_info (hov 2 pr_res ++ fnl ())
+
+VERNAC COMMAND EXTEND SsrSearchPattern CLASSIFIED AS QUERY
+| [ "Search" ssr_search_arg(a) ssr_modlocs(mr) ] ->
+ [ let hpat = interp_search_arg a in
+ let in_mod = interp_modloc mr in
+ let post_filter gr env typ = in_mod gr env typ && hpat gr env typ in
+ let display gr env typ =
+ if post_filter gr env typ then ssrdisplaysearch gr env typ
+ in
+ Search.generic_search None display ]
+END
+
+(* }}} *)
+
+(** View hint database and View application. {{{ ******************************)
+
+(* There are three databases of lemmas used to mediate the application *)
+(* of reflection lemmas: one for forward chaining, one for backward *)
+(* chaining, and one for secondary backward chaining. *)
+
+(* View hints *)
+
+let pr_raw_ssrhintref prc _ _ = let open CAst in function
+ | { v = CAppExpl ((None, r,x), args) } when isCHoles args ->
+ prc (CAst.make @@ CRef (r,x)) ++ str "|" ++ int (List.length args)
+ | { v = CApp ((_, { v = CRef _ }), _) } as c -> prc c
+ | { v = CApp ((_, c), args) } when isCxHoles args ->
+ prc c ++ str "|" ++ int (List.length args)
+ | c -> prc c
+
+let pr_rawhintref = let open CAst in function
+ | { v = GApp (f, args) } when isRHoles args ->
+ pr_glob_constr f ++ str "|" ++ int (List.length args)
+ | c -> pr_glob_constr c
+
+let pr_glob_ssrhintref _ _ _ (c, _) = pr_rawhintref c
+
+let pr_ssrhintref prc _ _ = prc
+
+let mkhintref ?loc c n = match c.CAst.v with
+ | CRef (r,x) -> CAst.make ?loc @@ CAppExpl ((None, r, x), mkCHoles ?loc n)
+ | _ -> mkAppC (c, mkCHoles ?loc n)
+
+ARGUMENT EXTEND ssrhintref
+ PRINTED BY pr_ssrhintref
+ RAW_TYPED AS constr RAW_PRINTED BY pr_raw_ssrhintref
+ GLOB_TYPED AS constr GLOB_PRINTED BY pr_glob_ssrhintref
+ | [ constr(c) ] -> [ c ]
+ | [ constr(c) "|" natural(n) ] -> [ mkhintref ~loc c n ]
+END
+
+(* View purpose *)
+
+let pr_viewpos = function
+ | 0 -> str " for move/"
+ | 1 -> str " for apply/"
+ | 2 -> str " for apply//"
+ | _ -> mt ()
+
+let pr_ssrviewpos _ _ _ = pr_viewpos
+
+let mapviewpos f n k = if n < 3 then f n else for i = 0 to k - 1 do f i done
+
+ARGUMENT EXTEND ssrviewpos TYPED AS int PRINTED BY pr_ssrviewpos
+ | [ "for" "move" "/" ] -> [ 0 ]
+ | [ "for" "apply" "/" ] -> [ 1 ]
+ | [ "for" "apply" "/" "/" ] -> [ 2 ]
+ | [ "for" "apply" "//" ] -> [ 2 ]
+ | [ ] -> [ 3 ]
+END
+
+let pr_ssrviewposspc _ _ _ i = pr_viewpos i ++ spc ()
+
+ARGUMENT EXTEND ssrviewposspc TYPED AS ssrviewpos PRINTED BY pr_ssrviewposspc
+ | [ ssrviewpos(i) ] -> [ i ]
+END
+
+let print_view_hints i =
+ let pp_viewname = str "Hint View" ++ pr_viewpos i ++ str " " in
+ let pp_hints = pr_list spc pr_rawhintref Ssrview.viewtab.(i) in
+ Feedback.msg_info (pp_viewname ++ hov 0 pp_hints ++ Pp.cut ())
+
+VERNAC COMMAND EXTEND PrintView CLASSIFIED AS QUERY
+| [ "Print" "Hint" "View" ssrviewpos(i) ] -> [ mapviewpos print_view_hints i 3 ]
+END
+
+
+VERNAC COMMAND EXTEND HintView CLASSIFIED AS SIDEFF
+ | [ "Hint" "View" ssrviewposspc(n) ne_ssrhintref_list(lvh) ] ->
+ [ mapviewpos (Ssrview.add_view_hints (Ssrview.glob_view_hints lvh)) n 2 ]
+END
+
+(* }}} *)
+
+(** Canonical Structure alias *)
+
+GEXTEND Gram
+ GLOBAL: gallina_ext;
+
+ gallina_ext:
+ (* Canonical structure *)
+ [[ IDENT "Canonical"; qid = Constr.global ->
+ Vernacexpr.VernacCanonical (AN qid)
+ | IDENT "Canonical"; ntn = Prim.by_notation ->
+ Vernacexpr.VernacCanonical (ByNotation ntn)
+ | IDENT "Canonical"; qid = Constr.global;
+ d = G_vernac.def_body ->
+ let s = coerce_reference_to_id qid in
+ Vernacexpr.VernacDefinition
+ ((Some Decl_kinds.Global,Decl_kinds.CanonicalStructure),
+ ((Loc.tag s),None),(d ))
+ ]];
+END
+
+(** Keyword compatibility fixes. *)
+
+(* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *)
+(* identifiers used as keywords. This is incompatible with ssreflect.v *)
+(* which makes "by" and "of" true keywords, because of technicalities *)
+(* in the internal lexer-parser API of Coq. We patch this here by *)
+(* adding new parsing rules that recognize the new keywords. *)
+(* To make matters worse, the Coq grammar for tactics fails to *)
+(* export the non-terminals we need to patch. Fortunately, the CamlP5 *)
+(* API provides a backdoor access (with loads of Obj.magic trickery). *)
+
+(* Coq v8.3 defines "by" as a keyword, some hacks are not needed any *)
+(* longer and thus comment out. Such comments are marked with v8.3 *)
+
+open Pltac
+
+GEXTEND Gram
+ GLOBAL: hypident;
+ hypident: [
+ [ "("; IDENT "type"; "of"; id = Prim.identref; ")" -> id, Locus.InHypTypeOnly
+ | "("; IDENT "value"; "of"; id = Prim.identref; ")" -> id, Locus.InHypValueOnly
+ ] ];
+END
+
+GEXTEND Gram
+ GLOBAL: hloc;
+hloc: [
+ [ "in"; "("; "Type"; "of"; id = ident; ")" ->
+ Tacexpr.HypLocation ((Loc.tag id), Locus.InHypTypeOnly)
+ | "in"; "("; IDENT "Value"; "of"; id = ident; ")" ->
+ Tacexpr.HypLocation ((Loc.tag id), Locus.InHypValueOnly)
+ ] ];
+END
+
+GEXTEND Gram
+ GLOBAL: constr_eval;
+ constr_eval: [
+ [ IDENT "type"; "of"; c = Constr.constr -> Genredexpr.ConstrTypeOf c ]
+ ];
+END
+
+(* We wipe out all the keywords generated by the grammar rules we defined. *)
+(* The user is supposed to Require Import ssreflect or Require ssreflect *)
+(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *)
+(* consequence the extended ssreflect grammar. *)
+let () = CLexer.set_keyword_state frozen_lexer ;;
+
+(* vim: set filetype=ocaml foldmethod=marker: *)
diff --git a/plugins/ssr/ssrvernac.mli b/plugins/ssr/ssrvernac.mli
new file mode 100644
index 000000000..58e81130c
--- /dev/null
+++ b/plugins/ssr/ssrvernac.mli
@@ -0,0 +1,9 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml
new file mode 100644
index 000000000..3c995b1bb
--- /dev/null
+++ b/plugins/ssr/ssrview.ml
@@ -0,0 +1,125 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Util
+open Names
+open Term
+open Ltac_plugin
+open Tacinterp
+open Glob_term
+open Tacmach
+open Tacticals
+
+open Ssrcommon
+
+(* The table and its display command *)
+
+(* FIXME this looks hackish *)
+
+let viewtab : glob_constr list array = Array.make 3 []
+
+let _ =
+ let init () = Array.fill viewtab 0 3 [] in
+ let freeze _ = Array.copy viewtab in
+ let unfreeze vt = Array.blit vt 0 viewtab 0 3 in
+ Summary.declare_summary "ssrview"
+ { Summary.freeze_function = freeze;
+ Summary.unfreeze_function = unfreeze;
+ Summary.init_function = init }
+
+(* Populating the table *)
+
+let cache_viewhint (_, (i, lvh)) =
+ let mem_raw h = List.exists (Glob_ops.glob_constr_eq h) in
+ let add_hint h hdb = if mem_raw h hdb then hdb else h :: hdb in
+ viewtab.(i) <- List.fold_right add_hint lvh viewtab.(i)
+
+let subst_viewhint ( subst, (i, lvh as ilvh)) =
+ let lvh' = List.smartmap (Detyping.subst_glob_constr subst) lvh in
+ if lvh' == lvh then ilvh else i, lvh'
+
+let classify_viewhint x = Libobject.Substitute x
+
+let in_viewhint =
+ Libobject.declare_object {(Libobject.default_object "VIEW_HINTS") with
+ Libobject.open_function = (fun i o -> if i = 1 then cache_viewhint o);
+ Libobject.cache_function = cache_viewhint;
+ Libobject.subst_function = subst_viewhint;
+ Libobject.classify_function = classify_viewhint }
+
+let glob_view_hints lvh =
+ List.map (Constrintern.intern_constr (Global.env ())) lvh
+
+let add_view_hints lvh i = Lib.add_anonymous_leaf (in_viewhint (i, lvh))
+
+let interp_view ist si env sigma gv v rid =
+ let open CAst in
+ match v with
+ | { v = GApp ( { v = GHole _ } , rargs); loc } ->
+ let rv = make ?loc @@ GApp (rid, rargs) in
+ snd (interp_open_constr ist (re_sig si sigma) (rv, None))
+ | rv ->
+ let interp rc rargs =
+ interp_open_constr ist (re_sig si sigma) (mkRApp rc rargs, None) in
+ let rec simple_view rargs n =
+ if n < 0 then view_error "use" gv else
+ try interp rv rargs with _ -> simple_view (mkRHole :: rargs) (n - 1) in
+ let view_nbimps = interp_view_nbimps ist (re_sig si sigma) rv in
+ let view_args = [mkRApp rv (mkRHoles view_nbimps); rid] in
+ let rec view_with = function
+ | [] -> simple_view [rid] (interp_nbargs ist (re_sig si sigma) rv)
+ | hint :: hints -> try interp hint view_args with _ -> view_with hints in
+ snd (view_with (if view_nbimps < 0 then [] else viewtab.(0)))
+
+
+let with_view ist ~next si env (gl0 : (Proof_type.goal * tac_ctx) Evd.sigma) c name cl prune (conclude : EConstr.t -> EConstr.t -> tac_ctx tac_a) clr =
+ let c2r ist x = { ist with lfun =
+ Id.Map.add top_id (Value.of_constr x) ist.lfun } in
+ let terminate (sigma, c') =
+ let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in
+ let c' = Reductionops.nf_evar sigma c' in
+ let n, c', _, ucst = without_ctx pf_abs_evars gl0 (sigma, c') in
+ let c' = if not prune then c' else without_ctx pf_abs_cterm gl0 n c' in
+ let gl0 = pf_merge_uc ucst gl0 in
+ let gl0, ap =
+ let gl0, ctx = pull_ctx gl0 in
+ let gl0, ap = pf_abs_prod name gl0 c' (Termops.prod_applist sigma cl [c]) in
+ push_ctx ctx gl0, ap in
+ let gl0 = pf_merge_uc_of sigma gl0 in
+ ap, c', gl0 in
+ let rec loop (sigma, c') = function
+ | [] ->
+ let ap, c', gl = terminate (sigma, c') in
+ ap, c', conclude ap c' gl
+ | f :: view ->
+ let ist, rid =
+ match EConstr.kind sigma c' with
+ | Var id -> ist,mkRVar id
+ | _ -> c2r ist c',mkRltacVar top_id in
+ let v = intern_term ist env f in
+ loop (interp_view ist si env sigma f v rid) view
+ in loop
+
+let pfa_with_view ist ?(next=ref []) (prune, view) cl c conclude clr gl =
+ let env, sigma, si =
+ without_ctx pf_env gl, Refiner.project gl, without_ctx sig_it gl in
+ with_view
+ ist ~next si env gl c (constr_name sigma c) cl prune conclude clr (sigma, c) view
+
+let pf_with_view_linear ist gl v cl c =
+ let x,y,gl =
+ pfa_with_view ist v cl c (fun _ _ -> tac_ctx tclIDTAC) []
+ (push_ctx (new_ctx ()) gl) in
+ let gl, _ = pull_ctxs gl in
+ assert(List.length (sig_it gl) = 1);
+ x,y,re_sig (List.hd (sig_it gl)) (Refiner.project gl)
+
+
+(* vim: set filetype=ocaml foldmethod=marker: *)
diff --git a/plugins/ssr/ssrview.mli b/plugins/ssr/ssrview.mli
new file mode 100644
index 000000000..6fd906ff4
--- /dev/null
+++ b/plugins/ssr/ssrview.mli
@@ -0,0 +1,36 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Ssrast
+open Ssrcommon
+
+val viewtab : Glob_term.glob_constr list array
+val add_view_hints : Glob_term.glob_constr list -> int -> unit
+val glob_view_hints : Constrexpr.constr_expr list -> Glob_term.glob_constr list
+
+val pfa_with_view :
+ ist ->
+ ?next:ssripats ref ->
+ bool * ssrterm list ->
+ EConstr.t ->
+ EConstr.t ->
+ (EConstr.t -> EConstr.t -> tac_ctx tac_a) ->
+ ssrhyps ->
+ (goal * tac_ctx) sigma -> EConstr.types * EConstr.t * (goal * tac_ctx) list sigma
+
+val pf_with_view_linear :
+ ist ->
+ goal sigma ->
+ bool * ssrterm list ->
+ EConstr.t ->
+ EConstr.t ->
+ EConstr.types * EConstr.t * goal sigma
+
+
diff --git a/plugins/ssr/vo.itarget b/plugins/ssr/vo.itarget
new file mode 100644
index 000000000..99f9f160b
--- /dev/null
+++ b/plugins/ssr/vo.itarget
@@ -0,0 +1,3 @@
+ssreflect.vo
+ssrfun.vo
+ssrbool.vo