summaryrefslogtreecommitdiff
path: root/plugins/ssrmatching
diff options
context:
space:
mode:
Diffstat (limited to 'plugins/ssrmatching')
-rw-r--r--plugins/ssrmatching/g_ssrmatching.ml4101
-rw-r--r--plugins/ssrmatching/ssrmatching.ml (renamed from plugins/ssrmatching/ssrmatching.ml4)184
-rw-r--r--plugins/ssrmatching/ssrmatching.mli258
-rw-r--r--plugins/ssrmatching/ssrmatching.v36
-rw-r--r--plugins/ssrmatching/ssrmatching_plugin.mlpack1
5 files changed, 460 insertions, 120 deletions
diff --git a/plugins/ssrmatching/g_ssrmatching.ml4 b/plugins/ssrmatching/g_ssrmatching.ml4
new file mode 100644
index 00000000..746c368a
--- /dev/null
+++ b/plugins/ssrmatching/g_ssrmatching.ml4
@@ -0,0 +1,101 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Ltac_plugin
+open Genarg
+open Pcoq
+open Pcoq.Constr
+open Ssrmatching
+open Ssrmatching.Internal
+
+(* 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 () ;;
+
+DECLARE PLUGIN "ssrmatching_plugin"
+
+let pr_rpattern _ _ _ = pr_rpattern
+
+ARGUMENT EXTEND rpattern
+ TYPED AS rpatternty
+ PRINTED BY pr_rpattern
+ INTERPRETED BY interp_rpattern
+ GLOBALIZED BY glob_rpattern
+ SUBSTITUTED BY subst_rpattern
+ | [ lconstr(c) ] -> [ mk_rpattern (T (mk_lterm c None)) ]
+ | [ "in" lconstr(c) ] -> [ mk_rpattern (In_T (mk_lterm c None)) ]
+ | [ lconstr(x) "in" lconstr(c) ] ->
+ [ mk_rpattern (X_In_T (mk_lterm x None, mk_lterm c None)) ]
+ | [ "in" lconstr(x) "in" lconstr(c) ] ->
+ [ mk_rpattern (In_X_In_T (mk_lterm x None, mk_lterm c None)) ]
+ | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] ->
+ [ mk_rpattern (E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) ]
+ | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] ->
+ [ mk_rpattern (E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) ]
+END
+
+let pr_ssrterm _ _ _ = pr_ssrterm
+
+ARGUMENT EXTEND cpattern
+ PRINTED BY pr_ssrterm
+ INTERPRETED BY interp_ssrterm
+ GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm
+ RAW_PRINTED BY pr_ssrterm
+ GLOB_PRINTED BY pr_ssrterm
+| [ "Qed" constr(c) ] -> [ mk_lterm c None ]
+END
+
+let input_ssrtermkind strm = match Util.stream_nth 0 strm with
+ | Tok.KEYWORD "(" -> '('
+ | Tok.KEYWORD "@" -> '@'
+ | _ -> ' '
+let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind
+
+GEXTEND Gram
+ GLOBAL: cpattern;
+ cpattern: [[ k = ssrtermkind; c = constr ->
+ let pattern = mk_term k c None in
+ if loc_of_cpattern pattern <> Some !@loc && k = '('
+ then mk_term 'x' c None
+ else pattern ]];
+END
+
+ARGUMENT EXTEND lcpattern
+ TYPED AS cpattern
+ PRINTED BY pr_ssrterm
+ INTERPRETED BY interp_ssrterm
+ GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm
+ RAW_PRINTED BY pr_ssrterm
+ GLOB_PRINTED BY pr_ssrterm
+| [ "Qed" lconstr(c) ] -> [ mk_lterm c None ]
+END
+
+GEXTEND Gram
+ GLOBAL: lcpattern;
+ lcpattern: [[ k = ssrtermkind; c = lconstr ->
+ let pattern = mk_term k c None in
+ if loc_of_cpattern pattern <> Some !@loc && k = '('
+ then mk_term 'x' c None
+ else pattern ]];
+END
+
+ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY pr_rpattern
+| [ rpattern(pat) ] -> [ pat ]
+END
+
+TACTIC EXTEND ssrinstoftpat
+| [ "ssrinstancesoftpat" cpattern(arg) ] -> [ Proofview.V82.tactic (ssrinstancesof arg) ]
+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 ;;
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml
index 307bc21a..4a63dd47 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml
@@ -10,10 +10,6 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
-(* 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 () ;;
-
open Ltac_plugin
open Names
open Pp
@@ -22,8 +18,6 @@ open Stdarg
open Term
module CoqConstr = Constr
open CoqConstr
-open Pcoq
-open Pcoq.Constr
open Vars
open Libnames
open Tactics
@@ -40,14 +34,12 @@ open Pretyping
open Ppconstr
open Printer
open Globnames
-open Misctypes
+open Namegen
open Decl_kinds
open Evar_kinds
open Constrexpr
open Constrexpr_ops
-DECLARE PLUGIN "ssrmatching_plugin"
-
let errorstrm = CErrors.user_err ~hdr:"ssrmatching"
let loc_error loc msg = CErrors.user_err ?loc ~hdr:msg (str msg)
let ppnl = Feedback.msg_info
@@ -131,9 +123,12 @@ let add_genarg tag pr =
(** Constructors for cast type *)
let dC t = CastConv t
(** Constructors for constr_expr *)
-let isCVar = function { CAst.v = CRef ({CAst.v=Ident _},_) } -> true | _ -> false
-let destCVar = function { CAst.v = CRef ({CAst.v=Ident id},_) } -> id | _ ->
- CErrors.anomaly (str"not a CRef.")
+let isCVar = function { CAst.v = CRef (qid,_) } -> qualid_is_ident qid | _ -> false
+let destCVar = function
+ | { CAst.v = CRef (qid,_) } when qualid_is_ident qid ->
+ qualid_basename qid
+ | _ ->
+ CErrors.anomaly (str"not a CRef.")
let isGLambda c = match DAst.get c with GLambda (Name _, _, _, _) -> true | _ -> false
let destGLambda c = match DAst.get c with GLambda (Name id, _, _, c) -> (id, c)
| _ -> CErrors.anomaly (str "not a GLambda")
@@ -283,7 +278,7 @@ exception NoProgress
(* comparison can be much faster than the HO one. *)
let unif_EQ env sigma p c =
- let evars = existential_opt_value sigma, Evd.universes sigma in
+ let evars = existential_opt_value0 sigma, Evd.universes sigma in
try let _ = Reduction.conv env p ~evars c in true with _ -> false
let unif_EQ_args env sigma pa a =
@@ -314,24 +309,27 @@ let unif_HO_args env ise0 pa i ca =
(* evars into metas, since 8.2 does not TC metas. This means some lossage *)
(* for HO evars, though hopefully Miller patterns can pick up some of *)
(* those cases, and HO matching will mop up the rest. *)
-let flags_FO =
+let flags_FO env =
+ let oracle = Environ.oracle env in
+ let ts = Conv_oracle.get_transp_state oracle in
let flags =
- { (Unification.default_no_delta_unify_flags ()).Unification.core_unify_flags
+ { (Unification.default_no_delta_unify_flags ts).Unification.core_unify_flags
with
Unification.modulo_conv_on_closed_terms = None;
Unification.modulo_eta = true;
Unification.modulo_betaiota = true;
- Unification.modulo_delta_types = full_transparent_state}
+ Unification.modulo_delta_types = ts }
in
{ Unification.core_unify_flags = flags;
Unification.merge_unify_flags = flags;
Unification.subterm_unify_flags = flags;
Unification.allow_K_in_toplevel_higher_order_unification = false;
Unification.resolve_evars =
- (Unification.default_no_delta_unify_flags ()).Unification.resolve_evars
+ (Unification.default_no_delta_unify_flags ts).Unification.resolve_evars
}
let unif_FO env ise p c =
- Unification.w_unify env ise Reduction.CONV ~flags:flags_FO (EConstr.of_constr p) (EConstr.of_constr c)
+ Unification.w_unify env ise Reduction.CONV ~flags:(flags_FO env)
+ (EConstr.of_constr p) (EConstr.of_constr c)
(* Perform evar substitution in main term and prune substitution. *)
let nf_open_term sigma0 ise c =
@@ -339,7 +337,7 @@ let nf_open_term sigma0 ise c =
let s = ise and s' = ref sigma0 in
let rec nf c' = match kind c' with
| Evar ex ->
- begin try nf (existential_value s ex) with _ ->
+ begin try nf (existential_value0 s ex) with _ ->
let k, a = ex in let a' = Array.map nf a in
if not (Evd.mem !s' k) then
s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k));
@@ -349,7 +347,9 @@ let nf_open_term sigma0 ise c =
let copy_def k evi () =
if evar_body evi != Evd.Evar_empty then () else
match Evd.evar_body (Evd.find s k) with
- | Evar_defined c' -> s' := Evd.define k (nf c') !s'
+ | Evar_defined c' ->
+ let c' = EConstr.of_constr (nf (EConstr.Unsafe.to_constr c')) in
+ s' := Evd.define k c' !s'
| _ -> () in
let c' = nf c in let _ = Evd.fold copy_def sigma0 () in
!s', Evd.evar_universe_context s, EConstr.of_constr c'
@@ -448,7 +448,7 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
let nenv = env_size env + if hack then 1 else 0 in
let rec put c = match kind c with
| Evar (k, a as ex) ->
- begin try put (existential_value !sigma ex)
+ begin try put (existential_value0 !sigma ex)
with NotInstantiatedEvar ->
if Evd.mem sigma0 k then map put c else
let evi = Evd.find !sigma k in
@@ -459,11 +459,13 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
| Context.Named.Declaration.LocalAssum (x, t) ->
mkVar x :: d, mkNamedProd x (put t) c in
let a, t =
- Context.Named.fold_inside abs_dc ~init:([], (put evi.evar_concl)) dc in
+ Context.Named.fold_inside abs_dc
+ ~init:([], (put @@ EConstr.Unsafe.to_constr evi.evar_concl))
+ (EConstr.Unsafe.to_named_context dc) in
let m = Evarutil.new_meta () in
- ise := meta_declare m t !ise;
- sigma := Evd.define k (applistc (mkMeta m) a) !sigma;
- put (existential_value !sigma ex)
+ ise := meta_declare m (EConstr.of_constr t) !ise;
+ sigma := Evd.define k (EConstr.of_constr (applistc (mkMeta m) a)) !sigma;
+ put (existential_value0 !sigma ex)
end
| _ -> map put c in
let c1 = put c0 in !ise, c1
@@ -543,7 +545,7 @@ let splay_app ise =
| App (f, a') -> loop f (Array.append a' a)
| Cast (c', _, _) -> loop c' a
| Evar ex ->
- (try loop (existential_value ise ex) a with _ -> c, a)
+ (try loop (existential_value0 ise ex) a with _ -> c, a)
| _ -> c, a in
fun c -> match kind c with
| App (f, a) -> loop f a
@@ -706,9 +708,9 @@ let match_upats_HO ~on_instance upats env sigma0 ise c =
;;
-let fixed_upat = function
+let fixed_upat evd = function
| {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false
-| {up_t = t} -> not (occur_existential Evd.empty (EConstr.of_constr t)) (** FIXME *)
+| {up_t = t} -> not (occur_existential evd (EConstr.of_constr t)) (** FIXME *)
let do_once r f = match !r with Some _ -> () | None -> r := Some (f ())
@@ -767,7 +769,7 @@ let mk_tpattern_matcher ?(all_instances=false)
let p2t p = mkApp(p.up_f,p.up_a) in
let source () = match upats_origin, upats with
| None, [p] ->
- (if fixed_upat p then str"term " else str"partial term ") ++
+ (if fixed_upat ise p then str"term " else str"partial term ") ++
pr_constr_pat (p2t p) ++ spc()
| Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++
pr_constr_pat rule ++ fnl() ++ ws 4 ++ pr_constr_pat (p2t p) ++ fnl()
@@ -854,7 +856,7 @@ let rec uniquize = function
let p' = mkApp (pf, pa) in
if max_occ <= !nocc then p', u.up_dir, (sigma, uc, u.up_t)
else errorstrm (str"Only " ++ int !nocc ++ str" < " ++ int max_occ ++
- str(String.plural !nocc " occurence") ++ match upats_origin with
+ str(String.plural !nocc " occurrence") ++ match upats_origin with
| None -> str" of" ++ spc() ++ pr_constr_pat p'
| Some (dir,rule) -> str" of the " ++ pr_dir_side dir ++ fnl() ++
ws 4 ++ pr_constr_pat p' ++ fnl () ++
@@ -900,7 +902,6 @@ let pr_pattern_aux pr_constr = function
let pp_pattern (sigma, p) =
pr_pattern_aux (fun t -> pr_constr_pat (EConstr.Unsafe.to_constr (pi3 (nf_open_term sigma sigma (EConstr.of_constr t))))) p
let pr_cpattern = pr_term
-let pr_rpattern _ _ _ = pr_pattern
let wit_rpatternty = add_genarg "rpatternty" pr_pattern
@@ -931,7 +932,7 @@ let glob_cpattern gs p =
| k, (v, Some t), _ as orig ->
if k = 'x' then glob_ssrterm gs ('(', (v, Some t), None) else
match t.CAst.v with
- | CNotation("( _ in _ )", ([t1; t2], [], [], [])) ->
+ | CNotation((InConstrEntrySomeLevel,"( _ in _ )"), ([t1; t2], [], [], [])) ->
(try match glob t1, glob t2 with
| (r1, None), (r2, None) -> encode k "In" [r1;r2]
| (r1, Some _), (r2, Some _) when isCVar t1 ->
@@ -939,11 +940,11 @@ let glob_cpattern gs p =
| (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2]
| _ -> CErrors.anomaly (str"where are we?.")
with _ when isCVar t1 -> encode k "In" [bind_in t1 t2])
- | CNotation("( _ in _ in _ )", ([t1; t2; t3], [], [], [])) ->
+ | CNotation((InConstrEntrySomeLevel,"( _ in _ in _ )"), ([t1; t2; t3], [], [], [])) ->
check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3]
- | CNotation("( _ as _ )", ([t1; t2], [], [], [])) ->
+ | CNotation((InConstrEntrySomeLevel,"( _ as _ )"), ([t1; t2], [], [], [])) ->
encode k "As" [fst (glob t1); fst (glob t2)]
- | CNotation("( _ as _ in _ )", ([t1; t2; t3], [], [], [])) ->
+ | CNotation((InConstrEntrySomeLevel,"( _ as _ in _ )"), ([t1; t2; t3], [], [], [])) ->
check_var t2; encode k "As" [fst (glob t1); bind_in t2 t3]
| _ -> glob_ssrterm gs orig
;;
@@ -980,27 +981,7 @@ let interp_rpattern s = function
| E_As_X_In_T(e,x,t) ->
E_As_X_In_T (interp_ssrterm s e,interp_ssrterm s x,interp_ssrterm s t)
-let interp_rpattern ist gl t = Tacmach.project gl, interp_rpattern ist t
-
-ARGUMENT EXTEND rpattern
- TYPED AS rpatternty
- PRINTED BY pr_rpattern
- INTERPRETED BY interp_rpattern
- GLOBALIZED BY glob_rpattern
- SUBSTITUTED BY subst_rpattern
- | [ lconstr(c) ] -> [ T (mk_lterm c None) ]
- | [ "in" lconstr(c) ] -> [ In_T (mk_lterm c None) ]
- | [ lconstr(x) "in" lconstr(c) ] ->
- [ X_In_T (mk_lterm x None, mk_lterm c None) ]
- | [ "in" lconstr(x) "in" lconstr(c) ] ->
- [ In_X_In_T (mk_lterm x None, mk_lterm c None) ]
- | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] ->
- [ E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None) ]
- | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] ->
- [ E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None) ]
-END
-
-
+let interp_rpattern0 ist gl t = Tacmach.project gl, interp_rpattern ist t
type cpattern = char * glob_constr_and_expr * Geninterp.interp_sign option
let tag_of_cpattern = pi1
@@ -1015,8 +996,10 @@ type pattern = Evd.evar_map * (constr, constr) ssrpattern
let id_of_cpattern (_, (c1, c2), _) =
let open CAst in
match DAst.get c1, c2 with
- | _, Some { v = CRef ({CAst.v=Ident x}, _) } -> Some x
- | _, Some { v = CAppExpl ((_, {CAst.v=Ident x}, _), []) } -> Some x
+ | _, Some { v = CRef (qid, _) } when qualid_is_ident qid ->
+ Some (qualid_basename qid)
+ | _, Some { v = CAppExpl ((_, qid, _), []) } when qualid_is_ident qid ->
+ Some (qualid_basename qid)
| GRef (VarRef x, _), None -> Some x
| _ -> None
let id_of_Cterm t = match id_of_cpattern t with
@@ -1042,52 +1025,9 @@ let interp_wit wit ist gl x =
let interp_open_constr ist gl gc =
interp_wit wit_open_constr ist gl gc
let pf_intern_term gl (_, c, ist) = glob_constr ist (pf_env gl) (project gl) c
-let pr_ssrterm _ _ _ = pr_term
-let input_ssrtermkind strm = match stream_nth 0 strm with
- | Tok.KEYWORD "(" -> '('
- | Tok.KEYWORD "@" -> '@'
- | _ -> ' '
-let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind
let interp_ssrterm ist gl t = Tacmach.project gl, interp_ssrterm ist t
-ARGUMENT EXTEND cpattern
- PRINTED BY pr_ssrterm
- INTERPRETED BY interp_ssrterm
- GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm
- RAW_PRINTED BY pr_ssrterm
- GLOB_PRINTED BY pr_ssrterm
-| [ "Qed" constr(c) ] -> [ mk_lterm c None ]
-END
-
-GEXTEND Gram
- GLOBAL: cpattern;
- cpattern: [[ k = ssrtermkind; c = constr ->
- let pattern = mk_term k c None in
- if loc_ofCG pattern <> Some !@loc && k = '('
- then mk_term 'x' c None
- else pattern ]];
-END
-
-ARGUMENT EXTEND lcpattern
- TYPED AS cpattern
- PRINTED BY pr_ssrterm
- INTERPRETED BY interp_ssrterm
- GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm
- RAW_PRINTED BY pr_ssrterm
- GLOB_PRINTED BY pr_ssrterm
-| [ "Qed" lconstr(c) ] -> [ mk_lterm c None ]
-END
-
-GEXTEND Gram
- GLOBAL: lcpattern;
- lcpattern: [[ k = ssrtermkind; c = lconstr ->
- let pattern = mk_term k c None in
- if loc_ofCG pattern <> Some !@loc && k = '('
- then mk_term 'x' c None
- else pattern ]];
-END
-
let interp_term gl = function
| (_, c, Some ist) ->
on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c)
@@ -1097,15 +1037,14 @@ let thin id sigma goal =
let ids = Id.Set.singleton id in
let env = Goal.V82.env sigma goal in
let cl = Goal.V82.concl sigma goal in
- let evdref = ref (Evd.clear_metas sigma) in
+ let sigma = Evd.clear_metas sigma in
let ans =
- try Some (Evarutil.clear_hyps_in_evi env evdref (Environ.named_context_val env) cl ids)
+ try Some (Evarutil.clear_hyps_in_evi env sigma (Environ.named_context_val env) cl ids)
with Evarutil.ClearDependencyError _ -> None
in
match ans with
| None -> sigma
- | Some (hyps, concl) ->
- let sigma = !evdref in
+ | Some (sigma, hyps, concl) ->
let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in
let sigma = Goal.V82.partial_solution_to sigma goal gl ev in
sigma
@@ -1257,7 +1196,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let fs sigma x = nf_evar sigma x in
let pop_evar sigma e p =
let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in
- let e_body = match e_body with Evar_defined c -> c
+ let e_body = match e_body with Evar_defined c -> EConstr.Unsafe.to_constr c
| _ -> errorstrm (str "Matching the pattern " ++ pr_constr_env env0 sigma0 p ++
str " did not instantiate ?" ++ int (Evar.repr e) ++ spc () ++
str "Does the variable bound by the \"in\" construct occur "++
@@ -1408,10 +1347,6 @@ let is_wildcard ((_, (l, r), _) : cpattern) : bool = match DAst.get l, r with
(* "ssrpattern" *)
-ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY pr_rpattern
-| [ rpattern(pat) ] -> [ pat ]
-END
-
let pr_rpattern = pr_pattern
let pf_merge_uc uc gl =
@@ -1420,6 +1355,9 @@ let pf_merge_uc uc gl =
let pf_unsafe_merge_uc uc gl =
re_sig (sig_it gl) (Evd.set_universe_context (project gl) uc)
+(** All the pattern types reuse the same dynamic toplevel tag *)
+let wit_ssrpatternarg = wit_rpatternty
+
let interp_rpattern = interp_rpattern ~wit_ssrpatternarg
let ssrpatterntac _ist arg gl =
@@ -1428,7 +1366,7 @@ let ssrpatterntac _ist arg gl =
let concl0 = pf_concl gl in
let concl0 = EConstr.Unsafe.to_constr concl0 in
let (t, uc), concl_x =
- fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in
+ fill_occ_pattern (pf_env gl) sigma0 concl0 pat noindex 1 in
let t = EConstr.of_constr t in
let concl_x = EConstr.of_constr concl_x in
let gl, tty = pf_type_of gl t in
@@ -1471,14 +1409,20 @@ let ssrinstancesof arg gl =
done; raise NoMatch
with NoMatch -> ppnl (str"END INSTANCES"); tclIDTAC gl
-TACTIC EXTEND ssrinstoftpat
-| [ "ssrinstancesoftpat" cpattern(arg) ] -> [ Proofview.V82.tactic (ssrinstancesof arg) ]
-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 ;;
+module Internal =
+struct
+ let wit_rpatternty = wit_rpatternty
+ let glob_rpattern = glob_rpattern
+ let subst_rpattern = subst_rpattern
+ let interp_rpattern = interp_rpattern0
+ let pr_rpattern = pr_rpattern
+ let mk_rpattern x = x
+ let mk_lterm = mk_lterm
+ let mk_term = mk_term
+ let glob_cpattern = glob_cpattern
+ let subst_ssrterm = subst_ssrterm
+ let interp_ssrterm = interp_ssrterm
+ let pr_ssrterm = pr_term
+end
(* vim: set filetype=ocaml foldmethod=marker: *)
diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli
new file mode 100644
index 00000000..9c79879d
--- /dev/null
+++ b/plugins/ssrmatching/ssrmatching.mli
@@ -0,0 +1,258 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(* (c) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+open Goal
+open Environ
+open Evd
+open Constr
+
+open Ltac_plugin
+open Tacexpr
+
+(** ******** Small Scale Reflection pattern matching facilities ************* *)
+
+(** Pattern parsing *)
+
+(** The type of context patterns, the patterns of the [set] tactic and
+ [:] tactical. These are patterns that identify a precise subterm. *)
+type cpattern
+val pr_cpattern : cpattern -> Pp.t
+
+(** The type of rewrite patterns, the patterns of the [rewrite] tactic.
+ These patterns also include patterns that identify all the subterms
+ of a context (i.e. "in" prefix) *)
+type rpattern
+val pr_rpattern : rpattern -> Pp.t
+
+(** Pattern interpretation and matching *)
+
+exception NoMatch
+exception NoProgress
+
+(** AST for [rpattern] (and consequently [cpattern]) *)
+type ('ident, 'term) ssrpattern =
+ | T of 'term
+ | In_T of 'term
+ | X_In_T of 'ident * 'term
+ | In_X_In_T of 'ident * 'term
+ | E_In_X_In_T of 'term * 'ident * 'term
+ | E_As_X_In_T of 'term * 'ident * 'term
+
+type pattern = evar_map * (constr, constr) ssrpattern
+val pp_pattern : pattern -> Pp.t
+
+(** Extracts the redex and applies to it the substitution part of the pattern.
+ @raise Anomaly if called on [In_T] or [In_X_In_T] *)
+val redex_of_pattern :
+ ?resolve_typeclasses:bool -> env -> pattern ->
+ constr Evd.in_evar_universe_context
+
+(** [interp_rpattern ise gl rpat] "internalizes" and "interprets" [rpat]
+ in the current [Ltac] interpretation signature [ise] and tactic input [gl]*)
+val interp_rpattern :
+ goal sigma ->
+ rpattern ->
+ pattern
+
+(** [interp_cpattern ise gl cpat ty] "internalizes" and "interprets" [cpat]
+ in the current [Ltac] interpretation signature [ise] and tactic input [gl].
+ [ty] is an optional type for the redex of [cpat] *)
+val interp_cpattern :
+ goal sigma ->
+ cpattern -> (glob_constr_and_expr * Geninterp.interp_sign) option ->
+ pattern
+
+(** The set of occurrences to be matched. The boolean is set to true
+ * to signal the complement of this set (i.e. \{-1 3\}) *)
+type occ = (bool * int list) option
+
+(** [subst e p t i]. [i] is the number of binders
+ traversed so far, [p] the term from the pattern, [t] the matched one *)
+type subst = env -> constr -> constr -> int -> constr
+
+(** [eval_pattern b env sigma t pat occ subst] maps [t] calling [subst] on every
+ [occ] occurrence of [pat]. The [int] argument is the number of
+ binders traversed. If [pat] is [None] then then subst is called on [t].
+ [t] must live in [env] and [sigma], [pat] must have been interpreted in
+ (an extension of) [sigma].
+ @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false])
+ @return [t] where all [occ] occurrences of [pat] have been mapped using
+ [subst] *)
+val eval_pattern :
+ ?raise_NoMatch:bool ->
+ env -> evar_map -> constr ->
+ pattern option -> occ -> subst ->
+ constr
+
+(** [fill_occ_pattern b env sigma t pat occ h] is a simplified version of
+ [eval_pattern].
+ It replaces all [occ] occurrences of [pat] in [t] with Rel [h].
+ [t] must live in [env] and [sigma], [pat] must have been interpreted in
+ (an extension of) [sigma].
+ @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false])
+ @return the instance of the redex of [pat] that was matched and [t]
+ transformed as described above. *)
+val fill_occ_pattern :
+ ?raise_NoMatch:bool ->
+ env -> evar_map -> constr ->
+ pattern -> occ -> int ->
+ constr Evd.in_evar_universe_context * constr
+
+(** *************************** Low level APIs ****************************** *)
+
+(* The primitive matching facility. It matches of a term with holes, like
+ the T pattern above, and calls a continuation on its occurrences. *)
+
+type ssrdir = L2R | R2L
+val pr_dir_side : ssrdir -> Pp.t
+
+(** a pattern for a term with wildcards *)
+type tpattern
+
+(** [mk_tpattern env sigma0 sigma_p ok p_origin dir t] compiles a term [t]
+ living in [env] [sigma] (an extension of [sigma0]) intro a [tpattern].
+ The [tpattern] can hold a (proof) term [p] and a diction [dir]. The [ok]
+ callback is used to filter occurrences.
+ @return the compiled [tpattern] and its [evar_map]
+ @raise UserEerror is the pattern is a wildcard *)
+val mk_tpattern :
+ ?p_origin:ssrdir * constr ->
+ env -> evar_map ->
+ evar_map * constr ->
+ (constr -> evar_map -> bool) ->
+ ssrdir -> constr ->
+ evar_map * tpattern
+
+(** [findP env t i k] is a stateful function that finds the next occurrence
+ of a tpattern and calls the callback [k] to map the subterm matched.
+ The [int] argument passed to [k] is the number of binders traversed so far
+ plus the initial value [i].
+ @return [t] where the subterms identified by the selected occurrences of
+ the patter have been mapped using [k]
+ @raise NoMatch if the raise_NoMatch flag given to [mk_tpattern_matcher] is
+ [true] and if the pattern did not match
+ @raise UserEerror if the raise_NoMatch flag given to [mk_tpattern_matcher] is
+ [false] and if the pattern did not match *)
+type find_P =
+ env -> constr -> int -> k:subst -> constr
+
+(** [conclude ()] asserts that all mentioned ocurrences have been visited.
+ @return the instance of the pattern, the evarmap after the pattern
+ instantiation, the proof term and the ssrdit stored in the tpattern
+ @raise UserEerror if too many occurrences were specified *)
+type conclude =
+ unit -> constr * ssrdir * (evar_map * UState.t * constr)
+
+(** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair
+ a function [find_P] and [conclude] with the behaviour explained above.
+ The flag [b] (default [false]) changes the error reporting behaviour
+ of [find_P] if none of the [tpattern] matches. The argument [o] can
+ be passed to tune the [UserError] eventually raised (useful if the
+ pattern is coming from the LHS/RHS of an equation) *)
+val mk_tpattern_matcher :
+ ?all_instances:bool ->
+ ?raise_NoMatch:bool ->
+ ?upats_origin:ssrdir * constr ->
+ evar_map -> occ -> evar_map * tpattern list ->
+ find_P * conclude
+
+(** Example of [mk_tpattern_matcher] to implement
+ [rewrite \{occ\}\[in t\]rules].
+ It first matches "in t" (called [pat]), then in all matched subterms
+ it matches the LHS of the rules using [find_R].
+ [concl0] is the initial goal, [concl] will be the goal where some terms
+ are replaced by a De Bruijn index. The [rw_progress] extra check
+ selects only occurrences that are not rewritten to themselves (e.g.
+ an occurrence "x + x" rewritten with the commutativity law of addition
+ is skipped) {[
+ let find_R, conclude = match pat with
+ | Some (_, In_T _) ->
+ let aux (sigma, pats) (d, r, lhs, rhs) =
+ let sigma, pat =
+ mk_tpattern env0 sigma0 (sigma, r) (rw_progress rhs) d lhs in
+ sigma, pats @ [pat] in
+ let rpats = List.fold_left aux (r_sigma, []) rules in
+ let find_R, end_R = mk_tpattern_matcher sigma0 occ rpats in
+ find_R ~k:(fun _ _ h -> mkRel h),
+ fun cl -> let rdx, d, r = end_R () in (d,r),rdx
+ | _ -> ... in
+ let concl = eval_pattern env0 sigma0 concl0 pat occ find_R in
+ let (d, r), rdx = conclude concl in ]} *)
+
+(* convenience shortcut: [pf_fill_occ_term gl occ (sigma,t)] returns
+ * the conclusion of [gl] where [occ] occurrences of [t] have been replaced
+ * by [Rel 1] and the instance of [t] *)
+val pf_fill_occ_term : goal sigma -> occ -> evar_map * EConstr.t -> EConstr.t * EConstr.t
+
+(* It may be handy to inject a simple term into the first form of cpattern *)
+val cpattern_of_term : char * glob_constr_and_expr -> Geninterp.interp_sign -> cpattern
+
+(** Helpers to make stateful closures. Example: a [find_P] function may be
+ called many times, but the pattern instantiation phase is performed only the
+ first time. The corresponding [conclude] has to return the instantiated
+ pattern redex. Since it is up to [find_P] to raise [NoMatch] if the pattern
+ has no instance, [conclude] considers it an anomaly if the pattern did
+ not match *)
+
+(** [do_once r f] calls [f] and updates the ref only once *)
+val do_once : 'a option ref -> (unit -> 'a) -> unit
+(** [assert_done r] return the content of r. @raise Anomaly is r is [None] *)
+val assert_done : 'a option ref -> 'a
+
+(** Very low level APIs.
+ these are calls to evarconv's [the_conv_x] followed by
+ [solve_unif_constraints_with_heuristics] and [resolve_typeclasses].
+ In case of failure they raise [NoMatch] *)
+
+val unify_HO : env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map
+val pf_unify_HO : goal sigma -> EConstr.constr -> EConstr.constr -> goal sigma
+
+(** Some more low level functions needed to implement the full SSR language
+ on top of the former APIs *)
+val tag_of_cpattern : cpattern -> char
+val loc_of_cpattern : cpattern -> Loc.t option
+val id_of_pattern : pattern -> Names.Id.t option
+val is_wildcard : cpattern -> bool
+val cpattern_of_id : Names.Id.t -> cpattern
+val pr_constr_pat : constr -> Pp.t
+val pf_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma
+val pf_unsafe_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma
+
+(* One can also "Set SsrMatchingDebug" from a .v *)
+val debug : bool -> unit
+
+(* One should delimit a snippet with "Set SsrMatchingProfiling" and
+ * "Unset SsrMatchingProfiling" to get timings *)
+val profile : bool -> unit
+
+val ssrinstancesof : cpattern -> Tacmach.tactic
+
+(** Functions used for grammar extensions. Do not use. *)
+
+module Internal :
+sig
+ val wit_rpatternty : (rpattern, rpattern, rpattern) Genarg.genarg_type
+ val glob_rpattern : Genintern.glob_sign -> rpattern -> rpattern
+ val subst_rpattern : Mod_subst.substitution -> rpattern -> rpattern
+ val interp_rpattern : Geninterp.interp_sign -> Proof_type.goal Evd.sigma -> rpattern -> Evd.evar_map * rpattern
+ val pr_rpattern : rpattern -> Pp.t
+ val mk_rpattern : (cpattern, cpattern) ssrpattern -> rpattern
+ val mk_lterm : Constrexpr.constr_expr -> Geninterp.interp_sign option -> cpattern
+ val mk_term : char -> Constrexpr.constr_expr -> Geninterp.interp_sign option -> cpattern
+
+ val glob_cpattern : Genintern.glob_sign -> cpattern -> cpattern
+ val subst_ssrterm : Mod_subst.substitution -> cpattern -> cpattern
+ val interp_ssrterm : Geninterp.interp_sign -> Proof_type.goal Evd.sigma -> cpattern -> Evd.evar_map * cpattern
+ val pr_ssrterm : cpattern -> Pp.t
+end
+
+(* eof *)
diff --git a/plugins/ssrmatching/ssrmatching.v b/plugins/ssrmatching/ssrmatching.v
new file mode 100644
index 00000000..fee09cc9
--- /dev/null
+++ b/plugins/ssrmatching/ssrmatching.v
@@ -0,0 +1,36 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(* (c) Copyright 2006-2015 Microsoft Corporation and Inria. *)
+
+Declare ML Module "ssrmatching_plugin".
+
+Module SsrMatchingSyntax.
+
+(* Reserve the notation for rewrite patterns so that the user is not allowed *)
+(* to declare it at a different level. *)
+Reserved Notation "( a 'in' b )" (at level 0).
+Reserved Notation "( a 'as' b )" (at level 0).
+Reserved Notation "( a 'in' b 'in' c )" (at level 0).
+Reserved Notation "( a 'as' b 'in' c )" (at level 0).
+
+(* Notation to define shortcuts for the "X in t" part of a pattern. *)
+Notation "( X 'in' t )" := (_ : fun X => t) : ssrpatternscope.
+Delimit Scope ssrpatternscope with pattern.
+
+(* Some shortcuts for recurrent "X in t" parts. *)
+Notation RHS := (X in _ = X)%pattern.
+Notation LHS := (X in X = _)%pattern.
+
+End SsrMatchingSyntax.
+
+Export SsrMatchingSyntax.
+
+Tactic Notation "ssrpattern" ssrpatternarg(p) := ssrpattern p .
diff --git a/plugins/ssrmatching/ssrmatching_plugin.mlpack b/plugins/ssrmatching/ssrmatching_plugin.mlpack
index 5fb1f156..02c75f14 100644
--- a/plugins/ssrmatching/ssrmatching_plugin.mlpack
+++ b/plugins/ssrmatching/ssrmatching_plugin.mlpack
@@ -1 +1,2 @@
Ssrmatching
+G_ssrmatching