diff options
| author | Hugo Herbelin | 2017-08-12 09:15:40 +0200 |
|---|---|---|
| committer | Hugo Herbelin | 2018-02-20 10:03:06 +0100 |
| commit | edd0d22429354a5f2c703a8c7bd1f775e2f97d9e (patch) | |
| tree | 865fd16d40f5641cac233f951f951a9a4502159f /interp/notation.ml | |
| parent | 398358618bb3eabfe822b79c669703c1c33b67e6 (diff) | |
Adding support for parsing subterms of a notation as "pattern".
This allows in particular to define notations with 'pat style binders.
E.g.:
A non-trivial change in this commit is storing binders and patterns
separately from terms.
This is not strictly necessary but has some advantages.
However, it is relatively common to have binders also used as terms,
or binders parsed as terms. Thus, it is already relatively common to
embed binders into terms (see e.g. notation for ETA in output test
Notations3.v) or to coerce terms to idents (see e.g. the notation for
{x|P} where x is parsed as a constr).
So, it is as simple to always store idents (and eventually patterns)
as terms.
Diffstat (limited to 'interp/notation.ml')
| -rw-r--r-- | interp/notation.ml | 29 |
1 files changed, 21 insertions, 8 deletions
diff --git a/interp/notation.ml b/interp/notation.ml index 31f16e1a90..e6186e08c2 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -82,18 +82,31 @@ let parenRelation_eq t1 t2 = match t1, t2 with | Prec l1, Prec l2 -> Int.equal l1 l2 | _ -> false -let notation_var_internalization_type_eq v1 v2 = match v1, v2 with -| NtnInternTypeConstr, NtnInternTypeConstr -> true -| NtnInternTypeBinder, NtnInternTypeBinder -> true -| NtnInternTypeIdent, NtnInternTypeIdent -> true -| (NtnInternTypeConstr | NtnInternTypeBinder | NtnInternTypeIdent), _ -> 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 v1 v2 = match v1, v2 with +| ETName, ETName -> true +| ETReference, ETReference -> true +| ETBigint, ETBigint -> true +| ETBinder b1, ETBinder b2 -> b1 == b2 +| ETConstr (n1,pp1), ETConstr (n2,pp2) -> production_level_eq n1 n2 && production_position_eq pp1 pp2 +| ETPattern n1, ETPattern n2 -> 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 _), _ -> false let level_eq (l1, t1, u1) (l2, t2, u2) = let tolerability_eq (i1, r1) (i2, r2) = Int.equal i1 i2 && parenRelation_eq r1 r2 in Int.equal l1 l2 && List.equal tolerability_eq t1 t2 - && List.equal notation_var_internalization_type_eq u1 u2 + && List.equal constr_entry_key_eq u1 u2 let declare_scope scope = try let _ = String.Map.find scope !scope_map in () @@ -611,10 +624,10 @@ let pair_eq f g (x1, y1) (x2, y2) = f x1 x2 && g y1 y2 let ntpe_eq t1 t2 = match t1, t2 with | NtnTypeConstr, NtnTypeConstr -> true -| NtnTypeOnlyBinder, NtnTypeOnlyBinder -> true +| NtnTypeBinder b1, NtnTypeBinder b2 -> b1 = (b2:bool) | NtnTypeConstrList, NtnTypeConstrList -> true | NtnTypeBinderList, NtnTypeBinderList -> true -| (NtnTypeConstr | NtnTypeOnlyBinder | NtnTypeConstrList | NtnTypeBinderList), _ -> false +| (NtnTypeConstr | NtnTypeBinder _ | NtnTypeConstrList | NtnTypeBinderList), _ -> false let var_attributes_eq (_, (sc1, tp1)) (_, (sc2, tp2)) = pair_eq (Option.equal String.equal) (List.equal String.equal) sc1 sc2 && |
