| author | wenzelm | 
| Fri, 07 Jul 2023 18:04:45 +0200 | |
| changeset 78265 | 03eb7f7bb990 | 
| parent 71911 | d25093536482 | 
| child 80875 | 2e33897071b6 | 
| permissions | -rw-r--r-- | 
| 59720 | 1 | (* Title: Pure/General/completion.ML | 
| 55672 | 2 | Author: Makarius | 
| 3 | ||
| 55674 | 4 | Semantic completion within the formal context. | 
| 55672 | 5 | *) | 
| 6 | ||
| 7 | signature COMPLETION = | |
| 8 | sig | |
| 69289 | 9 | type name = string * (string * string) | 
| 55687 | 10 | type T | 
| 69289 | 11 | val names: Position.T -> name list -> T | 
| 55672 | 12 | val none: T | 
| 69289 | 13 | val make: string * Position.T -> ((string -> bool) -> name list) -> T | 
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 14 | val encode: T -> XML.body | 
| 69347 
54b95d2ec040
more robust: avoid broken YXML due to Markup.empty;
 wenzelm parents: 
69289diff
changeset | 15 | val markup_element: T -> (Markup.T * XML.body) option | 
| 69289 | 16 | val markup_report: T list -> string | 
| 17 | val make_report: string * Position.T -> ((string -> bool) -> name list) -> string | |
| 55915 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 18 | val suppress_abbrevs: string -> Markup.T list | 
| 71911 | 19 | val check_item: string -> (string * 'a -> Markup.T) -> | 
| 20 | (string * 'a) list -> Proof.context -> string * Position.T -> string | |
| 21 | val check_entity: string -> (string * Position.T) list -> | |
| 22 | Proof.context -> string * Position.T -> string | |
| 67208 | 23 | val check_option: Options.T -> Proof.context -> string * Position.T -> string | 
| 24 | val check_option_value: | |
| 25 | Proof.context -> string * Position.T -> string * Position.T -> Options.T -> string * Options.T | |
| 55672 | 26 | end; | 
| 27 | ||
| 28 | structure Completion: COMPLETION = | |
| 29 | struct | |
| 30 | ||
| 59812 | 31 | (* completion of names *) | 
| 32 | ||
| 69289 | 33 | type name = string * (string * string); (*external name, kind, internal name*) | 
| 34 | ||
| 35 | abstype T = Completion of {pos: Position.T, total: int, names: name list}
 | |
| 55687 | 36 | with | 
| 55672 | 37 | |
| 55687 | 38 | fun dest (Completion args) = args; | 
| 55672 | 39 | |
| 55694 
a1184dfb8e00
clarified semantic completion: retain kind.full_name as official item name for history;
 wenzelm parents: 
55687diff
changeset | 40 | fun names pos names = | 
| 55687 | 41 | Completion | 
| 42 |    {pos = pos,
 | |
| 43 | total = length names, | |
| 44 | names = take (Options.default_int "completion_limit") names}; | |
| 55672 | 45 | |
| 46 | end; | |
| 55687 | 47 | |
| 55694 
a1184dfb8e00
clarified semantic completion: retain kind.full_name as official item name for history;
 wenzelm parents: 
55687diff
changeset | 48 | val none = names Position.none []; | 
| 55687 | 49 | |
| 59812 | 50 | fun make (name, pos) make_names = | 
| 51 | if Position.is_reported pos andalso name <> "" andalso name <> "_" | |
| 52 | then names pos (make_names (String.isPrefix (Name.clean name))) | |
| 53 | else none; | |
| 54 | ||
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 55 | fun encode completion = | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 56 | let | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 57 |     val {total, names, ...} = dest completion;
 | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 58 | open XML.Encode; | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 59 | in pair int (list (pair string (pair string string))) (total, names) end; | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 60 | |
| 69289 | 61 | fun markup_element completion = | 
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 62 |   let val {pos, names, ...} = dest completion in
 | 
| 55687 | 63 | if Position.is_reported pos andalso not (null names) then | 
| 69347 
54b95d2ec040
more robust: avoid broken YXML due to Markup.empty;
 wenzelm parents: 
69289diff
changeset | 64 | SOME (Position.markup pos Markup.completion, encode completion) | 
| 
54b95d2ec040
more robust: avoid broken YXML due to Markup.empty;
 wenzelm parents: 
69289diff
changeset | 65 | else NONE | 
| 55687 | 66 | end; | 
| 67 | ||
| 69289 | 68 | val markup_report = | 
| 69347 
54b95d2ec040
more robust: avoid broken YXML due to Markup.empty;
 wenzelm parents: 
69289diff
changeset | 69 | map_filter markup_element #> map XML.Elem #> YXML.string_of_body #> Markup.markup_report; | 
| 69289 | 70 | |
| 71 | val make_report = markup_report oo (single oo make); | |
| 72 | ||
| 55915 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 73 | |
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 74 | (* suppress short abbreviations *) | 
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 75 | |
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 76 | fun suppress_abbrevs s = | 
| 55917 
5438ed05e1c9
special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
 wenzelm parents: 
55915diff
changeset | 77 | if not (Symbol.is_ascii_identifier s) andalso (length (Symbol.explode s) <= 1 orelse s = "::") | 
| 55915 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 78 | then [Markup.no_completion] | 
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 79 | else []; | 
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 80 | |
| 67208 | 81 | |
| 71911 | 82 | (* check items *) | 
| 67208 | 83 | |
| 71911 | 84 | fun check_item kind markup items ctxt (name, pos) = | 
| 85 | (case AList.lookup (op =) items name of | |
| 86 | SOME x => (Context_Position.report ctxt pos (markup (name, x)); name) | |
| 87 | | NONE => | |
| 88 | let | |
| 89 | fun make_names completed = | |
| 90 | map_filter (fn (a, _) => if completed a then SOME (a, (kind, a)) else NONE) items; | |
| 91 | val completion_report = make_report (name, pos) make_names; | |
| 92 |       in error ("Bad " ^ kind ^ " " ^ quote name ^ Position.here pos ^ completion_report) end);
 | |
| 93 | ||
| 94 | fun check_entity kind = check_item kind (Position.entity_markup kind); | |
| 95 | ||
| 96 | ||
| 97 | val check_option = check_entity Markup.system_optionN o Options.dest; | |
| 67208 | 98 | |
| 99 | fun check_option_value ctxt (name, pos) (value, pos') options = | |
| 100 | let | |
| 101 | val _ = check_option options ctxt (name, pos); | |
| 102 | val options' = | |
| 103 | Options.update name value options | |
| 104 | handle ERROR msg => error (msg ^ Position.here pos'); | |
| 105 | in (name, options') end; | |
| 106 | ||
| 55687 | 107 | end; |