From a3ab8b07b912afd1b883ed60bd532b5a29bccd8f Mon Sep 17 00:00:00 2001 From: letouzey Date: Tue, 29 May 2012 11:09:01 +0000 Subject: Stuff about notation_constr (ex-aconstr) now in notation_ops.ml git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@15381 85f007b7-540e-0410-9357-904b9bb8a0f7 --- interp/notation_ops.mli | 63 +++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 63 insertions(+) create mode 100644 interp/notation_ops.mli (limited to 'interp/notation_ops.mli') diff --git a/interp/notation_ops.mli b/interp/notation_ops.mli new file mode 100644 index 000000000..9dfc656a0 --- /dev/null +++ b/interp/notation_ops.mli @@ -0,0 +1,63 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + (identifier * identifier) list -> glob_constr -> notation_constr + +(** Name of the special identifier used to encode recursive notations *) +val ldots_var : identifier + +(** Equality of [glob_constr] (warning: only partially implemented) *) +val eq_glob_constr : glob_constr -> glob_constr -> bool + +(** Re-interpret a notation as a [glob_constr], taking care of binders *) + +val glob_constr_of_notation_constr_with_binders : Pp.loc -> + ('a -> name -> 'a * name) -> + ('a -> notation_constr -> glob_constr) -> + 'a -> notation_constr -> glob_constr + +val glob_constr_of_notation_constr : Pp.loc -> notation_constr -> glob_constr + +(** [match_notation_constr] matches a [glob_constr] against a notation + interpretation; raise [No_match] if the matching fails *) + +exception No_match + +val match_notation_constr : bool -> glob_constr -> interpretation -> + (glob_constr * subscopes) list * (glob_constr list * subscopes) list * + (glob_decl list * subscopes) list + +val match_notation_constr_cases_pattern : + cases_pattern -> interpretation -> + ((cases_pattern * subscopes) list * (cases_pattern list * subscopes) list) * + (cases_pattern list) + +val match_notation_constr_ind_pattern : + inductive -> cases_pattern list -> interpretation -> + ((cases_pattern * subscopes) list * (cases_pattern list * subscopes) list) * + (cases_pattern list) + +(** Substitution of kernel names in interpretation data *) + +val subst_interpretation : + Mod_subst.substitution -> interpretation -> interpretation + +val add_patterns_for_params : inductive -> cases_pattern list -> cases_pattern list -- cgit v1.2.3