src/HOL/Boogie/Tools/boogie_loader.ML
author wenzelm
Thu, 09 Jun 2011 17:51:49 +0200
changeset 44208 47cf4bc789aa
parent 43246 774df7c59508
child 44211 84472e198515
permissions -rw-r--r--
simplified Name.variant -- discontinued builtin fold_map;
     1 (*  Title:      HOL/Boogie/Tools/boogie_loader.ML
     2     Author:     Sascha Boehme, TU Muenchen
     3 
     4 Loading and interpreting Boogie-generated files.
     5 *)
     6 
     7 signature BOOGIE_LOADER =
     8 sig
     9   val load_b2i: bool -> (string * int) list -> Path.T -> theory -> theory
    10 end
    11 
    12 structure Boogie_Loader: BOOGIE_LOADER =
    13 struct
    14 
    15 fun log verbose text args x =
    16   if verbose andalso not (null args)
    17   then (Pretty.writeln (Pretty.big_list text (map Pretty.str args)); x)
    18   else x
    19 
    20 val isabelle_name =
    21   let 
    22     fun purge s = if Symbol.is_letter s orelse Symbol.is_digit s then s else
    23       (case s of
    24         "." => "_o_"
    25       | "_" => "_n_"
    26       | "$" => "_S_"
    27       | "@" => "_G_"
    28       | "#" => "_H_"
    29       | "^" => "_T_"
    30       | _   => ("_" ^ string_of_int (ord s) ^ "_"))
    31   in prefix "b_" o translate_string purge end
    32 
    33 fun drop_underscore s =
    34   try (unsuffix "_") s
    35   |> Option.map drop_underscore
    36   |> the_default s
    37 
    38 val short_name =
    39   translate_string (fn s => if Symbol.is_letdig s then s else "") #>
    40   drop_underscore
    41 
    42 (* these prefixes must be distinct: *)
    43 val var_prefix = "V_"
    44 val label_prefix = "L_"
    45 val position_prefix = "P_"
    46 
    47 val no_label_name = label_prefix ^ "unknown"
    48 fun label_name line col =
    49   if line = 0 orelse col = 0 then no_label_name
    50   else label_prefix ^ string_of_int line ^ "_" ^ string_of_int col
    51 
    52 fun mk_syntax name i =
    53   let
    54     val syn = Syntax_Ext.escape name
    55     val args = space_implode ",/ " (replicate i "_")
    56   in
    57     if i = 0 then Mixfix (syn, [], 1000)
    58     else Mixfix (syn ^ "()'(/" ^ args ^ "')", replicate i 0, 1000)
    59   end
    60 
    61 
    62 datatype attribute_value = StringValue of string | TermValue of term
    63 
    64 
    65 
    66 local
    67   fun lookup_type_name thy name arity =
    68     let val intern = Sign.intern_type thy name
    69     in
    70       if Sign.declared_tyname thy intern
    71       then
    72         if Sign.arity_number thy intern = arity then SOME intern
    73         else error ("Boogie: type already declared with different arity: " ^
    74           quote name)
    75       else NONE
    76     end
    77 
    78   fun log_new bname name = bname ^ " (as " ^ name ^ ")"
    79   fun log_ex bname name = "[" ^ bname ^ " has already been declared as " ^
    80     name ^ "]"
    81 
    82   fun declare (name, arity) thy =
    83     let val isa_name = isabelle_name name
    84     in
    85       (case lookup_type_name thy isa_name arity of
    86         SOME type_name => (((name, type_name), log_ex name type_name), thy)
    87       | NONE =>
    88           let
    89             val args = map (rpair dummyS) (Name.invents Name.context "'a" arity)
    90             val (T, thy') =
    91               Typedecl.typedecl_global (Binding.name isa_name, args, mk_syntax name arity) thy
    92             val type_name = fst (Term.dest_Type T)
    93           in (((name, type_name), log_new name type_name), thy') end)
    94     end
    95 in
    96 fun declare_types verbose tys =
    97   fold_map declare tys #>> split_list #-> (fn (tds, logs) =>
    98   log verbose "Declared types:" logs #>
    99   rpair (Symtab.make tds))
   100 end
   101 
   102 
   103 
   104 local
   105   fun maybe_builtin T =
   106     let
   107       fun const name = SOME (Const (name, T))
   108       fun const2_abs name =
   109         let val U = Term.domain_type T
   110         in
   111           SOME (Abs (Name.uu, U, Abs (Name.uu, U,
   112             Const (name, T) $ Bound 0 $ Bound 1)))
   113         end
   114 
   115       fun choose builtin =
   116         (case builtin of
   117           "bvnot" => const @{const_name bitNOT}
   118         | "bvand" => const @{const_name bitAND}
   119         | "bvor" => const @{const_name bitOR}
   120         | "bvxor" => const @{const_name bitXOR}
   121         | "bvadd" => const @{const_name plus}
   122         | "bvneg" => const @{const_name uminus}
   123         | "bvsub" => const @{const_name minus}
   124         | "bvmul" => const @{const_name times}
   125 (* FIXME:
   126         | "bvudiv" => const @{const_name div}
   127         | "bvurem" => const @{const_name mod}
   128         | "bvsdiv" => const @{const_name sdiv}
   129         | "bvsrem" => const @{const_name srem}
   130         | "bvshl" => const @{const_name bv_shl}
   131         | "bvlshr" => const @{const_name bv_lshr}
   132         | "bvashr" => const @{const_name bv_ashr}
   133 *)
   134         | "bvult" => const @{const_name less}
   135         | "bvule" => const @{const_name less_eq}
   136         | "bvugt" => const2_abs @{const_name less}
   137         | "bvuge" => const2_abs @{const_name less_eq}
   138         | "bvslt" => const @{const_name word_sless}
   139         | "bvsle" => const @{const_name word_sle}
   140         | "bvsgt" => const2_abs @{const_name word_sless}
   141         | "bvsge" => const2_abs @{const_name word_sle}
   142         | "zero_extend" => const @{const_name ucast}
   143         | "sign_extend" => const @{const_name scast}
   144         | _ => NONE)
   145 
   146       fun is_builtin att =
   147         (case att of
   148           ("bvbuiltin", [StringValue builtin]) => choose builtin
   149         | ("bvint", [StringValue "ITE"]) => const @{const_name If}
   150         | _ => NONE)
   151     in get_first is_builtin end
   152 
   153   fun lookup_const thy name T =
   154     let val intern = Sign.intern_const thy name
   155     in
   156       if Sign.declared_const thy intern
   157       then
   158         if Sign.typ_instance thy (T, Sign.the_const_type thy intern)
   159         then SOME (Const (intern, T))
   160         else error ("Boogie: function already declared with different type: " ^
   161           quote name)
   162       else NONE
   163     end
   164 
   165   fun log_term thy t = Syntax.string_of_term_global thy t
   166   fun log_new thy name t = name ^ " (as " ^ log_term thy t ^ ")"
   167   fun log_ex thy name t = "[" ^ name ^ " has already been declared as " ^
   168     log_term thy t ^ "]"
   169   fun log_builtin thy name t = "[" ^ name ^ " has been identified as " ^
   170     log_term thy t ^ "]"
   171 
   172   fun declare' name isa_name T arity atts thy =
   173     (case lookup_const thy isa_name T of
   174       SOME t => (((name, t), log_ex thy name t), thy)
   175     | NONE =>
   176         (case maybe_builtin T atts of
   177           SOME t => (((name, t), log_builtin thy name t), thy)
   178         | NONE =>
   179             thy
   180             |> Sign.declare_const_global ((Binding.name isa_name, T),
   181                  mk_syntax name arity)
   182             |> (fn (t, thy') => (((name, t), log_new thy' name t), thy'))))
   183   fun declare (name, ((Ts, T), atts)) =
   184     declare' name (isabelle_name name) (Ts ---> T) (length Ts) atts
   185 
   186   fun uniques fns fds =
   187     let
   188       fun is_unique (name, (([], _), atts)) =
   189             (case AList.lookup (op =) atts "unique" of
   190               SOME _ => Symtab.lookup fds name
   191             | NONE => NONE)
   192         | is_unique _ = NONE
   193       fun mk_unique_axiom T ts =
   194         Const (@{const_name distinct}, HOLogic.listT T --> @{typ bool}) $
   195           HOLogic.mk_list T ts
   196     in
   197       map_filter is_unique fns
   198       |> map (swap o Term.dest_Const)
   199       |> AList.group (op =)
   200       |> map (fn (T, ns) => mk_unique_axiom T (map (Const o rpair T) ns))
   201     end
   202 in
   203 fun declare_functions verbose fns =
   204   fold_map declare fns #>> split_list #-> (fn (fds, logs) =>
   205   log verbose "Loaded constants:" logs #>
   206   rpair (` (uniques fns) (Symtab.make fds)))
   207 end
   208 
   209 
   210 
   211 local
   212   fun name_axioms axs =
   213     let fun mk_name idx = "axiom_" ^ string_of_int (idx + 1)
   214     in map_index (fn (idx, t) => (mk_name idx, HOLogic.mk_Trueprop t)) axs end
   215 
   216   datatype kind = Unused of thm | Used of thm | New of string
   217 
   218   fun mark (name, t) axs =
   219     (case Termtab.lookup axs t of
   220       SOME (Unused thm) => Termtab.update (t, Used thm) axs
   221     | NONE => Termtab.update (t, New name) axs
   222     | SOME _ => axs)
   223 
   224   val sort_fst_str = sort (prod_ord fast_string_ord (K EQUAL)) 
   225   fun split_list_kind thy axs =
   226     let
   227       fun split (_, Used thm) (used, new) = (thm :: used, new)
   228         | split (t, New name) (used, new) = (used, (name, t) :: new)
   229         | split (_, Unused thm) (used, new) =
   230            (warning (Pretty.str_of
   231              (Pretty.big_list "This background axiom has not been loaded:"
   232                [Display.pretty_thm_global thy thm]));
   233             (used, new))
   234     in apsnd sort_fst_str (fold split axs ([], [])) end
   235 
   236   fun mark_axioms thy axs =
   237     Boogie_Axioms.get (Proof_Context.init_global thy)
   238     |> Termtab.make o map (fn thm => (Thm.prop_of thm, Unused thm))
   239     |> fold mark axs
   240     |> split_list_kind thy o Termtab.dest
   241 in
   242 fun add_axioms verbose axs thy =
   243   let
   244     val (used, new) = mark_axioms thy (name_axioms axs)
   245   in
   246     thy
   247     |> fold_map (fn (n, t) => Specification.axiom ((Binding.name n, []), t)) new
   248     |-> Context.theory_map o fold (Boogie_Axioms.add_thm o Drule.export_without_context)
   249     |> log verbose "The following axioms were added:" (map fst new)
   250     |> (fn thy' => log verbose "The following axioms already existed:"
   251          (map (Display.string_of_thm_global thy') used) thy')
   252   end
   253 end
   254 
   255 
   256 
   257 local
   258   fun burrow_distinct eq f xs =
   259     let
   260       val ys = distinct eq xs
   261       val tab = ys ~~ f ys
   262     in map (the o AList.lookup eq tab) xs end
   263 
   264   fun indexed names =
   265     let
   266       val dup = member (op =) (duplicates (op =) (map fst names))
   267       fun make_name (n, i) = n ^ (if dup n then "_" ^ string_of_int i else "")
   268     in map make_name names end
   269 
   270   fun rename idx_names =
   271     idx_names
   272     |> burrow_fst (burrow_distinct (op =)
   273          (map short_name #> ` (duplicates (op =)) #-> Name.variant_list))
   274     |> indexed
   275 in
   276 fun add_vcs verbose vcs thy =
   277   let val vcs' = burrow_fst rename vcs
   278   in
   279     thy
   280     |> Boogie_VCs.set vcs'
   281     |> log verbose "The following verification conditions were loaded:"
   282          (map fst vcs')
   283   end
   284 end
   285 
   286 
   287 
   288 local
   289   fun mk_bitT i T =
   290     if i = 0
   291     then Type (@{type_name "Numeral_Type.bit0"}, [T])
   292     else Type (@{type_name "Numeral_Type.bit1"}, [T])
   293 
   294   fun mk_binT size = 
   295     if size = 0 then @{typ "Numeral_Type.num0"}
   296     else if size = 1 then @{typ "Numeral_Type.num1"}
   297     else let val (q, r) = Integer.div_mod size 2 in mk_bitT r (mk_binT q) end
   298 in
   299 fun mk_wordT size =
   300   if size >= 0 then Type (@{type_name "word"}, [mk_binT size])
   301   else raise TYPE ("mk_wordT: " ^ quote (string_of_int size), [], [])
   302 end
   303 
   304 local
   305   fun dest_binT T =
   306     (case T of
   307       Type (@{type_name "Numeral_Type.num0"}, _) => 0
   308     | Type (@{type_name "Numeral_Type.num1"}, _) => 1
   309     | Type (@{type_name "Numeral_Type.bit0"}, [T]) => 2 * dest_binT T
   310     | Type (@{type_name "Numeral_Type.bit1"}, [T]) => 1 + 2 * dest_binT T
   311     | _ => raise TYPE ("dest_binT", [T], []))
   312 in
   313 val dest_wordT = (fn
   314     Type (@{type_name "word"}, [T]) => dest_binT T
   315   | T => raise TYPE ("dest_wordT", [T], []))
   316 end
   317 
   318 fun mk_arrayT (Ts, T) = Type (@{type_name "fun"}, [HOLogic.mk_tupleT Ts, T])
   319 
   320 
   321 
   322 datatype token = Token of string | Newline | EOF
   323 
   324 fun tokenize path =
   325   let
   326     fun blank c = (c = #" " orelse c = #"\t" orelse c = #"\n" orelse c = #"\r")
   327     fun split line (i, tss) = (i + 1,
   328       map (pair i) (map Token (String.tokens blank line) @ [Newline]) :: tss)
   329   in apsnd (flat o rev) (File.fold_lines split path (1, [])) end
   330 
   331 fun stopper i = Scan.stopper (K (i, EOF)) (fn (_, EOF) => true | _ => false)
   332 
   333 fun scan_err msg [] = "Boogie (at end of input): " ^ msg
   334   | scan_err msg ((i, _) :: _) = "Boogie (at line " ^ string_of_int i ^ "): " ^
   335       msg
   336 
   337 fun scan_fail msg = Scan.fail_with (scan_err msg)
   338 
   339 fun finite scan path =
   340   let val (i, ts) = tokenize path
   341   in
   342     (case Scan.error (Scan.finite (stopper i) scan) ts of
   343       (x, []) => x
   344     | (_, ts') => error (scan_err "unparsed input" ts'))
   345   end
   346 
   347 fun read_int' s = (case read_int (raw_explode s) of (i, []) => SOME i | _ => NONE)
   348 
   349 fun $$$ s = Scan.one (fn (_, Token s') => s = s' | _ => false)
   350 fun str st = Scan.some (fn (_, Token s) => SOME s | _ => NONE) st
   351 fun num st = Scan.some (fn (_, Token s) => read_int' s | _ => NONE) st
   352 
   353 fun scan_line key scan =
   354   $$$ key |-- scan --| Scan.one (fn (_, Newline) => true | _ => false)
   355 fun scan_line' key = scan_line key (Scan.succeed ())
   356 
   357 fun scan_count scan i =
   358   if i > 0 then scan ::: scan_count scan (i - 1) else Scan.succeed []
   359 
   360 fun scan_lookup kind tab key =
   361   (case Symtab.lookup tab key of
   362     SOME value => Scan.succeed value
   363   | NONE => scan_fail ("undefined " ^ kind ^ ": " ^ quote key))
   364 
   365 fun typ tds =
   366   let
   367     fun tp st =
   368      (scan_line' "bool" >> K @{typ bool} ||
   369       scan_line' "int" >> K @{typ int} ||
   370       scan_line "bv" num >> mk_wordT ||
   371       scan_line "type-con" (str -- num) :|-- (fn (name, arity) =>
   372         scan_lookup "type constructor" tds name -- scan_count tp arity >>
   373         Type) ||
   374       scan_line "array" num :|-- (fn arity =>
   375         scan_count tp (arity - 1) -- tp >> mk_arrayT) ||
   376       scan_fail "illegal type") st
   377   in tp end
   378 
   379 local
   380   fun mk_nary _ t [] = t
   381     | mk_nary f _ ts = uncurry (fold_rev f) (split_last ts)
   382 
   383   fun mk_list T = HOLogic.mk_list T
   384 
   385   fun mk_distinct T ts =
   386     Const (@{const_name distinct}, HOLogic.listT T --> @{typ bool}) $
   387       mk_list T ts
   388 
   389   fun quant name f = scan_line name (num -- num -- num) >> pair f
   390   val quants =
   391     quant "forall" HOLogic.all_const ||
   392     quant "exists" HOLogic.exists_const ||
   393     scan_fail "illegal quantifier kind"
   394   fun mk_quant q (n, T) t = q T $ Term.absfree (n, T, t)
   395 
   396   val patternT = @{typ "SMT.pattern"}
   397   fun mk_pattern _ [] = raise TERM ("mk_pattern", [])
   398     | mk_pattern n ts =
   399         let fun mk_pat t = Const (n, Term.fastype_of t --> patternT) $ t
   400         in mk_list patternT (map mk_pat ts) end
   401   fun patt n c scan =
   402     scan_line n num :|-- scan_count scan >> (mk_pattern c)
   403   fun pattern scan =
   404     patt "pat" @{const_name "SMT.pat"} scan ||
   405     patt "nopat" @{const_name "SMT.nopat"} scan ||
   406     scan_fail "illegal pattern kind"
   407   fun mk_trigger [] t = t
   408     | mk_trigger ps t =
   409         @{term "SMT.trigger"} $ mk_list @{typ "SMT.pattern list"} ps $ t
   410 
   411   fun make_label (line, col) = Free (label_name line col, @{typ bool})
   412   fun labelled_by kind pos t = kind $ make_label pos $ t
   413   fun label offset =
   414     $$$ "pos" |-- num -- num >> (fn (line, col) =>
   415       if label_name line col = no_label_name then I
   416       else labelled_by @{term block_at} (line - offset, col)) ||
   417     $$$ "neg" |-- num -- num >> (fn (line, col) =>
   418       labelled_by @{term assert_at} (line - offset, col)) ||
   419     scan_fail "illegal label kind"
   420 
   421   fun mk_store ((m, k), v) =
   422     let
   423       val mT = Term.fastype_of m and kT = Term.fastype_of k
   424       val vT = Term.fastype_of v
   425     in Const (@{const_name fun_upd}, mT --> kT --> vT --> mT) $ m $ k $ v end
   426   
   427   fun mk_extract ((msb, lsb), t) =
   428     let
   429       val dT = Term.fastype_of t and rT = mk_wordT (msb - lsb)
   430       val nT = @{typ nat}
   431       val mk_nat_num = HOLogic.mk_number @{typ nat}
   432     in Const (@{const_name slice}, [nT, dT] ---> rT) $ mk_nat_num lsb $ t end
   433 
   434   fun mk_concat (t1, t2) =
   435     let
   436       val T1 = Term.fastype_of t1 and T2 = Term.fastype_of t2
   437       val U = mk_wordT (dest_wordT T1 + dest_wordT T2)
   438     in Const (@{const_name word_cat}, [T1, T2] ---> U) $ t1 $ t2 end
   439 
   440   fun unique_labels t =
   441     let
   442       fun names_of (@{term assert_at} $ Free (n, _) $ t) = cons n #> names_of t
   443         | names_of (t $ u) = names_of t #> names_of u
   444         | names_of (Abs (_, _, t)) = names_of t
   445         | names_of _ = I
   446       val nctxt = Name.make_context (duplicates (op =) (names_of t []))
   447 
   448       fun fresh (i, nctxt) = (position_prefix ^ string_of_int i, (i+1, nctxt))
   449       fun renamed n (i, nctxt) = Name.variant n nctxt ||> pair i
   450       fun mk_label (name, t) = @{term assert_at} $ Free (name, @{typ bool}) $ t
   451 
   452       fun unique t =
   453         (case t of
   454           @{term assert_at} $ Free (n, _) $ u =>
   455             if n = no_label_name
   456             then fresh ##>> unique u #>> mk_label
   457             else renamed n ##>> unique u #>> mk_label
   458         | u1 $ u2 => unique u1 ##>> unique u2 #>> (op $)
   459         | Abs (n, T, u) => unique u #>> (fn u' => Abs (n, T, u'))
   460         | _ => pair t)
   461     in fst (unique t (1, nctxt)) end
   462 
   463   val var_name = str >> prefix var_prefix
   464   val dest_var_name = unprefix var_prefix
   465   fun rename_variables t =
   466     let
   467       fun short_var_name n = short_name (dest_var_name n)
   468 
   469       val all_names = Term.add_free_names t []
   470       val (names, nctxt) =
   471         all_names
   472         |> map_filter (try (fn n => (n, short_var_name n)))
   473         |> split_list
   474         ||>> (fn names => fold_map Name.variant names (Name.make_context all_names))
   475         |>> Symtab.make o (op ~~)
   476 
   477       fun rename_free n = the_default n (Symtab.lookup names n)
   478       fun rename_abs n = Name.variant (short_var_name n)
   479 
   480       fun rename _ (Free (n, T)) = Free (rename_free n, T)
   481         | rename nctxt (Abs (n, T, t)) =
   482             let val (n', nctxt') = rename_abs n nctxt
   483             in Abs (n', T, rename nctxt' t) end
   484         | rename nctxt (t $ u) = rename nctxt t $ rename nctxt u
   485         | rename _ t = t
   486     in rename nctxt t end
   487 in
   488 fun expr offset tds fds =
   489   let
   490     fun binop t (u1, u2) = t $ u1 $ u2
   491     fun binexp s f = scan_line' s |-- exp -- exp >> f
   492 
   493     and exp st =
   494      (scan_line' "true" >> K @{term True} ||
   495       scan_line' "false" >> K @{term False} ||
   496       scan_line' "not" |-- exp >> HOLogic.mk_not ||
   497       scan_line "and" num :|-- scan_count exp >> 
   498         mk_nary (curry HOLogic.mk_conj) @{term True} ||
   499       scan_line "or" num :|-- scan_count exp >>
   500         mk_nary (curry HOLogic.mk_disj) @{term False} ||
   501       scan_line' "ite" |-- exp -- exp -- exp >> (fn ((c, t1), t2) =>
   502         let val T = Term.fastype_of t1
   503         in
   504           Const (@{const_name If}, [@{typ bool}, T, T] ---> T) $ c $ t1 $ t2
   505         end) ||
   506       binexp "implies" (binop @{term HOL.implies}) ||
   507       scan_line "distinct" num :|-- scan_count exp >>
   508         (fn [] => @{term True}
   509           | ts as (t :: _) => mk_distinct (Term.fastype_of t) ts) ||
   510       binexp "=" HOLogic.mk_eq ||
   511       scan_line "var" var_name -- typ tds >> Free ||
   512       scan_line "fun" (str -- num) :|-- (fn (name, arity) =>
   513         scan_lookup "constant" fds name -- scan_count exp arity >>
   514         Term.list_comb) ||
   515       quants :|-- (fn (q, ((n, k), i)) =>
   516         scan_count (scan_line "var" var_name -- typ tds) n --
   517         scan_count (pattern exp) k --
   518         scan_count (attribute offset tds fds) i --
   519         exp >> (fn (((vs, ps), _), t) =>
   520           fold_rev (mk_quant q) vs (mk_trigger ps t))) ||
   521       scan_line "label" (label offset) -- exp >> (fn (mk, t) => mk t) ||
   522       scan_line "int-num" num >> HOLogic.mk_number @{typ int} ||
   523       binexp "<" (binop @{term "op < :: int => _"}) ||
   524       binexp "<=" (binop @{term "op <= :: int => _"}) ||
   525       binexp ">" (binop @{term "op < :: int => _"} o swap) ||
   526       binexp ">=" (binop @{term "op <= :: int => _"} o swap) ||
   527       binexp "+" (binop @{term "op + :: int => _"}) ||
   528       binexp "-" (binop @{term "op - :: int => _"}) ||
   529       binexp "*" (binop @{term "op * :: int => _"}) ||
   530       binexp "/" (binop @{term boogie_div}) ||
   531       binexp "%" (binop @{term boogie_mod}) ||
   532       scan_line "select" num :|-- (fn arity =>
   533         exp -- (scan_count exp (arity - 1) >> HOLogic.mk_tuple) >> (op $)) ||
   534       scan_line "store" num :|-- (fn arity =>
   535         exp -- (scan_count exp (arity - 2) >> HOLogic.mk_tuple) -- exp >> 
   536           mk_store) ||
   537       scan_line "bv-num" (num -- num) >> (fn (size, i) =>
   538         HOLogic.mk_number (mk_wordT size) i) ||
   539       scan_line "bv-extract" (num -- num) -- exp >> mk_extract ||
   540       binexp "bv-concat" mk_concat ||
   541       scan_fail "illegal expression") st
   542   in exp >> (rename_variables o unique_labels) end
   543 
   544 and attribute offset tds fds =
   545   let
   546     val attr_val = 
   547       scan_line' "expr-attr" |-- expr offset tds fds >> TermValue ||
   548       scan_line "string-attr" (Scan.repeat1 str) >>
   549         (StringValue o space_implode " ") ||
   550       scan_fail "illegal attribute value"
   551   in
   552     scan_line "attribute" (str -- num) :|-- (fn (name, i) =>
   553       scan_count attr_val i >> pair name) ||
   554     scan_fail "illegal attribute"
   555   end
   556 end
   557 
   558 fun type_decls verbose = Scan.depend (fn thy => 
   559   Scan.repeat (scan_line "type-decl" (str -- num -- num) :|-- (fn (ty, i) =>
   560     scan_count (attribute 0 Symtab.empty Symtab.empty) i >> K ty)) >>
   561     (fn tys => declare_types verbose tys thy))
   562 
   563 fun fun_decls verbose tds = Scan.depend (fn thy =>
   564   Scan.repeat (scan_line "fun-decl" (str -- num -- num) :|--
   565     (fn ((name, arity), i) =>
   566       scan_count (typ tds) (arity - 1) -- typ tds --
   567       scan_count (attribute 0 tds Symtab.empty) i >> pair name)) >>
   568     (fn fns => declare_functions verbose fns thy))
   569 
   570 fun axioms verbose tds fds unique_axs = Scan.depend (fn thy =>
   571   Scan.repeat (scan_line "axiom" num :|-- (fn i =>
   572     expr 0 tds fds --| scan_count (attribute 0 tds fds) i)) >>
   573     (fn axs => (add_axioms verbose (unique_axs @ axs) thy, ())))
   574 
   575 fun var_decls tds fds = Scan.depend (fn thy =>
   576   Scan.repeat (scan_line "var-decl" (str -- num) :|-- (fn (_, i) =>
   577     typ tds -- scan_count (attribute 0 tds fds) i >> K ())) >> K (thy, ()))
   578 
   579 fun local_vc_offset offsets vc_name =
   580    Integer.add ~1 (the_default 1 (AList.lookup (op =) offsets vc_name))
   581 
   582 fun vcs verbose offsets tds fds = Scan.depend (fn thy =>
   583   Scan.repeat (scan_line "vc" (str -- num) :-- (fn (name, _) =>
   584     (expr (local_vc_offset offsets name) tds fds))) >> 
   585     (fn vcs => ((), add_vcs verbose vcs thy)))
   586 
   587 fun parse verbose offsets thy = Scan.pass thy
   588  (type_decls verbose :|-- (fn tds =>
   589   fun_decls verbose tds :|-- (fn (unique_axs, fds) =>
   590   axioms verbose tds fds unique_axs |--
   591   var_decls tds fds |--
   592   vcs verbose offsets tds fds)))
   593 
   594 fun load_b2i verbose offsets path thy = finite (parse verbose offsets thy) path
   595 
   596 end