| author | paulson <lp15@cam.ac.uk> | 
| Mon, 24 Sep 2018 14:33:08 +0100 | |
| changeset 69109 | c9ea9290880f | 
| parent 49998 | ad8a6db0b480 | 
| child 80910 | 406a85a25189 | 
| permissions | -rw-r--r-- | 
| 41561 | 1 | (* Title: HOL/SPARK/Tools/fdl_parser.ML | 
| 2 | Author: Stefan Berghofer | |
| 3 | Copyright: secunet Security Networks AG | |
| 4 | ||
| 5 | Parser for fdl files. | |
| 6 | *) | |
| 7 | ||
| 8 | signature FDL_PARSER = | |
| 9 | sig | |
| 10 | datatype expr = | |
| 11 | Ident of string | |
| 12 | | Number of int | |
| 13 | | Quantifier of string * string list * string * expr | |
| 14 | | Funct of string * expr list | |
| 15 | | Element of expr * expr list | |
| 16 | | Update of expr * expr list * expr | |
| 17 | | Record of string * (string * expr) list | |
| 18 | | Array of string * expr option * | |
| 19 | ((expr * expr option) list list * expr) list; | |
| 20 | ||
| 21 | datatype fdl_type = | |
| 22 | Basic_Type of string | |
| 23 | | Enum_Type of string list | |
| 24 | | Array_Type of string list * string | |
| 25 | | Record_Type of (string list * string) list | |
| 26 | | Pending_Type; | |
| 27 | ||
| 28 | datatype fdl_rule = | |
| 29 | Inference_Rule of expr list * expr | |
| 30 | | Substitution_Rule of expr list * expr * expr; | |
| 31 | ||
| 32 | type rules = | |
| 33 | ((string * int) * fdl_rule) list * | |
| 34 | (string * (expr * (string * string) list) list) list; | |
| 35 | ||
| 36 | type vcs = (string * (string * | |
| 37 | ((string * expr) list * (string * expr) list)) list) list; | |
| 38 | ||
| 39 | type 'a tab = 'a Symtab.table * (string * 'a) list; | |
| 40 | ||
| 41 | val lookup: 'a tab -> string -> 'a option; | |
| 42 | val update: string * 'a -> 'a tab -> 'a tab; | |
| 43 | val items: 'a tab -> (string * 'a) list; | |
| 44 | ||
| 45 | type decls = | |
| 46 |     {types: fdl_type tab,
 | |
| 47 | vars: string tab, | |
| 48 | consts: string tab, | |
| 49 | funs: (string list * string) tab}; | |
| 50 | ||
| 51 | val parse_vcs: (Fdl_Lexer.chars -> 'a * Fdl_Lexer.chars) -> Position.T -> | |
| 52 | string -> 'a * ((string * string) * vcs); | |
| 53 | ||
| 54 | val parse_declarations: Position.T -> string -> (string * string) * decls; | |
| 55 | ||
| 56 | val parse_rules: Position.T -> string -> rules; | |
| 57 | end; | |
| 58 | ||
| 59 | structure Fdl_Parser: FDL_PARSER = | |
| 60 | struct | |
| 61 | ||
| 62 | (** error handling **) | |
| 63 | ||
| 64 | fun beginning n cs = | |
| 65 | let val dots = if length cs > n then " ..." else ""; | |
| 66 | in | |
| 67 | space_implode " " (take n cs) ^ dots | |
| 68 | end; | |
| 69 | ||
| 70 | fun !!! scan = | |
| 71 | let | |
| 48911 
5debc3e4fa81
tuned messages: end-of-input rarely means physical end-of-file from the past;
 wenzelm parents: 
47083diff
changeset | 72 | fun get_pos [] = " (end-of-input)" | 
| 41561 | 73 | | get_pos (tok :: _) = Fdl_Lexer.pos_of tok; | 
| 74 | ||
| 43947 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 wenzelm parents: 
41620diff
changeset | 75 | fun err (syms, msg) = fn () => | 
| 41561 | 76 | "Syntax error" ^ get_pos syms ^ " at " ^ | 
| 43947 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 wenzelm parents: 
41620diff
changeset | 77 | beginning 10 (map Fdl_Lexer.unparse syms) ^ | 
| 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 wenzelm parents: 
41620diff
changeset | 78 | (case msg of NONE => "" | SOME m => "\n" ^ m () ^ " expected"); | 
| 41561 | 79 | in Scan.!! err scan end; | 
| 80 | ||
| 81 | fun parse_all p = | |
| 82 | Scan.repeat (Scan.unless (Scan.one Fdl_Lexer.is_eof) (!!! p)); | |
| 83 | ||
| 84 | ||
| 85 | (** parsers **) | |
| 86 | ||
| 43947 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 wenzelm parents: 
41620diff
changeset | 87 | fun group s p = p || Scan.fail_with (K (fn () => s)); | 
| 41561 | 88 | |
| 89 | fun $$$ s = group s | |
| 90 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Keyword andalso | |
| 91 | Fdl_Lexer.content_of t = s) >> K s); | |
| 92 | ||
| 93 | val identifier = group "identifier" | |
| 94 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident) >> | |
| 95 | Fdl_Lexer.content_of); | |
| 96 | ||
| 97 | val long_identifier = group "long identifier" | |
| 98 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Long_Ident) >> | |
| 99 | Fdl_Lexer.content_of); | |
| 100 | ||
| 101 | fun the_identifier s = group s | |
| 102 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident andalso | |
| 103 | Fdl_Lexer.content_of t = s) >> K s); | |
| 104 | ||
| 105 | fun prfx_identifier g s = group g | |
| 106 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident andalso | |
| 107 | can (unprefix s) (Fdl_Lexer.content_of t)) >> | |
| 108 | (unprefix s o Fdl_Lexer.content_of)); | |
| 109 | ||
| 110 | val mk_identifier = prfx_identifier "identifier \"mk__...\"" "mk__"; | |
| 111 | val hyp_identifier = prfx_identifier "hypothesis identifer" "H"; | |
| 112 | val concl_identifier = prfx_identifier "conclusion identifier" "C"; | |
| 113 | ||
| 114 | val number = group "number" | |
| 115 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Number) >> | |
| 116 | (the o Int.fromString o Fdl_Lexer.content_of)); | |
| 117 | ||
| 118 | val traceability = group "traceability information" | |
| 119 | (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Traceability) >> | |
| 120 | Fdl_Lexer.content_of); | |
| 121 | ||
| 122 | fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan); | |
| 123 | fun enum sep scan = enum1 sep scan || Scan.succeed []; | |
| 124 | ||
| 125 | fun list1 scan = enum1 "," scan; | |
| 126 | fun list scan = enum "," scan; | |
| 127 | ||
| 128 | ||
| 129 | (* expressions, see section 4.4 of SPARK Proof Manual *) | |
| 130 | ||
| 131 | datatype expr = | |
| 132 | Ident of string | |
| 133 | | Number of int | |
| 134 | | Quantifier of string * string list * string * expr | |
| 135 | | Funct of string * expr list | |
| 136 | | Element of expr * expr list | |
| 137 | | Update of expr * expr list * expr | |
| 138 | | Record of string * (string * expr) list | |
| 139 | | Array of string * expr option * | |
| 140 | ((expr * expr option) list list * expr) list; | |
| 141 | ||
| 142 | fun unop (f, x) = Funct (f, [x]); | |
| 143 | ||
| 144 | fun binop p q = p :|-- (fn x => Scan.optional | |
| 145 | (q -- !!! p >> (fn (f, y) => Funct (f, [x, y]))) x); | |
| 146 | ||
| 147 | (* left-associative *) | |
| 148 | fun binops opp argp = | |
| 149 | argp -- Scan.repeat (opp -- !!! argp) >> (fn (x, fys) => | |
| 150 | fold (fn (f, y) => fn x => Funct (f, [x, y])) fys x); | |
| 151 | ||
| 152 | (* right-associative *) | |
| 153 | fun binops' f p = enum1 f p >> foldr1 (fn (x, y) => Funct (f, [x, y])); | |
| 154 | ||
| 155 | val multiplying_operator = $$$ "*" || $$$ "/" || $$$ "div" || $$$ "mod"; | |
| 156 | ||
| 157 | val adding_operator = $$$ "+" || $$$ "-"; | |
| 158 | ||
| 159 | val relational_operator = | |
| 160 | $$$ "=" || $$$ "<>" | |
| 161 | || $$$ "<" || $$$ ">" | |
| 162 | || $$$ "<="|| $$$ ">="; | |
| 163 | ||
| 164 | val quantification_kind = $$$ "for_all" || $$$ "for_some"; | |
| 165 | ||
| 166 | val quantification_generator = | |
| 167 | list1 identifier --| $$$ ":" -- identifier; | |
| 168 | ||
| 49998 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 169 | fun opt_parens p = $$$ "(" |-- p --| $$$ ")" || p;
 | 
| 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 170 | |
| 41561 | 171 | fun expression xs = group "expression" | 
| 172 | (binop disjunction ($$$ "->" || $$$ "<->")) xs | |
| 173 | ||
| 174 | and disjunction xs = binops' "or" conjunction xs | |
| 175 | ||
| 176 | and conjunction xs = binops' "and" negation xs | |
| 177 | ||
| 178 | and negation xs = | |
| 179 | ( $$$ "not" -- !!! relation >> unop | |
| 180 | || relation) xs | |
| 181 | ||
| 182 | and relation xs = binop sum relational_operator xs | |
| 183 | ||
| 184 | and sum xs = binops adding_operator term xs | |
| 185 | ||
| 186 | and term xs = binops multiplying_operator factor xs | |
| 187 | ||
| 188 | and factor xs = | |
| 189 | ( $$$ "+" |-- !!! primary | |
| 190 | || $$$ "-" -- !!! primary >> unop | |
| 191 | || binop primary ($$$ "**")) xs | |
| 192 | ||
| 193 | and primary xs = group "primary" | |
| 194 | ( number >> Number | |
| 195 |    || $$$ "(" |-- !!! (expression --| $$$ ")")
 | |
| 196 | || quantified_expression | |
| 197 | || function_designator | |
| 198 | || identifier >> Ident) xs | |
| 199 | ||
| 200 | and quantified_expression xs = (quantification_kind -- | |
| 201 |   !!! ($$$ "(" |-- quantification_generator --|  $$$ "," --
 | |
| 202 | expression --| $$$ ")") >> (fn (q, ((xs, T), e)) => | |
| 203 | Quantifier (q, xs, T, e))) xs | |
| 204 | ||
| 205 | and function_designator xs = | |
| 206 |   (   mk_identifier --| $$$ "(" :|--
 | |
| 207 | (fn s => record_args s || array_args s) --| $$$ ")" | |
| 208 |    || $$$ "element" |-- !!! ($$$ "(" |-- expression --| $$$ "," --| $$$ "[" --
 | |
| 209 | list1 expression --| $$$ "]" --| $$$ ")") >> Element | |
| 210 |    || $$$ "update" |-- !!! ($$$ "(" |-- expression --| $$$ "," --| $$$ "[" --
 | |
| 211 | list1 expression --| $$$ "]" --| $$$ "," -- expression --| $$$ ")") >> | |
| 212 | (fn ((A, xs), x) => Update (A, xs, x)) | |
| 213 |    || identifier --| $$$ "(" -- !!! (list1 expression --| $$$ ")") >> Funct) xs
 | |
| 214 | ||
| 215 | and record_args s xs = | |
| 216 | (list1 (identifier --| $$$ ":=" -- !!! expression) >> (pair s #> Record)) xs | |
| 217 | ||
| 218 | and array_args s xs = | |
| 49998 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 219 | ( array_associations >> (fn assocs => Array (s, NONE, assocs)) | 
| 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 220 | || expression -- Scan.optional ($$$ "," |-- !!! array_associations) [] >> | 
| 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 221 | (fn (default, assocs) => Array (s, SOME default, assocs))) xs | 
| 41561 | 222 | |
| 223 | and array_associations xs = | |
| 49998 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 224 | (list1 (opt_parens (enum1 "&" ($$$ "[" |-- | 
| 41561 | 225 | !!! (list1 (expression -- Scan.option ($$$ ".." |-- !!! expression)) --| | 
| 49998 
ad8a6db0b480
Allow parentheses around left-hand sides of array associations
 berghofe parents: 
48911diff
changeset | 226 | $$$ "]"))) --| $$$ ":=" -- expression)) xs; | 
| 41561 | 227 | |
| 228 | ||
| 229 | (* verification conditions *) | |
| 230 | ||
| 231 | type vcs = (string * (string * | |
| 232 | ((string * expr) list * (string * expr) list)) list) list; | |
| 233 | ||
| 234 | val vc = | |
| 235 | identifier --| $$$ "." -- !!! | |
| 236 | ( $$$ "***" |-- !!! (the_identifier "true" --| $$$ ".") >> | |
| 237 | (Ident #> pair "1" #> single #> pair []) | |
| 238 | || $$$ "!!!" |-- !!! (the_identifier "false" --| $$$ ".") >> | |
| 239 | (Ident #> pair "1" #> single #> pair []) | |
| 240 | || Scan.repeat1 (hyp_identifier -- | |
| 241 | !!! ($$$ ":" |-- expression --| $$$ ".")) --| $$$ "->" -- | |
| 242 | Scan.repeat1 (concl_identifier -- | |
| 243 | !!! ($$$ ":" |-- expression --| $$$ "."))); | |
| 244 | ||
| 245 | val subprogram_kind = $$$ "function" || $$$ "procedure"; | |
| 246 | ||
| 247 | val vcs = | |
| 248 | subprogram_kind -- (long_identifier || identifier) -- | |
| 249 | parse_all (traceability -- !!! (Scan.repeat1 vc)); | |
| 250 | ||
| 251 | fun parse_vcs header pos s = | |
| 252 | s |> | |
| 253 | Fdl_Lexer.tokenize header Fdl_Lexer.c_comment pos ||> | |
| 254 | filter Fdl_Lexer.is_proper ||> | |
| 255 | Scan.finite Fdl_Lexer.stopper (Scan.error (!!! vcs)) ||> | |
| 256 | fst; | |
| 257 | ||
| 258 | ||
| 259 | (* fdl declarations, see section 4.3 of SPARK Proof Manual *) | |
| 260 | ||
| 261 | datatype fdl_type = | |
| 262 | Basic_Type of string | |
| 263 | | Enum_Type of string list | |
| 264 | | Array_Type of string list * string | |
| 265 | | Record_Type of (string list * string) list | |
| 266 | | Pending_Type; | |
| 267 | ||
| 268 | (* also store items in a list to preserve order *) | |
| 269 | type 'a tab = 'a Symtab.table * (string * 'a) list; | |
| 270 | ||
| 271 | fun lookup ((tab, _) : 'a tab) = Symtab.lookup tab; | |
| 272 | fun update decl (tab, items) = (Symtab.update_new decl tab, decl :: items); | |
| 273 | fun items ((_, items) : 'a tab) = rev items; | |
| 274 | ||
| 275 | type decls = | |
| 276 |   {types: fdl_type tab,
 | |
| 277 | vars: string tab, | |
| 278 | consts: string tab, | |
| 279 | funs: (string list * string) tab}; | |
| 280 | ||
| 281 | val empty_decls : decls = | |
| 282 |   {types = (Symtab.empty, []), vars = (Symtab.empty, []),
 | |
| 283 | consts = (Symtab.empty, []), funs = (Symtab.empty, [])}; | |
| 284 | ||
| 285 | fun add_type_decl decl {types, vars, consts, funs} =
 | |
| 286 |   {types = update decl types,
 | |
| 287 | vars = vars, consts = consts, funs = funs} | |
| 288 |   handle Symtab.DUP s => error ("Duplicate type " ^ s);
 | |
| 289 | ||
| 290 | fun add_var_decl (vs, ty) {types, vars, consts, funs} =
 | |
| 291 |   {types = types,
 | |
| 292 | vars = fold (update o rpair ty) vs vars, | |
| 293 | consts = consts, funs = funs} | |
| 294 |   handle Symtab.DUP s => error ("Duplicate variable " ^ s);
 | |
| 295 | ||
| 296 | fun add_const_decl decl {types, vars, consts, funs} =
 | |
| 297 |   {types = types, vars = vars,
 | |
| 298 | consts = update decl consts, | |
| 299 | funs = funs} | |
| 300 |   handle Symtab.DUP s => error ("Duplicate constant " ^ s);
 | |
| 301 | ||
| 302 | fun add_fun_decl decl {types, vars, consts, funs} =
 | |
| 303 |   {types = types, vars = vars, consts = consts,
 | |
| 304 | funs = update decl funs} | |
| 305 |   handle Symtab.DUP s => error ("Duplicate function " ^ s);
 | |
| 306 | ||
| 41620 | 307 | fun type_decl x = ($$$ "type" |-- !!! (identifier --| $$$ "=" -- | 
| 41561 | 308 | ( identifier >> Basic_Type | 
| 309 |    || $$$ "(" |-- !!! (list1 identifier --| $$$ ")") >> Enum_Type
 | |
| 310 | || $$$ "array" |-- !!! ($$$ "[" |-- list1 identifier --| $$$ "]" --| | |
| 311 | $$$ "of" -- identifier) >> Array_Type | |
| 312 | || $$$ "record" |-- !!! (enum1 ";" | |
| 313 | (list1 identifier -- !!! ($$$ ":" |-- identifier)) --| | |
| 314 | $$$ "end") >> Record_Type | |
| 41620 | 315 | || $$$ "pending" >> K Pending_Type)) >> add_type_decl) x; | 
| 41561 | 316 | |
| 41620 | 317 | fun const_decl x = ($$$ "const" |-- !!! (identifier --| $$$ ":" -- identifier --| | 
| 318 | $$$ "=" --| $$$ "pending") >> add_const_decl) x; | |
| 41561 | 319 | |
| 41620 | 320 | fun var_decl x = ($$$ "var" |-- !!! (list1 identifier --| $$$ ":" -- identifier) >> | 
| 321 | add_var_decl) x; | |
| 41561 | 322 | |
| 41620 | 323 | fun fun_decl x = ($$$ "function" |-- !!! (identifier -- | 
| 41561 | 324 |   (Scan.optional ($$$ "(" |-- !!! (list1 identifier --| $$$ ")")) [] --|
 | 
| 41620 | 325 | $$$ ":" -- identifier)) >> add_fun_decl) x; | 
| 41561 | 326 | |
| 41620 | 327 | fun declarations x = | 
| 47083 | 328 | (the_identifier "title" |-- subprogram_kind -- identifier --| $$$ ";" -- | 
| 41561 | 329 | (Scan.repeat ((type_decl || const_decl || var_decl || fun_decl) --| | 
| 46778 | 330 | !!! ($$$ ";")) >> (fn ds => fold I ds empty_decls)) --| | 
| 41620 | 331 | $$$ "end" --| $$$ ";") x; | 
| 41561 | 332 | |
| 333 | fun parse_declarations pos s = | |
| 334 | s |> | |
| 335 | Fdl_Lexer.tokenize (Scan.succeed ()) Fdl_Lexer.curly_comment pos |> | |
| 336 | snd |> filter Fdl_Lexer.is_proper |> | |
| 337 | Scan.finite Fdl_Lexer.stopper (Scan.error (!!! declarations)) |> | |
| 338 | fst; | |
| 339 | ||
| 340 | ||
| 341 | (* rules, see section 5 of SPADE Proof Checker Rules Manual *) | |
| 342 | ||
| 343 | datatype fdl_rule = | |
| 344 | Inference_Rule of expr list * expr | |
| 345 | | Substitution_Rule of expr list * expr * expr; | |
| 346 | ||
| 347 | type rules = | |
| 348 | ((string * int) * fdl_rule) list * | |
| 349 | (string * (expr * (string * string) list) list) list; | |
| 350 | ||
| 351 | val condition_list = $$$ "[" |-- list expression --| $$$ "]"; | |
| 352 | val if_condition_list = $$$ "if" |-- !!! condition_list; | |
| 353 | ||
| 354 | val rule = | |
| 355 |   identifier -- !!! ($$$ "(" |-- number --| $$$ ")" --| $$$ ":" --
 | |
| 356 | (expression :|-- (fn e => | |
| 357 | $$$ "may_be_deduced" >> K (Inference_Rule ([], e)) | |
| 358 | || $$$ "may_be_deduced_from" |-- | |
| 359 | !!! condition_list >> (Inference_Rule o rpair e) | |
| 360 | || $$$ "may_be_replaced_by" |-- !!! (expression -- | |
| 361 | Scan.optional if_condition_list []) >> (fn (e', cs) => | |
| 362 | Substitution_Rule (cs, e, e')) | |
| 363 | || $$$ "&" |-- !!! (expression --| $$$ "are_interchangeable" -- | |
| 364 | Scan.optional if_condition_list []) >> (fn (e', cs) => | |
| 365 | Substitution_Rule (cs, e, e')))) --| $$$ ".") >> | |
| 366 | (fn (id, (n, rl)) => ((id, n), rl)); | |
| 367 | ||
| 368 | val rule_family = | |
| 369 | $$$ "rule_family" |-- identifier --| $$$ ":" -- | |
| 370 | enum1 "&" (expression -- !!! ($$$ "requires" |-- $$$ "[" |-- | |
| 371 | list (identifier -- !!! ($$$ ":" |-- identifier)) --| $$$ "]")) --| | |
| 372 | $$$ "."; | |
| 373 | ||
| 374 | val rules = | |
| 375 | parse_all (rule >> (apfst o cons) || rule_family >> (apsnd o cons)) >> | |
| 46778 | 376 | (fn rls => fold_rev I rls ([], [])); | 
| 41561 | 377 | |
| 378 | fun parse_rules pos s = | |
| 379 | s |> | |
| 380 | Fdl_Lexer.tokenize (Scan.succeed ()) | |
| 381 | (Fdl_Lexer.c_comment || Fdl_Lexer.percent_comment) pos |> | |
| 382 | snd |> filter Fdl_Lexer.is_proper |> | |
| 383 | Scan.finite Fdl_Lexer.stopper (Scan.error (!!! rules)) |> | |
| 384 | fst; | |
| 385 | ||
| 386 | end; |