diff options
| author | Hugo Herbelin | 2018-03-08 18:22:50 +0100 |
|---|---|---|
| committer | Hugo Herbelin | 2018-12-04 11:50:02 +0100 |
| commit | 06fc6caa49b67526cf3521d1b5885aae6710358b (patch) | |
| tree | c893a474c67d9f569e1e19c7ccaedb1a02ada4f5 | |
| parent | 74dfaaa8555f53bfc75216205823a8020e80b6a1 (diff) | |
Addressing issues with PR#873: performance and use of abbreviation for printing.
We do a couple of changes:
- Splitting notation keys into more categories to make table smaller.
This should (a priori) make printing faster (see #6416).
- Abbreviations are treated for printing like single notations: they
are pushed to the scope stack, so that in a situation such as
Open Scope foo_scope.
Notation foo := term.
Open Scope bar_scope.
one looks for notations first in scope bar_scope, then try to use
foo, they try for notations in scope foo_scope.
- We seize the opportunity of this commit to simplify
availability_of_notation which is now integrated to
uninterp_notation and which does not have to be called explicitly
anymore.
| -rw-r--r-- | dev/doc/changes.md | 6 | ||||
| -rw-r--r-- | interp/constrextern.ml | 33 | ||||
| -rw-r--r-- | interp/notation.ml | 205 | ||||
| -rw-r--r-- | interp/notation.mli | 14 | ||||
| -rw-r--r-- | test-suite/output/Inductive.out | 2 | ||||
| -rw-r--r-- | test-suite/output/Notations.out | 2 | ||||
| -rw-r--r-- | test-suite/output/Notations2.out | 2 |
7 files changed, 168 insertions, 96 deletions
diff --git a/dev/doc/changes.md b/dev/doc/changes.md index acb0d80c18..c0f15f02a5 100644 --- a/dev/doc/changes.md +++ b/dev/doc/changes.md @@ -152,6 +152,12 @@ Termops - Internal printing functions have been placed under the `Termops.Internal` namespace. +Notations: + +- Notation.availability_of_notation is not anymore needed: if a + delimiter is needed, it is provided by Notation.uninterp_notation + which fails in case the notation is not available. + ### Unit testing The test suite now allows writing unit tests against OCaml code in the Coq diff --git a/interp/constrextern.ml b/interp/constrextern.ml index faff0d71e0..f1d8d858a1 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -110,13 +110,13 @@ let deactivate_notation nr = (* shouldn't we check wether it is well defined? *) inactive_notations_table := IRuleSet.add nr !inactive_notations_table | NotationRule (scopt, ntn) -> - match availability_of_notation (scopt, ntn) (scopt, []) with - | None -> user_err ~hdr:"Notation" + if not (exists_notation_interpretation_in_scope scopt ntn) then + user_err ~hdr:"Notation" (pr_notation ntn ++ spc () ++ str "does not exist" ++ (match scopt with | None -> spc () ++ str "in the empty scope." | Some _ -> show_scope scopt ++ str ".")) - | Some _ -> + else if IRuleSet.mem nr !inactive_notations_table then Feedback.msg_warning (str "Notation" ++ spc () ++ pr_notation ntn ++ spc () @@ -443,16 +443,12 @@ let rec extern_cases_pattern_in_scope (custom,scopes as allscopes) vars pat = and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) (custom, (tmp_scope, scopes) as allscopes) vars = function - | NotationRule (sc,ntn) -> + | NotationRule (sc,ntn),key -> begin match availability_of_entry_coercion custom (fst ntn) with | None -> raise No_match | Some coercion -> - match availability_of_notation (sc,ntn) (tmp_scope,scopes) with - (* Uninterpretation is not allowed in current context *) - | None -> raise No_match - (* Uninterpretation is allowed in current context *) - | Some (scopt,key) -> + let scopt = match key with Some _ -> sc | _ -> None in let scopes' = Option.List.cons scopt scopes in let l = List.map (fun (c,(subentry,(scopt,scl))) -> @@ -474,7 +470,8 @@ and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) (insert_pat_delimiters ?loc (make_pat_notation ?loc ntn (l,ll) l2') key) end - | SynDefRule kn -> + | SynDefRule kn,key -> + assert (key = None); match availability_of_entry_coercion custom InConstrEntrySomeLevel with | None -> raise No_match | Some coercion -> @@ -494,7 +491,7 @@ and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) insert_pat_coercion ?loc coercion (mkPat ?loc qid (List.rev_append l1 l2')) and extern_notation_pattern allscopes vars t = function | [] -> raise No_match - | (keyrule,pat,n as _rule)::rules -> + | (keyrule,pat,n as _rule,key)::rules -> try if is_inactive_rule keyrule then raise No_match; let loc = t.loc in @@ -502,7 +499,7 @@ and extern_notation_pattern allscopes vars t = function | PatCstr (cstr,args,na) -> let t = if na = Anonymous then t else DAst.make ?loc (PatCstr (cstr,args,Anonymous)) in let p = apply_notation_to_pattern ?loc (ConstructRef cstr) - (match_notation_constr_cases_pattern t pat) allscopes vars keyrule in + (match_notation_constr_cases_pattern t pat) allscopes vars (keyrule,key) in insert_pat_alias ?loc p na | PatVar Anonymous -> CAst.make ?loc @@ CPatAtom None | PatVar (Name id) -> CAst.make ?loc @@ CPatAtom (Some (qualid_of_ident ?loc id)) @@ -511,11 +508,11 @@ and extern_notation_pattern allscopes vars t = function let rec extern_notation_ind_pattern allscopes vars ind args = function | [] -> raise No_match - | (keyrule,pat,n as _rule)::rules -> + | (keyrule,pat,n as _rule,key)::rules -> try if is_inactive_rule keyrule then raise No_match; apply_notation_to_pattern (IndRef ind) - (match_notation_constr_ind_pattern ind args pat) allscopes vars keyrule + (match_notation_constr_ind_pattern ind args pat) allscopes vars (keyrule,key) with No_match -> extern_notation_ind_pattern allscopes vars ind args rules @@ -1058,7 +1055,7 @@ and extern_eqn inctx scopes vars {CAst.loc;v=(ids,pll,c)} = and extern_notation (custom,scopes as allscopes) vars t = function | [] -> raise No_match - | (keyrule,pat,n as _rule)::rules -> + | (keyrule,pat,n as _rule,key)::rules -> let loc = Glob_ops.loc_of_glob_constr t in try if is_inactive_rule keyrule then raise No_match; @@ -1106,11 +1103,7 @@ and extern_notation (custom,scopes as allscopes) vars t = function (match availability_of_entry_coercion custom (fst ntn) with | None -> raise No_match | Some coercion -> - match availability_of_notation (sc,ntn) scopes with - (* Uninterpretation is not allowed in current context *) - | None -> raise No_match - (* Uninterpretation is allowed in current context *) - | Some (scopt,key) -> + let scopt = match key with Some _ -> sc | None -> None in let scopes' = Option.List.cons scopt (snd scopes) in let l = List.map (fun (c,(subentry,(scopt,scl))) -> diff --git a/interp/notation.ml b/interp/notation.ml index 49403c6e34..d5b43860c1 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -156,6 +156,8 @@ let scope_eq s1 s2 = match s1, s2 with | Scope _, SingleNotation _ | SingleNotation _, Scope _ -> false +(* Scopes for interpretation *) + let scope_stack = ref [] let current_scopes () = !scope_stack @@ -165,14 +167,56 @@ let scope_is_open_in_scopes sc l = let scope_is_open sc = scope_is_open_in_scopes sc (!scope_stack) +(* Uninterpretation tables *) + +type interp_rule = + | NotationRule of scope_name option * notation + | SynDefRule of KerName.t + +type notation_rule_core = interp_rule * interpretation * int option +type notation_rule = notation_rule_core * delimiters option + +(* Scopes for uninterpretation: includes abbreviations (i.e. syntactic definitions) and *) + +type uninterp_scope_elem = + | UninterpScope of scope_name + | UninterpSingle of notation_rule_core + +let uninterp_scope_eq_weak s1 s2 = match s1, s2 with +| UninterpScope s1, UninterpScope s2 -> String.equal s1 s2 +| UninterpSingle s1, UninterpSingle s2 -> false +| (UninterpSingle _ | UninterpScope _), _ -> false + +module ScopeOrd = + struct + type t = scope_name option + let compare = Pervasives.compare + end + +module ScopeMap = CMap.Make(ScopeOrd) + +let uninterp_scope_stack = ref [] + +let push_uninterp_scope sc scopes = UninterpScope sc :: scopes + +let push_uninterp_scopes = List.fold_right push_uninterp_scope + +let make_current_uninterp_scopes (tmp_scope,scopes) = + Option.fold_right push_uninterp_scope tmp_scope + (push_uninterp_scopes scopes !uninterp_scope_stack) + (* TODO: push nat_scope, z_scope, ... in scopes summary *) (* Exportation of scopes *) let open_scope i (_,(local,op,sc)) = - if Int.equal i 1 then + if Int.equal i 1 then begin scope_stack := - if op then sc :: !scope_stack - else List.except scope_eq sc !scope_stack + if op then Scope sc :: !scope_stack + else List.except scope_eq (Scope sc) !scope_stack; + uninterp_scope_stack := + if op then UninterpScope sc :: !uninterp_scope_stack + else List.except uninterp_scope_eq_weak (UninterpScope sc) !uninterp_scope_stack + end let cache_scope o = open_scope 1 o @@ -187,7 +231,7 @@ let discharge_scope (_,(local,_,_ as o)) = let classify_scope (local,_,_ as o) = if local then Dispose else Substitute o -let inScope : bool * bool * scope_elem -> obj = +let inScope : bool * bool * scope_name -> obj = declare_object {(default_object "SCOPE") with cache_function = cache_scope; open_function = open_scope; @@ -196,7 +240,7 @@ let inScope : bool * bool * scope_elem -> obj = classify_function = classify_scope } let open_close_scope (local,opening,sc) = - Lib.add_anonymous_leaf (inScope (local,opening,Scope (normalize_scope sc))) + Lib.add_anonymous_leaf (inScope (local,opening,normalize_scope sc)) let empty_scope_stack = [] @@ -250,40 +294,55 @@ let find_delimiters_scope ?loc key = user_err ?loc ~hdr:"find_delimiters" (str "Unknown scope delimiting key " ++ str key ++ str ".") -(* Uninterpretation tables *) - -type interp_rule = - | NotationRule of scope_name option * notation - | SynDefRule of KerName.t - (* We define keys for glob_constr and aconstr to split the syntax entries according to the key of the pattern (adapted from Chet Murthy by HH) *) type key = | RefKey of GlobRef.t + | LambdaKey + | ProdKey | Oth let key_compare k1 k2 = match k1, k2 with | RefKey gr1, RefKey gr2 -> GlobRef.Ordered.compare gr1 gr2 -| RefKey _, Oth -> -1 -| Oth, RefKey _ -> 1 -| Oth, Oth -> 0 +| RefKey _, _ -> -1 +| _, RefKey _ -> 1 +| k1, k2 -> Pervasives.compare k1 k2 module KeyOrd = struct type t = key let compare = key_compare end module KeyMap = Map.Make(KeyOrd) -type notation_rule = interp_rule * interpretation * int option +let keymap_add key sc interp map = + let oldkeymap = try ScopeMap.find sc map with Not_found -> KeyMap.empty in + let oldscmap = try KeyMap.find key oldkeymap with Not_found -> [] in + let newscmap = KeyMap.add key (interp :: oldscmap) oldkeymap in + ScopeMap.add sc newscmap map -let keymap_add key interp map = - let old = try KeyMap.find key map with Not_found -> [] in - KeyMap.add key (interp :: old) map +let keymap_extract keys sc map = + let keymap, map = + try ScopeMap.find sc map, ScopeMap.remove sc map + with Not_found -> KeyMap.empty, map in + let add_scope rule = (rule,None) in + List.map_append (fun key -> try List.map add_scope (KeyMap.find key keymap) with Not_found -> []) keys, map -let keymap_find key map = - try KeyMap.find key map - with Not_found -> [] +let find_with_delimiters = function + | None -> None + | Some scope -> + match (String.Map.find scope !scope_map).delimiters with + | Some key -> Some (Some key) + | None -> None + +let keymap_extract_remainder keys map = + ScopeMap.fold (fun sc keymap acc -> + match find_with_delimiters sc with + | None -> acc + | Some delim -> + let add_scope rule = (rule,delim) in + let l = List.map_append (fun key -> try List.map add_scope (KeyMap.find key keymap) with Not_found -> []) keys in + l @ acc) map [] (* Scopes table : interpretation -> scope_name *) -let notations_key_table = ref (KeyMap.empty : notation_rule list KeyMap.t) +let notations_key_table = ref (ScopeMap.empty : notation_rule_core list KeyMap.t ScopeMap.t) let glob_prim_constr_key c = match DAst.get c with | GRef (ref, _) -> Some (canonical_gr ref) @@ -295,12 +354,14 @@ let glob_prim_constr_key c = match DAst.get c with | _ -> None let glob_constr_keys c = match DAst.get c with + | GRef (ref,_) -> [RefKey (canonical_gr ref)] | GApp (c, _) -> begin match DAst.get c with | GRef (ref, _) -> [RefKey (canonical_gr ref); Oth] | _ -> [Oth] end - | GRef (ref,_) -> [RefKey (canonical_gr ref)] + | GLambda _ -> [LambdaKey] + | GProd _ -> [ProdKey] | _ -> [Oth] let cases_pattern_key c = match DAst.get c with @@ -314,6 +375,8 @@ let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *) RefKey (canonical_gr ref), Some (List.length args) | NRef ref -> RefKey(canonical_gr ref), None | NApp (_,args) -> Oth, Some (List.length args) + | NLambda _ | NBinderList (_,_,NLambda _,_,_) | NList (_,_,NLambda _,_,_) -> LambdaKey, None + | NProd _ | NBinderList (_,_,NProd _,_,_) | NList (_,_,NProd _,_,_) -> ProdKey, None | _ -> Oth, None (**********************************************************************) @@ -839,19 +902,12 @@ let check_required_module ?loc sc (sp,d) = (* Look if some notation or numeral printer in [scope] can be used in the scope stack [scopes], and if yes, using delimiters or not *) -let find_with_delimiters = function - | None -> None - | Some scope -> - match (String.Map.find scope !scope_map).delimiters with - | Some key -> Some (Some scope, Some key) - | None -> None - let rec find_without_delimiters find (ntn_scope,ntn) = function - | Scope scope :: scopes -> + | UninterpScope scope :: scopes -> (* Is the expected ntn/numpr attached to the most recently open scope? *) begin match ntn_scope with | Some scope' when String.equal scope scope' -> - Some (None,None) + Some None | _ -> (* If the most recently open scope has a notation/numeral printer but not the expected one then we need delimiters *) @@ -860,13 +916,14 @@ let rec find_without_delimiters find (ntn_scope,ntn) = function else find_without_delimiters find (ntn_scope,ntn) scopes end - | SingleNotation ntn' :: scopes -> + | UninterpSingle (NotationRule (_,ntn'),_,_) :: scopes -> begin match ntn_scope, ntn with | None, Some ntn when notation_eq ntn ntn' -> - Some (None, None) + Some None | _ -> find_without_delimiters find (ntn_scope,ntn) scopes end + | UninterpSingle (SynDefRule _,_,_) :: scopes -> find_without_delimiters find (ntn_scope,ntn) scopes | [] -> (* Can we switch to [scope]? Yes if it has defined delimiters *) find_with_delimiters ntn_scope @@ -902,9 +959,19 @@ let declare_notation_interpretation ntn scopt pat df ~onlyprint = | Some _ -> () end +let scope_of_rule = function + | NotationRule (None,_) | SynDefRule _ -> None + | NotationRule (Some sc as sco,_) -> sco + +let uninterp_scope_to_add pat n = function + | NotationRule (None,_) | SynDefRule _ as rule -> Some (UninterpSingle (rule,pat,n)) + | NotationRule (Some sc,_) -> None + let declare_uninterpretation rule (metas,c as pat) = let (key,n) = notation_constr_key c in - notations_key_table := keymap_add key (rule,pat,n) !notations_key_table + let sc = scope_of_rule rule in + notations_key_table := keymap_add key sc (rule,pat,n) !notations_key_table; + uninterp_scope_stack := Option.List.cons (uninterp_scope_to_add pat n rule) !uninterp_scope_stack let rec find_interpretation ntn find = function | [] -> raise Not_found @@ -982,43 +1049,27 @@ let interp_notation ?loc ntn local_scopes = user_err ?loc (str "Unknown interpretation for notation " ++ pr_notation ntn ++ str ".") -let sort_notations scopes l = - let extract_scope l = function - | Scope sc -> List.partitioni (fun _i x -> - match x with - | NotationRule (Some sc',_),_,_ -> String.equal sc sc' - | _ -> false) l - | SingleNotation ntn -> List.partitioni (fun _i x -> - match x with - | NotationRule (None,ntn'),_,_ -> notation_eq ntn ntn' - | _ -> false) l in - let rec aux l scopes = - if l == [] then [] (* shortcut *) else - match scopes with - | sc :: scopes -> let ntn_in_sc,l = extract_scope l sc in ntn_in_sc @ aux l scopes - | [] -> l in - aux l scopes +let extract_notations scopes keys = + if keys == [] then [] (* shortcut *) else + let rec aux scopes map = + match scopes with + | UninterpScope sc :: scopes -> + let l, map = keymap_extract keys (Some sc) map in l @ aux scopes map + | UninterpSingle rule :: scopes -> (rule,None) :: aux scopes map + | [] -> keymap_extract_remainder keys map + in aux scopes !notations_key_table let uninterp_notations scopes c = - let scopes = make_current_scopes scopes in - let keys = glob_constr_keys c in - let maps = List.map_append (fun key -> keymap_find key !notations_key_table) keys in - sort_notations scopes maps + let scopes = make_current_uninterp_scopes scopes in + extract_notations scopes (glob_constr_keys c) let uninterp_cases_pattern_notations scopes c = - let scopes = make_current_scopes scopes in - let maps = keymap_find (cases_pattern_key c) !notations_key_table in - sort_notations scopes maps + let scopes = make_current_uninterp_scopes scopes in + extract_notations scopes [cases_pattern_key c] let uninterp_ind_pattern_notations scopes ind = - let scopes = make_current_scopes scopes in - let maps = keymap_find (RefKey (canonical_gr (IndRef ind))) !notations_key_table in - sort_notations scopes maps - -let availability_of_notation (ntn_scope,ntn) scopes = - let f scope = - NotationMap.mem ntn (String.Map.find scope !scope_map).notations in - find_without_delimiters f (ntn_scope,Some ntn) (make_current_scopes scopes) + let scopes = make_current_uninterp_scopes scopes in + extract_notations scopes [RefKey (canonical_gr (IndRef ind))] (* We support coercions from a custom entry at some level to an entry at some level (possibly the same), and from and to the constr entry. E.g.: @@ -1172,8 +1223,8 @@ let availability_of_prim_token n printer_scope local_scopes = | _ -> false with Not_found -> false in - let scopes = make_current_scopes local_scopes in - Option.map snd (find_without_delimiters f (Some printer_scope,None) scopes) + let scopes = make_current_uninterp_scopes local_scopes in + find_without_delimiters f (Some printer_scope,None) scopes (* Miscellaneous *) @@ -1209,6 +1260,14 @@ let exists_notation_in_scope scopt ntn onlyprint r = interpretation_eq n.not_interp r with Not_found -> false +let exists_notation_interpretation_in_scope scopt ntn = + let scope = match scopt with Some s -> s | None -> default_scope in + try + let sc = String.Map.find scope !scope_map in + let _ = NotationMap.find ntn sc.notations in + true + with Not_found -> false + let isNVar_or_NHole = function NVar _ | NHole _ -> true | _ -> false (**********************************************************************) @@ -1690,17 +1749,18 @@ let pr_visibility prglob = function (* Synchronisation with reset *) let freeze _ = - (!scope_map, !scope_stack, !arguments_scope, + (!scope_map, !scope_stack, !uninterp_scope_stack, !arguments_scope, !delimiters_map, !notations_key_table, !scope_class_map, !prim_token_interp_infos, !prim_token_uninterp_infos, !entry_coercion_map, !entry_has_global_map, !entry_has_ident_map) -let unfreeze (scm,scs,asc,dlm,fkm,clsc,ptii,ptui,coe,globs,ids) = +let unfreeze (scm,scs,uscs,asc,dlm,fkm,clsc,ptii,ptui,coe,globs,ids) = scope_map := scm; scope_stack := scs; - delimiters_map := dlm; + uninterp_scope_stack := uscs; arguments_scope := asc; + delimiters_map := dlm; notations_key_table := fkm; scope_class_map := clsc; prim_token_interp_infos := ptii; @@ -1711,8 +1771,9 @@ let unfreeze (scm,scs,asc,dlm,fkm,clsc,ptii,ptui,coe,globs,ids) = let init () = init_scope_map (); + uninterp_scope_stack := []; delimiters_map := String.Map.empty; - notations_key_table := KeyMap.empty; + notations_key_table := ScopeMap.empty; scope_class_map := initial_scope_class_map; prim_token_interp_infos := String.Map.empty; prim_token_uninterp_infos := GlobRef.Map.empty diff --git a/interp/notation.mli b/interp/notation.mli index 40b7f2c2c9..8cc69f4d1b 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -211,18 +211,27 @@ val declare_uninterpretation : interp_rule -> interpretation -> unit val interp_notation : ?loc:Loc.t -> notation -> subscopes -> interpretation * (notation_location * scope_name option) -type notation_rule = interp_rule * interpretation * int option +type notation_rule_core = + interp_rule (* kind of notation *) + * interpretation (* pattern associated to the notation *) + * int option (* number of expected arguments *) + +type notation_rule = + notation_rule_core + * delimiters option (* delimiter to possibly add *) (** Return the possible notations for a given term *) val uninterp_notations : subscopes -> 'a glob_constr_g -> notation_rule list val uninterp_cases_pattern_notations : subscopes -> 'a cases_pattern_g -> notation_rule list val uninterp_ind_pattern_notations : subscopes -> inductive -> notation_rule list +(* (** Test if a notation is available in the scopes context [scopes]; if available, the result is not None; the first argument is itself not None if a delimiters is needed *) val availability_of_notation : scope_name option * notation -> subscopes -> (scope_name option * delimiters option) option + *) (** {6 Miscellaneous} *) @@ -233,6 +242,9 @@ val interp_notation_as_global_reference : ?loc:Loc.t -> (GlobRef.t -> bool) -> val exists_notation_in_scope : scope_name option -> notation -> bool -> interpretation -> bool +(** Checks for already existing notations *) +val exists_notation_interpretation_in_scope : scope_name option -> notation -> bool + (** Declares and looks for scopes associated to arguments of a global ref *) val declare_arguments_scope : bool (** true=local *) -> GlobRef.t -> scope_name option list -> unit diff --git a/test-suite/output/Inductive.out b/test-suite/output/Inductive.out index 6d65db9e22..5a548cfae4 100644 --- a/test-suite/output/Inductive.out +++ b/test-suite/output/Inductive.out @@ -1,6 +1,6 @@ The command has indeed failed with message: Last occurrence of "list'" must have "A" as 1st argument in - "A -> list' A -> list' (A * A)%type". + "A -> list' A -> list' (A * A)". Monomorphic Inductive foo (A : Type) (x : A) (y : A := x) : Prop := Foo : foo A x diff --git a/test-suite/output/Notations.out b/test-suite/output/Notations.out index e3816d65fd..bec4fc1579 100644 --- a/test-suite/output/Notations.out +++ b/test-suite/output/Notations.out @@ -64,7 +64,7 @@ The command has indeed failed with message: Cannot find where the recursive pattern starts. The command has indeed failed with message: Both ends of the recursive pattern are the same. -(nat * nat + nat)%type +SUM (nat * nat) nat : Set FST (0; 1) : Z diff --git a/test-suite/output/Notations2.out b/test-suite/output/Notations2.out index 672a742f24..41d1593758 100644 --- a/test-suite/output/Notations2.out +++ b/test-suite/output/Notations2.out @@ -51,7 +51,7 @@ end match n with | nil => 2 | 0 :: _ => 2 -| 1 :: nil => 0 +| list1 => 0 | 1 :: _ :: _ => 2 | plus2 _ :: _ => 2 end |
