From 63b530234e0b19323a50c52434a7439518565c81 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Thu, 24 May 2018 03:15:17 +0200 Subject: [notations] Split interpretation and parsing of notations Previously to this patch, `Notation_term` contained information about both parsing and notation interpretation. We split notation grammar to a file `parsing/notation_gram` as to make `interp/` not to depend on some parsing structures such as entries. --- interp/constrintern.ml | 4 +- interp/interp.mllib | 1 - interp/notation.ml | 99 ++----------------------------------------------- interp/notation.mli | 23 ------------ interp/notation_ops.ml | 2 +- interp/notation_term.ml | 37 +++--------------- interp/ppextend.ml | 43 --------------------- interp/ppextend.mli | 36 ------------------ 8 files changed, 12 insertions(+), 233 deletions(-) delete mode 100644 interp/ppextend.ml delete mode 100644 interp/ppextend.mli (limited to 'interp') diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 3a88284e4..45c0e9c42 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -820,11 +820,11 @@ let split_by_type ids subst = | NtnTypeConstr -> let terms,terms' = bind id scl terms terms' in (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists') - | NtnTypeBinder NtnBinderParsedAsConstr (Extend.AsIdentOrPattern | Extend.AsStrictPattern) -> + | NtnTypeBinder NtnBinderParsedAsConstr (AsIdentOrPattern | AsStrictPattern) -> let a,terms = match terms with a::terms -> a,terms | _ -> assert false in let binders' = Id.Map.add id (coerce_to_cases_pattern_expr a,(false,scl)) binders' in (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists') - | NtnTypeBinder NtnBinderParsedAsConstr Extend.AsIdent -> + | NtnTypeBinder NtnBinderParsedAsConstr AsIdent -> let a,terms = match terms with a::terms -> a,terms | _ -> assert false in let binders' = Id.Map.add id (cases_pattern_of_name (coerce_to_name a),(true,scl)) binders' in (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists') diff --git a/interp/interp.mllib b/interp/interp.mllib index 8ed2b9fad..3668455ae 100644 --- a/interp/interp.mllib +++ b/interp/interp.mllib @@ -10,7 +10,6 @@ Notation Syntax_def Smartlocate Constrexpr_ops -Ppextend Dumpglob Reserve Impargs diff --git a/interp/notation.ml b/interp/notation.ml index 192c477be..05fcd0e7f 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -20,7 +20,6 @@ open Constrexpr open Notation_term open Glob_term open Glob_ops -open Ppextend open Context.Named.Declaration (*i*) @@ -56,9 +55,6 @@ type scope = { delimiters: delimiters option } -(* Uninterpreted notation map: notation -> level * DirPath.t *) -let notation_level_map = ref String.Map.empty - (* Scopes table: scope_name -> symbol_interpretation *) let scope_map = ref String.Map.empty @@ -75,44 +71,6 @@ let default_scope = "" (* empty name, not available from outside *) let init_scope_map () = scope_map := String.Map.add default_scope empty_scope !scope_map -(**********************************************************************) -(* Operations on scopes *) - -let parenRelation_eq t1 t2 = match t1, t2 with -| L, L | E, E | Any, Any -> true -| Prec l1, Prec l2 -> Int.equal l1 l2 -| _ -> false - -open Extend - -let production_level_eq l1 l2 = true (* (l1 = l2) *) - -let production_position_eq pp1 pp2 = true (* pp1 = pp2 *) (* match pp1, pp2 with -| NextLevel, NextLevel -> true -| NumLevel n1, NumLevel n2 -> Int.equal n1 n2 -| (NextLevel | NumLevel _), _ -> false *) - -let constr_entry_key_eq eq v1 v2 = match v1, v2 with -| ETName, ETName -> true -| ETReference, ETReference -> true -| ETBigint, ETBigint -> true -| ETBinder b1, ETBinder b2 -> b1 == b2 -| ETConstr lev1, ETConstr lev2 -> eq lev1 lev2 -| ETConstrAsBinder (bk1,lev1), ETConstrAsBinder (bk2,lev2) -> eq lev1 lev2 && bk1 = bk2 -| ETPattern (b1,n1), ETPattern (b2,n2) -> b1 = b2 && Option.equal Int.equal n1 n2 -| ETOther (s1,s1'), ETOther (s2,s2') -> String.equal s1 s2 && String.equal s1' s2' -| (ETName | ETReference | ETBigint | ETBinder _ | ETConstr _ | ETPattern _ | ETOther _ | ETConstrAsBinder _), _ -> false - -let level_eq_gen strict (l1, t1, u1) (l2, t2, u2) = - let tolerability_eq (i1, r1) (i2, r2) = Int.equal i1 i2 && parenRelation_eq r1 r2 in - let prod_eq (l1,pp1) (l2,pp2) = - if strict then production_level_eq l1 l2 && production_position_eq pp1 pp2 - else production_level_eq l1 l2 in - Int.equal l1 l2 && List.equal tolerability_eq t1 t2 - && List.equal (constr_entry_key_eq prod_eq) u1 u2 - -let level_eq = level_eq_gen false - let declare_scope scope = try let _ = String.Map.find scope !scope_map in () with Not_found -> @@ -427,18 +385,6 @@ let rec find_without_delimiters find (ntn_scope,ntn) = function (* Can we switch to [scope]? Yes if it has defined delimiters *) find_with_delimiters ntn_scope -(* Uninterpreted notation levels *) - -let declare_notation_level ?(onlyprint=false) ntn level = - if String.Map.mem ntn !notation_level_map then - anomaly (str "Notation " ++ str ntn ++ str " is already assigned a level."); - notation_level_map := String.Map.add ntn (level,onlyprint) !notation_level_map - -let level_of_notation ?(onlyprint=false) ntn = - let (level,onlyprint') = String.Map.find ntn !notation_level_map in - if onlyprint' && not onlyprint then raise Not_found; - level - (* The mapping between notations and their interpretation *) let warn_notation_overridden = @@ -1112,64 +1058,25 @@ let pr_visibility prglob = function | Some scope -> pr_scope_stack prglob (push_scope scope !scope_stack) | None -> pr_scope_stack prglob !scope_stack -(**********************************************************************) -(* Mapping notations to concrete syntax *) - -type unparsing_rule = unparsing list * precedence -type extra_unparsing_rules = (string * string) list -(* Concrete syntax for symbolic-extension table *) -let notation_rules = - ref (String.Map.empty : (unparsing_rule * extra_unparsing_rules * notation_grammar) String.Map.t) - -let declare_notation_rule ntn ~extra unpl gram = - notation_rules := String.Map.add ntn (unpl,extra,gram) !notation_rules - -let find_notation_printing_rule ntn = - try pi1 (String.Map.find ntn !notation_rules) - with Not_found -> anomaly (str "No printing rule found for " ++ str ntn ++ str ".") -let find_notation_extra_printing_rules ntn = - try pi2 (String.Map.find ntn !notation_rules) - with Not_found -> [] -let find_notation_parsing_rules ntn = - try pi3 (String.Map.find ntn !notation_rules) - with Not_found -> anomaly (str "No parsing rule found for " ++ str ntn ++ str ".") - -let get_defined_notations () = - String.Set.elements @@ String.Map.domain !notation_rules - -let add_notation_extra_printing_rule ntn k v = - try - notation_rules := - let p, pp, gr = String.Map.find ntn !notation_rules in - String.Map.add ntn (p, (k,v) :: pp, gr) !notation_rules - with Not_found -> - user_err ~hdr:"add_notation_extra_printing_rule" - (str "No such Notation.") - (**********************************************************************) (* Synchronisation with reset *) let freeze _ = - (!scope_map, !notation_level_map, !scope_stack, !arguments_scope, - !delimiters_map, !notations_key_table, !notation_rules, - !scope_class_map) + (!scope_map, !scope_stack, !arguments_scope, + !delimiters_map, !notations_key_table, !scope_class_map) -let unfreeze (scm,nlm,scs,asc,dlm,fkm,pprules,clsc) = +let unfreeze (scm,scs,asc,dlm,fkm,clsc) = scope_map := scm; - notation_level_map := nlm; scope_stack := scs; delimiters_map := dlm; arguments_scope := asc; notations_key_table := fkm; - notation_rules := pprules; scope_class_map := clsc let init () = init_scope_map (); - notation_level_map := String.Map.empty; delimiters_map := String.Map.empty; notations_key_table := KeyMap.empty; - notation_rules := String.Map.empty; scope_class_map := initial_scope_class_map let _ = diff --git a/interp/notation.mli b/interp/notation.mli index ccc67fe49..b177b7f1e 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -14,7 +14,6 @@ open Libnames open Constrexpr open Glob_term open Notation_term -open Ppextend (** Notations *) @@ -32,8 +31,6 @@ val declare_scope : scope_name -> unit val current_scopes : unit -> scopes -val level_eq : level -> level -> bool - (** Check where a scope is opened or not in a scope list, or in * the current opened scopes *) val scope_is_open_in_scopes : scope_name -> scopes -> bool @@ -135,11 +132,6 @@ val uninterp_ind_pattern_notations : inductive -> notation_rule list val availability_of_notation : scope_name option * notation -> local_scopes -> (scope_name option * delimiters option) option -(** {6 Declare and test the level of a (possibly uninterpreted) notation } *) - -val declare_notation_level : ?onlyprint:bool -> notation -> level -> unit -val level_of_notation : ?onlyprint:bool -> notation -> level (** raise [Not_found] if no level or not respecting onlyprint *) - (** {6 Miscellaneous} *) val interp_notation_as_global_reference : ?loc:Loc.t -> (GlobRef.t -> bool) -> @@ -200,21 +192,6 @@ val locate_notation : (glob_constr -> Pp.t) -> notation -> val pr_visibility: (glob_constr -> Pp.t) -> scope_name option -> Pp.t -(** {6 Printing rules for notations} *) - -(** Declare and look for the printing rule for symbolic notations *) -type unparsing_rule = unparsing list * precedence -type extra_unparsing_rules = (string * string) list -val declare_notation_rule : - notation -> extra:extra_unparsing_rules -> unparsing_rule -> notation_grammar -> unit -val find_notation_printing_rule : notation -> unparsing_rule -val find_notation_extra_printing_rules : notation -> extra_unparsing_rules -val find_notation_parsing_rules : notation -> notation_grammar -val add_notation_extra_printing_rule : notation -> string -> string -> unit - -(** Returns notations with defined parsing/printing rules *) -val get_defined_notations : unit -> notation list - (** Rem: printing rules for primitive token are canonical *) val with_notation_protection : ('a -> 'b) -> 'a -> 'b diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index b0480aa70..f208b23fb 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -687,7 +687,7 @@ let is_onlybinding_meta id metas = let is_onlybinding_pattern_like_meta isvar id metas = try match Id.List.assoc id metas with | _,NtnTypeBinder (NtnBinderParsedAsConstr - (Extend.AsIdentOrPattern | Extend.AsStrictPattern)) -> true + (AsIdentOrPattern | AsStrictPattern)) -> true | _,NtnTypeBinder (NtnParsedAsPattern strict) -> not (strict && isvar) | _ -> false with Not_found -> false diff --git a/interp/notation_term.ml b/interp/notation_term.ml index 1a46746cc..52a6354a0 100644 --- a/interp/notation_term.ml +++ b/interp/notation_term.ml @@ -62,6 +62,11 @@ type subscopes = tmp_scope_name option * scope_name list (** Type of the meta-variables of an notation_constr: in a recursive pattern x..y, x carries the sequence of objects bound to the list x..y *) +type constr_as_binder_kind = + | AsIdent + | AsIdentOrPattern + | AsStrictPattern + type notation_binder_source = (* This accepts only pattern *) (* NtnParsedAsPattern true means only strict pattern (no single variable) at printing *) @@ -69,7 +74,7 @@ type notation_binder_source = (* This accepts only ident *) | NtnParsedAsIdent (* This accepts ident, or pattern, or both *) - | NtnBinderParsedAsConstr of Extend.constr_as_binder_kind + | NtnBinderParsedAsConstr of constr_as_binder_kind type notation_var_instance_type = | NtnTypeConstr | NtnTypeBinder of notation_binder_source | NtnTypeConstrList | NtnTypeBinderList @@ -91,33 +96,3 @@ type notation_interp_env = { ninterp_var_type : notation_var_internalization_type Id.Map.t; ninterp_rec_vars : Id.t Id.Map.t; } - -type grammar_constr_prod_item = - | GramConstrTerminal of Tok.t - | GramConstrNonTerminal of Extend.constr_prod_entry_key * Id.t option - | GramConstrListMark of int * bool * int - (* tells action rule to make a list of the n previous parsed items; - concat with last parsed list when true; additionally release - the p last items as if they were parsed autonomously *) - -(** Dealing with precedences *) - -type precedence = int -type parenRelation = L | E | Any | Prec of precedence -type tolerability = precedence * parenRelation - -type level = precedence * tolerability list * Extend.constr_entry_key list - -(** Grammar rules for a notation *) - -type one_notation_grammar = { - notgram_level : level; - notgram_assoc : Extend.gram_assoc option; - notgram_notation : Constrexpr.notation; - notgram_prods : grammar_constr_prod_item list list; -} - -type notation_grammar = { - notgram_onlyprinting : bool; - notgram_rules : one_notation_grammar list -} diff --git a/interp/ppextend.ml b/interp/ppextend.ml deleted file mode 100644 index c75d9e12f..000000000 --- a/interp/ppextend.ml +++ /dev/null @@ -1,43 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* h n - | PpHOVB n -> hov n - | PpHVB n -> hv n - | PpVB n -> v n - -let ppcmd_of_cut = function - | PpFnl -> fnl () - | PpBrk(n1,n2) -> brk(n1,n2) - -type unparsing = - | UnpMetaVar of int * parenRelation - | UnpBinderMetaVar of int * parenRelation - | UnpListMetaVar of int * parenRelation * unparsing list - | UnpBinderListMetaVar of int * bool * unparsing list - | UnpTerminal of string - | UnpBox of ppbox * unparsing Loc.located list - | UnpCut of ppcut diff --git a/interp/ppextend.mli b/interp/ppextend.mli deleted file mode 100644 index c81058e72..000000000 --- a/interp/ppextend.mli +++ /dev/null @@ -1,36 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* Pp.t -> Pp.t - -val ppcmd_of_cut : ppcut -> Pp.t - -type unparsing = - | UnpMetaVar of int * parenRelation - | UnpBinderMetaVar of int * parenRelation - | UnpListMetaVar of int * parenRelation * unparsing list - | UnpBinderListMetaVar of int * bool * unparsing list - | UnpTerminal of string - | UnpBox of ppbox * unparsing Loc.located list - | UnpCut of ppcut -- cgit v1.2.3