| author | wenzelm | 
| Mon, 02 Jul 2018 16:43:06 +0200 | |
| changeset 68572 | c8bf6077a87d | 
| parent 67208 | 16519cd83ed4 | 
| child 69289 | bf6937af7fe8 | 
| 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 | |
| 55687 | 9 | type T | 
| 55977 | 10 | val names: Position.T -> (string * (string * string)) list -> T | 
| 55672 | 11 | val none: T | 
| 59812 | 12 | val make: string * Position.T -> ((string -> bool) -> (string * (string * string)) list) -> T | 
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 13 | val encode: T -> XML.body | 
| 55840 
2982d233d798
consider completion report as part of error message -- less stateful, may get handled;
 wenzelm parents: 
55694diff
changeset | 14 | val reported_text: T -> string | 
| 55915 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 15 | val suppress_abbrevs: string -> Markup.T list | 
| 67208 | 16 | val check_option: Options.T -> Proof.context -> string * Position.T -> string | 
| 17 | val check_option_value: | |
| 18 | Proof.context -> string * Position.T -> string * Position.T -> Options.T -> string * Options.T | |
| 55672 | 19 | end; | 
| 20 | ||
| 21 | structure Completion: COMPLETION = | |
| 22 | struct | |
| 23 | ||
| 59812 | 24 | (* completion of names *) | 
| 25 | ||
| 55977 | 26 | abstype T = | 
| 27 |   Completion of {pos: Position.T, total: int, names: (string * (string * string)) list}
 | |
| 55687 | 28 | with | 
| 55672 | 29 | |
| 55687 | 30 | fun dest (Completion args) = args; | 
| 55672 | 31 | |
| 55694 
a1184dfb8e00
clarified semantic completion: retain kind.full_name as official item name for history;
 wenzelm parents: 
55687diff
changeset | 32 | fun names pos names = | 
| 55687 | 33 | Completion | 
| 34 |    {pos = pos,
 | |
| 35 | total = length names, | |
| 36 | names = take (Options.default_int "completion_limit") names}; | |
| 55672 | 37 | |
| 38 | end; | |
| 55687 | 39 | |
| 55694 
a1184dfb8e00
clarified semantic completion: retain kind.full_name as official item name for history;
 wenzelm parents: 
55687diff
changeset | 40 | val none = names Position.none []; | 
| 55687 | 41 | |
| 59812 | 42 | fun make (name, pos) make_names = | 
| 43 | if Position.is_reported pos andalso name <> "" andalso name <> "_" | |
| 44 | then names pos (make_names (String.isPrefix (Name.clean name))) | |
| 45 | else none; | |
| 46 | ||
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 47 | fun encode completion = | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 48 | let | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 49 |     val {total, names, ...} = dest completion;
 | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 50 | open XML.Encode; | 
| 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 51 | 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 | 52 | |
| 55840 
2982d233d798
consider completion report as part of error message -- less stateful, may get handled;
 wenzelm parents: 
55694diff
changeset | 53 | fun reported_text completion = | 
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 54 |   let val {pos, names, ...} = dest completion in
 | 
| 55687 | 55 | if Position.is_reported pos andalso not (null names) then | 
| 56 | let | |
| 55694 
a1184dfb8e00
clarified semantic completion: retain kind.full_name as official item name for history;
 wenzelm parents: 
55687diff
changeset | 57 | val markup = Position.markup pos Markup.completion; | 
| 60731 
4ac4b314d93c
additional ML parse tree components for Poly/ML 5.5.3, or later;
 wenzelm parents: 
59812diff
changeset | 58 | in YXML.string_of (XML.Elem (markup, encode completion)) end | 
| 55840 
2982d233d798
consider completion report as part of error message -- less stateful, may get handled;
 wenzelm parents: 
55694diff
changeset | 59 | else "" | 
| 55687 | 60 | end; | 
| 61 | ||
| 55915 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 62 | |
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 63 | (* suppress short abbreviations *) | 
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 64 | |
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 65 | fun suppress_abbrevs s = | 
| 55917 
5438ed05e1c9
special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
 wenzelm parents: 
55915diff
changeset | 66 | 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 | 67 | then [Markup.no_completion] | 
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 68 | else []; | 
| 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 wenzelm parents: 
55840diff
changeset | 69 | |
| 67208 | 70 | |
| 71 | (* system options *) | |
| 72 | ||
| 73 | fun check_option options ctxt (name, pos) = | |
| 74 | let | |
| 75 | val markup = | |
| 76 | Options.markup options (name, pos) handle ERROR msg => | |
| 77 | let | |
| 78 | val completion = | |
| 79 | make (name, pos) (fn completed => | |
| 80 | Options.names options | |
| 81 | |> filter completed | |
| 82 |                 |> map (fn a => (a, ("system_option", a))));
 | |
| 83 | val report = Markup.markup_report (reported_text completion); | |
| 84 | in error (msg ^ report) end; | |
| 85 | val _ = Context_Position.report ctxt pos markup; | |
| 86 | in name end; | |
| 87 | ||
| 88 | fun check_option_value ctxt (name, pos) (value, pos') options = | |
| 89 | let | |
| 90 | val _ = check_option options ctxt (name, pos); | |
| 91 | val options' = | |
| 92 | Options.update name value options | |
| 93 | handle ERROR msg => error (msg ^ Position.here pos'); | |
| 94 | in (name, options') end; | |
| 95 | ||
| 55687 | 96 | end; |