src/HOL/SPARK/Tools/spark_vcs.ML
author berghofe
Mon, 20 Feb 2012 16:09:58 +0100
changeset 47438 8421b6cf2a33
parent 46777 0aaeb5520f2f
child 47595 5759ecd5c905
permissions -rw-r--r--
Fixed bugs:
- set_env no longer modifies pfuns field in theory data record.
Instead, a copy of this field is now contained in the env field.
- add_type_def now checks whether type associated with SPARK enumeration
type is really a datatype with no parameters.
- check_pfuns_types now properly strips off prefixes of proof function names.
     1 (*  Title:      HOL/SPARK/Tools/spark_vcs.ML
     2     Author:     Stefan Berghofer
     3     Copyright:  secunet Security Networks AG
     4 
     5 Store for verification conditions generated by SPARK/Ada.
     6 *)
     7 
     8 signature SPARK_VCS =
     9 sig
    10   val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules ->
    11     (string * string) * Fdl_Parser.vcs -> Path.T -> string -> theory -> theory
    12   val add_proof_fun: (typ option -> 'a -> term) ->
    13     string * ((string list * string) option * 'a) ->
    14     theory -> theory
    15   val add_type: string * typ -> theory -> theory
    16   val lookup_vc: theory -> string -> (Element.context_i list *
    17     (string * thm list option * Element.context_i * Element.statement_i)) option
    18   val get_vcs: theory ->
    19     Element.context_i list * (binding * thm) list * (string *
    20     (string * thm list option * Element.context_i * Element.statement_i)) list
    21   val mark_proved: string -> thm list -> theory -> theory
    22   val close: theory -> theory
    23   val is_closed: theory -> bool
    24 end;
    25 
    26 structure SPARK_VCs: SPARK_VCS =
    27 struct
    28 
    29 open Fdl_Parser;
    30 
    31 
    32 (** theory data **)
    33 
    34 fun err_unfinished () = error "An unfinished SPARK environment is still open."
    35 
    36 val strip_number = pairself implode o take_suffix Fdl_Lexer.is_digit o raw_explode;
    37 
    38 val name_ord = prod_ord string_ord (option_ord int_ord) o
    39   pairself (strip_number ##> Int.fromString);
    40 
    41 structure VCtab = Table(type key = string val ord = name_ord);
    42 
    43 structure VCs = Theory_Data
    44 (
    45   type T =
    46     {pfuns: ((string list * string) option * term) Symtab.table,
    47      type_map: typ Symtab.table,
    48      env:
    49        {ctxt: Element.context_i list,
    50         defs: (binding * thm) list,
    51         types: fdl_type tab,
    52         funs: (string list * string) tab,
    53         pfuns: ((string list * string) option * term) Symtab.table,
    54         ids: (term * string) Symtab.table * Name.context,
    55         proving: bool,
    56         vcs: (string * thm list option *
    57           (string * expr) list * (string * expr) list) VCtab.table,
    58         path: Path.T,
    59         prefix: string} option}
    60   val empty : T = {pfuns = Symtab.empty, type_map = Symtab.empty, env = NONE}
    61   val extend = I
    62   fun merge ({pfuns = pfuns1, type_map = type_map1, env = NONE},
    63         {pfuns = pfuns2, type_map = type_map2, env = NONE}) =
    64         {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2),
    65          type_map = Symtab.merge (op =) (type_map1, type_map2),
    66          env = NONE}
    67     | merge _ = err_unfinished ()
    68 )
    69 
    70 
    71 (** utilities **)
    72 
    73 val to_lower = raw_explode #> map Symbol.to_ascii_lower #> implode;
    74 
    75 val lcase_eq = (op =) o pairself (to_lower o Long_Name.base_name);
    76 
    77 fun lookup_prfx "" tab s = Symtab.lookup tab s
    78   | lookup_prfx prfx tab s = (case Symtab.lookup tab s of
    79         NONE => Symtab.lookup tab (prfx ^ "__" ^ s)
    80       | x => x);
    81 
    82 fun strip_prfx s =
    83   let
    84     fun strip ys [] = ("", implode ys)
    85       | strip ys ("_" :: "_" :: xs) = (implode (rev xs), implode ys)
    86       | strip ys (x :: xs) = strip (x :: ys) xs
    87   in strip [] (rev (raw_explode s)) end;
    88 
    89 fun mk_unop s t =
    90   let val T = fastype_of t
    91   in Const (s, T --> T) $ t end;
    92 
    93 fun mk_times (t, u) =
    94   let
    95     val setT = fastype_of t;
    96     val T = HOLogic.dest_setT setT;
    97     val U = HOLogic.dest_setT (fastype_of u)
    98   in
    99     Const (@{const_name Sigma}, setT --> (T --> HOLogic.mk_setT U) -->
   100       HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u)
   101   end;
   102 
   103 fun get_type thy prfx ty =
   104   let val {type_map, ...} = VCs.get thy
   105   in lookup_prfx prfx type_map ty end;
   106 
   107 fun mk_type _ _ "integer" = HOLogic.intT
   108   | mk_type _ _ "boolean" = HOLogic.boolT
   109   | mk_type thy prfx ty =
   110       (case get_type thy prfx ty of
   111          NONE =>
   112            Syntax.check_typ (Proof_Context.init_global thy)
   113              (Type (Sign.full_name thy (Binding.name ty), []))
   114        | SOME T => T);
   115 
   116 val booleanN = "boolean";
   117 val integerN = "integer";
   118 
   119 fun define_overloaded (def_name, eq) lthy =
   120   let
   121     val ((c, _), rhs) = eq |> Syntax.check_term lthy |>
   122       Logic.dest_equals |>> dest_Free;
   123     val ((_, (_, thm)), lthy') = Local_Theory.define
   124       ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy
   125     val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy');
   126     val thm' = singleton (Proof_Context.export lthy' ctxt_thy) thm
   127   in (thm', lthy') end;
   128 
   129 fun strip_underscores s =
   130   strip_underscores (unsuffix "_" s) handle Fail _ => s;
   131 
   132 fun strip_tilde s =
   133   unsuffix "~" s ^ "_init" handle Fail _ => s;
   134 
   135 val mangle_name = strip_underscores #> strip_tilde;
   136 
   137 fun mk_variables thy prfx xs ty (tab, ctxt) =
   138   let
   139     val T = mk_type thy prfx ty;
   140     val (ys, ctxt') = fold_map Name.variant (map mangle_name xs) ctxt;
   141     val zs = map (Free o rpair T) ys;
   142   in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end;
   143 
   144 fun get_record_info thy T = (case Record.dest_recTs T of
   145     [(tyname, [@{typ unit}])] =>
   146       Record.get_info thy (Long_Name.qualifier tyname)
   147   | _ => NONE);
   148 
   149 fun find_field fname = find_first (curry lcase_eq fname o fst);
   150 
   151 fun find_field' fname = get_first (fn (flds, fldty) =>
   152   if member (op =) flds fname then SOME fldty else NONE);
   153 
   154 fun assoc_ty_err thy T s msg =
   155   error ("Type " ^ Syntax.string_of_typ_global thy T ^
   156     " associated with SPARK type " ^ s ^ "\n" ^ msg);
   157 
   158 
   159 (** generate properties of enumeration types **)
   160 
   161 fun add_enum_type tyname tyname' thy =
   162   let
   163     val {case_name, ...} = the (Datatype.get_info thy tyname');
   164     val cs = map Const (the (Datatype.get_constrs thy tyname'));
   165     val k = length cs;
   166     val T = Type (tyname', []);
   167     val p = Const (@{const_name pos}, T --> HOLogic.intT);
   168     val v = Const (@{const_name val}, HOLogic.intT --> T);
   169     val card = Const (@{const_name card},
   170       HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T;
   171 
   172     fun mk_binrel_def s f = Logic.mk_equals
   173       (Const (s, T --> T --> HOLogic.boolT),
   174        Abs ("x", T, Abs ("y", T,
   175          Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $
   176            (f $ Bound 1) $ (f $ Bound 0))));
   177 
   178     val (((def1, def2), def3), lthy) = thy |>
   179 
   180       Class.instantiation ([tyname'], [], @{sort spark_enum}) |>
   181 
   182       define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals
   183         (p,
   184          list_comb (Const (case_name, replicate k HOLogic.intT @
   185              [T] ---> HOLogic.intT),
   186            map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>>
   187 
   188       define_overloaded ("less_eq_" ^ tyname ^ "_def",
   189         mk_binrel_def @{const_name less_eq} p) ||>>
   190       define_overloaded ("less_" ^ tyname ^ "_def",
   191         mk_binrel_def @{const_name less} p);
   192 
   193     val UNIV_eq = Goal.prove lthy [] []
   194       (HOLogic.mk_Trueprop (HOLogic.mk_eq
   195          (HOLogic.mk_UNIV T, HOLogic.mk_set T cs)))
   196       (fn _ =>
   197          rtac @{thm subset_antisym} 1 THEN
   198          rtac @{thm subsetI} 1 THEN
   199          Datatype_Aux.exh_tac (K (#exhaust (Datatype.the_info
   200            (Proof_Context.theory_of lthy) tyname'))) 1 THEN
   201          ALLGOALS (asm_full_simp_tac (simpset_of lthy)));
   202 
   203     val finite_UNIV = Goal.prove lthy [] []
   204       (HOLogic.mk_Trueprop (Const (@{const_name finite},
   205          HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T))
   206       (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1);
   207 
   208     val card_UNIV = Goal.prove lthy [] []
   209       (HOLogic.mk_Trueprop (HOLogic.mk_eq
   210          (card, HOLogic.mk_number HOLogic.natT k)))
   211       (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1);
   212 
   213     val range_pos = Goal.prove lthy [] []
   214       (HOLogic.mk_Trueprop (HOLogic.mk_eq
   215          (Const (@{const_name image}, (T --> HOLogic.intT) -->
   216             HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $
   217               p $ HOLogic.mk_UNIV T,
   218           Const (@{const_name atLeastLessThan}, HOLogic.intT -->
   219             HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $
   220               HOLogic.mk_number HOLogic.intT 0 $
   221               (@{term int} $ card))))
   222       (fn _ =>
   223          simp_tac (simpset_of lthy addsimps [card_UNIV]) 1 THEN
   224          simp_tac (simpset_of lthy addsimps [UNIV_eq, def1]) 1 THEN
   225          rtac @{thm subset_antisym} 1 THEN
   226          simp_tac (simpset_of lthy) 1 THEN
   227          rtac @{thm subsetI} 1 THEN
   228          asm_full_simp_tac (simpset_of lthy addsimps @{thms interval_expand}
   229            delsimps @{thms atLeastLessThan_iff}) 1);
   230 
   231     val lthy' =
   232       Class.prove_instantiation_instance (fn _ =>
   233         Class.intro_classes_tac [] THEN
   234         rtac finite_UNIV 1 THEN
   235         rtac range_pos 1 THEN
   236         simp_tac (HOL_basic_ss addsimps [def3]) 1 THEN
   237         simp_tac (HOL_basic_ss addsimps [def2]) 1) lthy;
   238 
   239     val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) =>
   240       let
   241         val n = HOLogic.mk_number HOLogic.intT i;
   242         val th = Goal.prove lthy' [] []
   243           (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n)))
   244           (fn _ => simp_tac (simpset_of lthy' addsimps [def1]) 1);
   245         val th' = Goal.prove lthy' [] []
   246           (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c)))
   247           (fn _ =>
   248              rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN
   249              simp_tac (simpset_of lthy' addsimps
   250                [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
   251       in (th, th') end) cs);
   252 
   253     val first_el = Goal.prove lthy' [] []
   254       (HOLogic.mk_Trueprop (HOLogic.mk_eq
   255          (Const (@{const_name first_el}, T), hd cs)))
   256       (fn _ => simp_tac (simpset_of lthy' addsimps
   257          [@{thm first_el_def}, hd val_eqs]) 1);
   258 
   259     val last_el = Goal.prove lthy' [] []
   260       (HOLogic.mk_Trueprop (HOLogic.mk_eq
   261          (Const (@{const_name last_el}, T), List.last cs)))
   262       (fn _ => simp_tac (simpset_of lthy' addsimps
   263          [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
   264   in
   265     lthy' |>
   266     Local_Theory.note
   267       ((Binding.name (tyname ^ "_card"), @{attributes [simp]}), [card_UNIV]) ||>>
   268     Local_Theory.note
   269       ((Binding.name (tyname ^ "_pos"), @{attributes [simp]}), pos_eqs) ||>>
   270     Local_Theory.note
   271       ((Binding.name (tyname ^ "_val"), @{attributes [simp]}), val_eqs) ||>>
   272     Local_Theory.note
   273       ((Binding.name (tyname ^ "_first_el"), @{attributes [simp]}), [first_el]) ||>>
   274     Local_Theory.note
   275       ((Binding.name (tyname ^ "_last_el"), @{attributes [simp]}), [last_el]) |> snd |>
   276     Local_Theory.exit_global
   277   end;
   278 
   279 
   280 fun check_no_assoc thy prfx s = case get_type thy prfx s of
   281     NONE => ()
   282   | SOME _ => error ("Cannot associate a type with " ^ s ^
   283       "\nsince it is no record or enumeration type");
   284 
   285 fun check_enum [] [] = NONE 
   286   | check_enum els [] = SOME ("has no element(s) " ^ commas els)
   287   | check_enum [] cs = SOME ("has extra element(s) " ^
   288       commas (map (Long_Name.base_name o fst) cs))
   289   | check_enum (el :: els) ((cname, _) :: cs) =
   290       if lcase_eq (el, cname) then check_enum els cs
   291       else SOME ("either has no element " ^ el ^
   292         " or it is at the wrong position");
   293 
   294 fun add_type_def prfx (s, Basic_Type ty) (ids, thy) =
   295       (check_no_assoc thy prfx s;
   296        (ids,
   297         Typedecl.abbrev_global (Binding.name s, [], NoSyn)
   298           (mk_type thy prfx ty) thy |> snd))
   299 
   300   | add_type_def prfx (s, Enum_Type els) ((tab, ctxt), thy) =
   301       let
   302         val (thy', tyname) = (case get_type thy prfx s of
   303             NONE =>
   304               let
   305                 val tyb = Binding.name s;
   306                 val tyname = Sign.full_name thy tyb
   307               in
   308                 (thy |>
   309                  Datatype.add_datatype {strict = true, quiet = true}
   310                    [((tyb, [], NoSyn), map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
   311                  add_enum_type s tyname,
   312                  tyname)
   313               end
   314           | SOME (T as Type (tyname, [])) =>
   315               (case Datatype.get_constrs thy tyname of
   316                  NONE => assoc_ty_err thy T s "is not a datatype"
   317                | SOME cs =>
   318                    let
   319                      val (prfx', _) = strip_prfx s;
   320                      val els' =
   321                        if prfx' = "" then els
   322                        else map (unprefix (prfx' ^ "__")) els
   323                          handle Fail _ => error ("Bad enumeration type " ^ s)
   324                    in
   325                      case check_enum els' cs of
   326                        NONE => (thy, tyname)
   327                      | SOME msg => assoc_ty_err thy T s msg
   328                    end)
   329           | SOME T => assoc_ty_err thy T s "is not a datatype");
   330         val cs = map Const (the (Datatype.get_constrs thy' tyname));
   331       in
   332         ((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab,
   333           fold Name.declare els ctxt),
   334          thy')
   335       end
   336 
   337   | add_type_def prfx (s, Array_Type (argtys, resty)) (ids, thy) =
   338       (check_no_assoc thy prfx s;
   339        (ids,
   340         Typedecl.abbrev_global (Binding.name s, [], NoSyn)
   341           (foldr1 HOLogic.mk_prodT (map (mk_type thy prfx) argtys) -->
   342              mk_type thy prfx resty) thy |> snd))
   343 
   344   | add_type_def prfx (s, Record_Type fldtys) (ids, thy) =
   345       (ids,
   346        let val fldTs = maps (fn (flds, ty) =>
   347          map (rpair (mk_type thy prfx ty)) flds) fldtys
   348        in case get_type thy prfx s of
   349            NONE =>
   350              Record.add_record ([], Binding.name s) NONE
   351                (map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy
   352          | SOME rT =>
   353              (case get_record_info thy rT of
   354                 NONE => assoc_ty_err thy rT s "is not a record type"
   355               | SOME {fields, ...} =>
   356                   (case subtract (lcase_eq o pairself fst) fldTs fields of
   357                      [] => ()
   358                    | flds => assoc_ty_err thy rT s ("has extra field(s) " ^
   359                        commas (map (Long_Name.base_name o fst) flds));
   360                    map (fn (fld, T) =>
   361                      case AList.lookup lcase_eq fields fld of
   362                        NONE => assoc_ty_err thy rT s ("has no field " ^ fld)
   363                      | SOME U => T = U orelse assoc_ty_err thy rT s
   364                          ("has field " ^
   365                           fld ^ " whose type\n" ^
   366                           Syntax.string_of_typ_global thy U ^
   367                           "\ndoes not match declared type\n" ^
   368                           Syntax.string_of_typ_global thy T)) fldTs;
   369                    thy))
   370        end)
   371 
   372   | add_type_def prfx (s, Pending_Type) (ids, thy) =
   373       (ids,
   374        case get_type thy prfx s of
   375          SOME _ => thy
   376        | NONE => Typedecl.typedecl_global
   377            (Binding.name s, [], NoSyn) thy |> snd);
   378 
   379 
   380 fun term_of_expr thy prfx types pfuns =
   381   let
   382     fun tm_of vs (Funct ("->", [e, e'])) =
   383           (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   384 
   385       | tm_of vs (Funct ("<->", [e, e'])) =
   386           (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   387 
   388       | tm_of vs (Funct ("or", [e, e'])) =
   389           (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   390 
   391       | tm_of vs (Funct ("and", [e, e'])) =
   392           (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   393 
   394       | tm_of vs (Funct ("not", [e])) =
   395           (HOLogic.mk_not (fst (tm_of vs e)), booleanN)
   396 
   397       | tm_of vs (Funct ("=", [e, e'])) =
   398           (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   399 
   400       | tm_of vs (Funct ("<>", [e, e'])) = (HOLogic.mk_not
   401           (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN)
   402 
   403       | tm_of vs (Funct ("<", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
   404           (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   405 
   406       | tm_of vs (Funct (">", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
   407           (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
   408 
   409       | tm_of vs (Funct ("<=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
   410           (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
   411 
   412       | tm_of vs (Funct (">=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
   413           (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
   414 
   415       | tm_of vs (Funct ("+", [e, e'])) = (HOLogic.mk_binop @{const_name plus}
   416           (fst (tm_of vs e), fst (tm_of vs e')), integerN)
   417 
   418       | tm_of vs (Funct ("-", [e, e'])) = (HOLogic.mk_binop @{const_name minus}
   419           (fst (tm_of vs e), fst (tm_of vs e')), integerN)
   420 
   421       | tm_of vs (Funct ("*", [e, e'])) = (HOLogic.mk_binop @{const_name times}
   422           (fst (tm_of vs e), fst (tm_of vs e')), integerN)
   423 
   424       | tm_of vs (Funct ("/", [e, e'])) = (HOLogic.mk_binop @{const_name divide}
   425           (fst (tm_of vs e), fst (tm_of vs e')), integerN)
   426 
   427       | tm_of vs (Funct ("div", [e, e'])) = (HOLogic.mk_binop @{const_name sdiv}
   428           (fst (tm_of vs e), fst (tm_of vs e')), integerN)
   429 
   430       | tm_of vs (Funct ("mod", [e, e'])) = (HOLogic.mk_binop @{const_name mod}
   431           (fst (tm_of vs e), fst (tm_of vs e')), integerN)
   432 
   433       | tm_of vs (Funct ("-", [e])) =
   434           (mk_unop @{const_name uminus} (fst (tm_of vs e)), integerN)
   435 
   436       | tm_of vs (Funct ("**", [e, e'])) =
   437           (Const (@{const_name power}, HOLogic.intT --> HOLogic.natT -->
   438              HOLogic.intT) $ fst (tm_of vs e) $
   439                (@{const nat} $ fst (tm_of vs e')), integerN)
   440 
   441       | tm_of (tab, _) (Ident s) =
   442           (case Symtab.lookup tab s of
   443              SOME t_ty => t_ty
   444            | NONE => (case lookup_prfx prfx pfuns s of
   445                SOME (SOME ([], resty), t) => (t, resty)
   446              | _ => error ("Undeclared identifier " ^ s)))
   447 
   448       | tm_of _ (Number i) = (HOLogic.mk_number HOLogic.intT i, integerN)
   449 
   450       | tm_of vs (Quantifier (s, xs, ty, e)) =
   451           let
   452             val (ys, vs') = mk_variables thy prfx xs ty vs;
   453             val q = (case s of
   454                 "for_all" => HOLogic.mk_all
   455               | "for_some" => HOLogic.mk_exists)
   456           in
   457             (fold_rev (fn Free (x, T) => fn t => q (x, T, t))
   458                ys (fst (tm_of vs' e)),
   459              booleanN)
   460           end
   461 
   462       | tm_of vs (Funct (s, es)) =
   463 
   464           (* record field selection *)
   465           (case try (unprefix "fld_") s of
   466              SOME fname => (case es of
   467                [e] =>
   468                  let
   469                    val (t, rcdty) = tm_of vs e;
   470                    val rT = mk_type thy prfx rcdty
   471                  in case (get_record_info thy rT, lookup types rcdty) of
   472                      (SOME {fields, ...}, SOME (Record_Type fldtys)) =>
   473                        (case (find_field fname fields,
   474                             find_field' fname fldtys) of
   475                           (SOME (fname', fT), SOME fldty) =>
   476                             (Const (fname', rT --> fT) $ t, fldty)
   477                         | _ => error ("Record " ^ rcdty ^
   478                             " has no field named " ^ fname))
   479                    | _ => error (rcdty ^ " is not a record type")
   480                  end
   481              | _ => error ("Function " ^ s ^ " expects one argument"))
   482            | NONE =>
   483 
   484           (* record field update *)
   485           (case try (unprefix "upf_") s of
   486              SOME fname => (case es of
   487                [e, e'] =>
   488                  let
   489                    val (t, rcdty) = tm_of vs e;
   490                    val rT = mk_type thy prfx rcdty;
   491                    val (u, fldty) = tm_of vs e';
   492                    val fT = mk_type thy prfx fldty
   493                  in case get_record_info thy rT of
   494                      SOME {fields, ...} =>
   495                        (case find_field fname fields of
   496                           SOME (fname', fU) =>
   497                             if fT = fU then
   498                               (Const (fname' ^ "_update",
   499                                  (fT --> fT) --> rT --> rT) $
   500                                    Abs ("x", fT, u) $ t,
   501                                rcdty)
   502                             else error ("Type\n" ^
   503                               Syntax.string_of_typ_global thy fT ^
   504                               "\ndoes not match type\n" ^
   505                               Syntax.string_of_typ_global thy fU ^
   506                               "\nof field " ^ fname)
   507                         | NONE => error ("Record " ^ rcdty ^
   508                             " has no field named " ^ fname))
   509                    | _ => error (rcdty ^ " is not a record type")
   510                  end
   511              | _ => error ("Function " ^ s ^ " expects two arguments"))
   512            | NONE =>
   513 
   514           (* enumeration type to integer *)
   515           (case try (unsuffix "__pos") s of
   516              SOME tyname => (case es of
   517                [e] => (Const (@{const_name pos},
   518                    mk_type thy prfx tyname --> HOLogic.intT) $ fst (tm_of vs e),
   519                  integerN)
   520              | _ => error ("Function " ^ s ^ " expects one argument"))
   521            | NONE =>
   522 
   523           (* integer to enumeration type *)
   524           (case try (unsuffix "__val") s of
   525              SOME tyname => (case es of
   526                [e] => (Const (@{const_name val},
   527                    HOLogic.intT --> mk_type thy prfx tyname) $ fst (tm_of vs e),
   528                  tyname)
   529              | _ => error ("Function " ^ s ^ " expects one argument"))
   530            | NONE =>
   531 
   532           (* successor / predecessor of enumeration type element *)
   533           if s = "succ" orelse s = "pred" then (case es of
   534               [e] =>
   535                 let
   536                   val (t, tyname) = tm_of vs e;
   537                   val T = mk_type thy prfx tyname
   538                 in (Const
   539                   (if s = "succ" then @{const_name succ}
   540                    else @{const_name pred}, T --> T) $ t, tyname)
   541                 end
   542             | _ => error ("Function " ^ s ^ " expects one argument"))
   543 
   544           (* user-defined proof function *)
   545           else
   546             (case lookup_prfx prfx pfuns s of
   547                SOME (SOME (_, resty), t) =>
   548                  (list_comb (t, map (fst o tm_of vs) es), resty)
   549              | _ => error ("Undeclared proof function " ^ s))))))
   550 
   551       | tm_of vs (Element (e, es)) =
   552           let val (t, ty) = tm_of vs e
   553           in case lookup types ty of
   554               SOME (Array_Type (_, elty)) =>
   555                 (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty)
   556             | _ => error (ty ^ " is not an array type")
   557           end
   558 
   559       | tm_of vs (Update (e, es, e')) =
   560           let val (t, ty) = tm_of vs e
   561           in case lookup types ty of
   562               SOME (Array_Type (idxtys, elty)) =>
   563                 let
   564                   val T = foldr1 HOLogic.mk_prodT
   565                     (map (mk_type thy prfx) idxtys);
   566                   val U = mk_type thy prfx elty;
   567                   val fT = T --> U
   568                 in
   569                   (Const (@{const_name fun_upd}, fT --> T --> U --> fT) $
   570                      t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $
   571                        fst (tm_of vs e'),
   572                    ty)
   573                 end
   574             | _ => error (ty ^ " is not an array type")
   575           end
   576 
   577       | tm_of vs (Record (s, flds)) =
   578           let
   579             val T = mk_type thy prfx s;
   580             val {extension = (ext_name, _), fields, ...} =
   581               (case get_record_info thy T of
   582                  NONE => error (s ^ " is not a record type")
   583                | SOME info => info);
   584             val flds' = map (apsnd (tm_of vs)) flds;
   585             val fnames = map (Long_Name.base_name o fst) fields;
   586             val fnames' = map fst flds;
   587             val (fvals, ftys) = split_list (map (fn s' =>
   588               case AList.lookup lcase_eq flds' s' of
   589                 SOME fval_ty => fval_ty
   590               | NONE => error ("Field " ^ s' ^ " missing in record " ^ s))
   591                   fnames);
   592             val _ = (case subtract lcase_eq fnames fnames' of
   593                 [] => ()
   594               | xs => error ("Extra field(s) " ^ commas xs ^
   595                   " in record " ^ s));
   596             val _ = (case duplicates (op =) fnames' of
   597                 [] => ()
   598               | xs => error ("Duplicate field(s) " ^ commas xs ^
   599                   " in record " ^ s))
   600           in
   601             (list_comb
   602                (Const (ext_name,
   603                   map (mk_type thy prfx) ftys @ [HOLogic.unitT] ---> T),
   604                 fvals @ [HOLogic.unit]),
   605              s)
   606           end
   607 
   608       | tm_of vs (Array (s, default, assocs)) =
   609           (case lookup types s of
   610              SOME (Array_Type (idxtys, elty)) =>
   611                let
   612                  val Ts = map (mk_type thy prfx) idxtys;
   613                  val T = foldr1 HOLogic.mk_prodT Ts;
   614                  val U = mk_type thy prfx elty;
   615                  fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)]
   616                    | mk_idx' T (e, SOME e') = Const (@{const_name atLeastAtMost},
   617                        T --> T --> HOLogic.mk_setT T) $
   618                          fst (tm_of vs e) $ fst (tm_of vs e');
   619                  fun mk_idx idx =
   620                    if length Ts <> length idx then
   621                      error ("Arity mismatch in construction of array " ^ s)
   622                    else foldr1 mk_times (map2 mk_idx' Ts idx);
   623                  fun mk_upd (idxs, e) t =
   624                    if length idxs = 1 andalso forall (is_none o snd) (hd idxs)
   625                    then
   626                      Const (@{const_name fun_upd}, (T --> U) -->
   627                          T --> U --> T --> U) $ t $
   628                        foldl1 HOLogic.mk_prod
   629                          (map (fst o tm_of vs o fst) (hd idxs)) $
   630                        fst (tm_of vs e)
   631                    else
   632                      Const (@{const_name fun_upds}, (T --> U) -->
   633                          HOLogic.mk_setT T --> U --> T --> U) $ t $
   634                        foldl1 (HOLogic.mk_binop @{const_name sup})
   635                          (map mk_idx idxs) $
   636                        fst (tm_of vs e)
   637                in
   638                  (fold mk_upd assocs
   639                     (case default of
   640                        SOME e => Abs ("x", T, fst (tm_of vs e))
   641                      | NONE => Const (@{const_name undefined}, T --> U)),
   642                   s)
   643                end
   644            | _ => error (s ^ " is not an array type"))
   645 
   646   in tm_of end;
   647 
   648 
   649 fun term_of_rule thy prfx types pfuns ids rule =
   650   let val tm_of = fst o term_of_expr thy prfx types pfuns ids
   651   in case rule of
   652       Inference_Rule (es, e) => Logic.list_implies
   653         (map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e))
   654     | Substitution_Rule (es, e, e') => Logic.list_implies
   655         (map (HOLogic.mk_Trueprop o tm_of) es,
   656          HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e')))
   657   end;
   658 
   659 
   660 val builtin = Symtab.make (map (rpair ())
   661   ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=",
   662    "+", "-", "*", "/", "div", "mod", "**"]);
   663 
   664 fun complex_expr (Number _) = false
   665   | complex_expr (Ident _) = false 
   666   | complex_expr (Funct (s, es)) =
   667       not (Symtab.defined builtin s) orelse exists complex_expr es
   668   | complex_expr (Quantifier (_, _, _, e)) = complex_expr e
   669   | complex_expr _ = true;
   670 
   671 fun complex_rule (Inference_Rule (es, e)) =
   672       complex_expr e orelse exists complex_expr es
   673   | complex_rule (Substitution_Rule (es, e, e')) =
   674       complex_expr e orelse complex_expr e' orelse
   675       exists complex_expr es;
   676 
   677 val is_pfun =
   678   Symtab.defined builtin orf
   679   can (unprefix "fld_") orf can (unprefix "upf_") orf
   680   can (unsuffix "__pos") orf can (unsuffix "__val") orf
   681   equal "succ" orf equal "pred";
   682 
   683 fun fold_opt f = the_default I o Option.map f;
   684 fun fold_pair f g (x, y) = f x #> g y;
   685 
   686 fun fold_expr f g (Funct (s, es)) = f s #> fold (fold_expr f g) es
   687   | fold_expr f g (Ident s) = g s
   688   | fold_expr f g (Number _) = I
   689   | fold_expr f g (Quantifier (_, _, _, e)) = fold_expr f g e
   690   | fold_expr f g (Element (e, es)) =
   691       fold_expr f g e #> fold (fold_expr f g) es
   692   | fold_expr f g (Update (e, es, e')) =
   693       fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e'
   694   | fold_expr f g (Record (_, flds)) = fold (fold_expr f g o snd) flds
   695   | fold_expr f g (Array (_, default, assocs)) =
   696       fold_opt (fold_expr f g) default #>
   697       fold (fold_pair
   698         (fold (fold (fold_pair
   699           (fold_expr f g) (fold_opt (fold_expr f g)))))
   700         (fold_expr f g)) assocs;
   701 
   702 fun add_expr_pfuns funs = fold_expr
   703   (fn s => if is_pfun s then I else insert (op =) s)
   704   (fn s => if is_none (lookup funs s) then I else insert (op =) s);
   705 
   706 val add_expr_idents = fold_expr (K I) (insert (op =));
   707 
   708 fun pfun_type thy prfx (argtys, resty) =
   709   map (mk_type thy prfx) argtys ---> mk_type thy prfx resty;
   710 
   711 fun check_pfun_type thy prfx s t optty1 optty2 =
   712   let
   713     val T = fastype_of t;
   714     fun check ty =
   715       let val U = pfun_type thy prfx ty
   716       in
   717         T = U orelse
   718         error ("Type\n" ^
   719           Syntax.string_of_typ_global thy T ^
   720           "\nof function " ^
   721           Syntax.string_of_term_global thy t ^
   722           " associated with proof function " ^ s ^
   723           "\ndoes not match declared type\n" ^
   724           Syntax.string_of_typ_global thy U)
   725       end
   726   in (Option.map check optty1; Option.map check optty2; ()) end;
   727 
   728 fun upd_option x y = if is_some x then x else y;
   729 
   730 fun unprefix_pfun "" s = s
   731   | unprefix_pfun prfx s =
   732       let val (prfx', s') = strip_prfx s
   733       in if prfx = prfx' then s' else s end;
   734 
   735 fun check_pfuns_types thy prfx funs =
   736   Symtab.map (fn s => fn (optty, t) =>
   737    let val optty' = lookup funs (unprefix_pfun prfx s)
   738    in
   739      (check_pfun_type thy prfx s t optty optty';
   740       (NONE |> upd_option optty |> upd_option optty', t))
   741    end);
   742 
   743 
   744 (** the VC store **)
   745 
   746 fun err_vcs names = error (Pretty.string_of
   747   (Pretty.big_list "The following verification conditions have not been proved:"
   748     (map Pretty.str names)))
   749 
   750 fun set_env ctxt defs types funs ids vcs path prefix thy = VCs.map (fn
   751     {pfuns, type_map, env = NONE} =>
   752       {pfuns = pfuns,
   753        type_map = type_map,
   754        env = SOME
   755          {ctxt = ctxt, defs = defs, types = types, funs = funs,
   756           pfuns = check_pfuns_types thy prefix funs pfuns,
   757           ids = ids, proving = false, vcs = vcs, path = path,
   758           prefix = prefix}}
   759   | _ => err_unfinished ()) thy;
   760 
   761 fun mk_pat s = (case Int.fromString s of
   762     SOME i => [HOLogic.mk_Trueprop (Var (("C", i), HOLogic.boolT))]
   763   | NONE => error ("Bad conclusion identifier: C" ^ s));
   764 
   765 fun mk_vc thy prfx types pfuns ids (tr, proved, ps, cs) =
   766   let val prop_of =
   767     HOLogic.mk_Trueprop o fst o term_of_expr thy prfx types pfuns ids
   768   in
   769     (tr, proved,
   770      Element.Assumes (map (fn (s', e) =>
   771        ((Binding.name ("H" ^ s'), []), [(prop_of e, [])])) ps),
   772      Element.Shows (map (fn (s', e) =>
   773        (Attrib.empty_binding, [(prop_of e, mk_pat s')])) cs))
   774   end;
   775 
   776 fun fold_vcs f vcs =
   777   VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs;
   778 
   779 fun pfuns_of_vcs prfx funs pfuns vcs =
   780   fold_vcs (add_expr_pfuns funs o snd) vcs [] |>
   781   filter (is_none o lookup_prfx prfx pfuns);
   782 
   783 fun declare_missing_pfuns thy prfx funs pfuns vcs (tab, ctxt) =
   784   let
   785     val (fs, (tys, Ts)) =
   786       pfuns_of_vcs prfx funs pfuns vcs |>
   787       map_filter (fn s => lookup funs s |>
   788         Option.map (fn ty => (s, (SOME ty, pfun_type thy prfx ty)))) |>
   789       split_list ||> split_list;
   790     val (fs', ctxt') = fold_map Name.variant fs ctxt
   791   in
   792     (fold Symtab.update_new (fs ~~ (tys ~~ map Free (fs' ~~ Ts))) pfuns,
   793      Element.Fixes (map2 (fn s => fn T =>
   794        (Binding.name s, SOME T, NoSyn)) fs' Ts),
   795      (tab, ctxt'))
   796   end;
   797 
   798 fun map_pfuns f {pfuns, type_map, env} =
   799   {pfuns = f pfuns, type_map = type_map, env = env}
   800 
   801 fun map_pfuns_env f {pfuns, type_map,
   802       env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env,
   803         ids, proving, vcs, path, prefix}} =
   804   if proving then err_unfinished ()
   805   else
   806     {pfuns = pfuns, type_map = type_map,
   807      env = SOME {ctxt = ctxt, defs = defs, types = types, funs = funs,
   808        pfuns = f pfuns_env, ids = ids, proving = false, vcs = vcs,
   809        path = path, prefix = prefix}};
   810 
   811 fun add_proof_fun prep (s, (optty, raw_t)) thy =
   812   VCs.map (fn data as {env, ...} =>
   813     let
   814       val (optty', prfx, map_pf) = (case env of
   815           SOME {funs, prefix, ...} =>
   816             (lookup funs (unprefix_pfun prefix s),
   817              prefix, map_pfuns_env)
   818         | NONE => (NONE, "", map_pfuns));
   819       val optty'' = NONE |> upd_option optty |> upd_option optty';
   820       val t = prep (Option.map (pfun_type thy prfx) optty'') raw_t;
   821       val _ = (case fold_aterms (fn u =>
   822           if is_Var u orelse is_Free u then insert (op =) u else I) t [] of
   823           [] => ()
   824         | ts => error ("Term\n" ^ Syntax.string_of_term_global thy t ^
   825             "\nto be associated with proof function " ^ s ^
   826             " contains free variable(s) " ^
   827             commas (map (Syntax.string_of_term_global thy) ts)));
   828     in
   829       (check_pfun_type thy prfx s t optty optty';
   830        if is_some optty'' orelse is_none env then
   831          map_pf (Symtab.update_new (s, (optty'', t))) data
   832            handle Symtab.DUP _ => error ("Proof function " ^ s ^
   833              " already associated with function")
   834        else error ("Undeclared proof function " ^ s))
   835     end) thy;
   836 
   837 fun add_type (s, T as Type (tyname, Ts)) thy =
   838       thy |>
   839       VCs.map (fn
   840           {env = SOME _, ...} => err_unfinished ()
   841         | {pfuns, type_map, env} =>
   842             {pfuns = pfuns,
   843              type_map = Symtab.update_new (s, T) type_map,
   844              env = env}
   845               handle Symtab.DUP _ => error ("SPARK type " ^ s ^
   846                 " already associated with type")) |>
   847       (fn thy' =>
   848          case Datatype.get_constrs thy' tyname of
   849            NONE => thy'
   850          | SOME cs =>
   851              if null Ts then
   852                (map
   853                   (fn (_, Type (_, [])) => ()
   854                     | (cname, _) => assoc_ty_err thy T s
   855                         ("has illegal constructor " ^
   856                          Long_Name.base_name cname)) cs;
   857                 add_enum_type s tyname thy')
   858              else assoc_ty_err thy T s "is illegal")
   859   | add_type (s, T) thy = assoc_ty_err thy T s "is illegal";
   860 
   861 val is_closed = is_none o #env o VCs.get;
   862 
   863 fun lookup_vc thy name =
   864   (case VCs.get thy of
   865     {env = SOME {vcs, types, funs, pfuns, ids, ctxt, prefix, ...}, ...} =>
   866       (case VCtab.lookup vcs name of
   867          SOME vc =>
   868            let val (pfuns', ctxt', ids') =
   869              declare_missing_pfuns thy prefix funs pfuns vcs ids
   870            in SOME (ctxt @ [ctxt'], mk_vc thy prefix types pfuns' ids' vc) end
   871        | NONE => NONE)
   872   | _ => NONE);
   873 
   874 fun get_vcs thy = (case VCs.get thy of
   875     {env = SOME {vcs, types, funs, pfuns, ids, ctxt, defs, prefix, ...}, ...} =>
   876       let val (pfuns', ctxt', ids') =
   877         declare_missing_pfuns thy prefix funs pfuns vcs ids
   878       in
   879         (ctxt @ [ctxt'], defs,
   880          VCtab.dest vcs |>
   881          map (apsnd (mk_vc thy prefix types pfuns' ids')))
   882       end
   883   | _ => ([], [], []));
   884 
   885 fun mark_proved name thms = VCs.map (fn
   886     {pfuns, type_map,
   887      env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env,
   888         ids, vcs, path, prefix, ...}} =>
   889       {pfuns = pfuns,
   890        type_map = type_map,
   891        env = SOME {ctxt = ctxt, defs = defs,
   892          types = types, funs = funs, pfuns = pfuns_env,
   893          ids = ids,
   894          proving = true,
   895          vcs = VCtab.map_entry name (fn (trace, _, ps, cs) =>
   896            (trace, SOME thms, ps, cs)) vcs,
   897          path = path,
   898          prefix = prefix}}
   899   | x => x);
   900 
   901 fun close thy =
   902   thy |>
   903   VCs.map (fn
   904       {pfuns, type_map, env = SOME {vcs, path, ...}} =>
   905         (case VCtab.fold_rev (fn vc as (_, (_, p, _, _)) =>
   906              (if is_some p then apfst else apsnd) (cons vc)) vcs ([], []) of
   907            (proved, []) =>
   908              (Thm.join_proofs (maps (the o #2 o snd) proved);
   909               File.write (Path.ext "prv" path)
   910                 (implode (map (fn (s, _) => snd (strip_number s) ^
   911                    " -- proved by " ^ Distribution.version ^ "\n") proved));
   912               {pfuns = pfuns, type_map = type_map, env = NONE})
   913          | (_, unproved) => err_vcs (map fst unproved))
   914     | _ => error "No SPARK environment is currently open") |>
   915   Sign.parent_path;
   916 
   917 
   918 (** set up verification conditions **)
   919 
   920 fun partition_opt f =
   921   let
   922     fun part ys zs [] = (rev ys, rev zs)
   923       | part ys zs (x :: xs) = (case f x of
   924             SOME y => part (y :: ys) zs xs
   925           | NONE => part ys (x :: zs) xs)
   926   in part [] [] end;
   927 
   928 fun dest_def (id, (Substitution_Rule ([], Ident s, rhs))) = SOME (id, (s, rhs))
   929   | dest_def _ = NONE;
   930 
   931 fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i);
   932 
   933 fun add_const prfx (s, ty) ((tab, ctxt), thy) =
   934   let
   935     val T = mk_type thy prfx ty;
   936     val b = Binding.name s;
   937     val c = Const (Sign.full_name thy b, T)
   938   in
   939     (c,
   940      ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt),
   941       Sign.add_consts_i [(b, T, NoSyn)] thy))
   942   end;
   943 
   944 fun add_def prfx types pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) =
   945   (case lookup consts s of
   946      SOME ty =>
   947        let
   948          val (t, ty') = term_of_expr thy prfx types pfuns ids e;
   949          val T = mk_type thy prfx ty;
   950          val T' = mk_type thy prfx ty';
   951          val _ = T = T' orelse
   952            error ("Declared type " ^ ty ^ " of " ^ s ^
   953              "\ndoes not match type " ^ ty' ^ " in definition");
   954          val id' = mk_rulename id;
   955          val lthy = Named_Target.theory_init thy;
   956          val ((t', (_, th)), lthy') = Specification.definition
   957            (NONE, ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq
   958              (Free (s, T), t)))) lthy;
   959          val phi = Proof_Context.export_morphism lthy' lthy
   960        in
   961          ((id', Morphism.thm phi th),
   962           ((Symtab.update (s, (Morphism.term phi t', ty)) tab,
   963             Name.declare s ctxt),
   964            Local_Theory.exit_global lthy'))
   965        end
   966    | NONE => error ("Undeclared constant " ^ s));
   967 
   968 fun add_var prfx (s, ty) (ids, thy) =
   969   let val ([Free p], ids') = mk_variables thy prfx [s] ty ids
   970   in (p, (ids', thy)) end;
   971 
   972 fun add_init_vars prfx vcs (ids_thy as ((tab, _), _)) =
   973   fold_map (add_var prfx)
   974     (map_filter
   975        (fn s => case try (unsuffix "~") s of
   976           SOME s' => (case Symtab.lookup tab s' of
   977             SOME (_, ty) => SOME (s, ty)
   978           | NONE => error ("Undeclared identifier " ^ s'))
   979         | NONE => NONE)
   980        (fold_vcs (add_expr_idents o snd) vcs []))
   981     ids_thy;
   982 
   983 fun is_trivial_vc ([], [(_, Ident "true")]) = true
   984   | is_trivial_vc _ = false;
   985 
   986 fun rulenames rules = commas
   987   (map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules);
   988 
   989 (* sort definitions according to their dependency *)
   990 fun sort_defs _ _ _ _ [] sdefs = rev sdefs
   991   | sort_defs prfx funs pfuns consts defs sdefs =
   992       (case find_first (fn (_, (_, e)) =>
   993          forall (is_some o lookup_prfx prfx pfuns)
   994            (add_expr_pfuns funs e []) andalso
   995          forall (fn id =>
   996            member (fn (s, (_, (s', _))) => s = s') sdefs id orelse
   997            consts id)
   998              (add_expr_idents e [])) defs of
   999          SOME d => sort_defs prfx funs pfuns consts
  1000            (remove (op =) d defs) (d :: sdefs)
  1001        | NONE => error ("Bad definitions: " ^ rulenames defs));
  1002 
  1003 fun set_vcs ({types, vars, consts, funs} : decls)
  1004       (rules, _) ((_, ident), vcs) path prfx thy =
  1005   let
  1006     val {pfuns, ...} = VCs.get thy;
  1007     val (defs, rules') = partition_opt dest_def rules;
  1008     val consts' =
  1009       subtract (fn ((_, (s, _)), (s', _)) => s = s') defs (items consts);
  1010     (* ignore all complex rules in rls files *)
  1011     val (rules'', other_rules) =
  1012       List.partition (complex_rule o snd) rules';
  1013     val _ = if null rules'' then ()
  1014       else warning ("Ignoring rules: " ^ rulenames rules'');
  1015 
  1016     val vcs' = VCtab.make (maps (fn (tr, vcs) =>
  1017       map (fn (s, (ps, cs)) => (s, (tr, NONE, ps, cs)))
  1018         (filter_out (is_trivial_vc o snd) vcs)) vcs);
  1019 
  1020     val _ = (case filter_out (is_some o lookup funs)
  1021         (pfuns_of_vcs prfx funs pfuns vcs') of
  1022         [] => ()
  1023       | fs => error ("Undeclared proof function(s) " ^ commas fs));
  1024 
  1025     val (((defs', vars''), ivars), (ids, thy')) =
  1026       ((Symtab.empty |>
  1027         Symtab.update ("false", (@{term False}, booleanN)) |>
  1028         Symtab.update ("true", (@{term True}, booleanN)),
  1029         Name.context),
  1030        thy |> Sign.add_path (Long_Name.base_name ident)) |>
  1031       fold (add_type_def prfx) (items types) |>
  1032       fold (snd oo add_const prfx) consts' |> (fn ids_thy as ((tab, _), _) =>
  1033         ids_thy |>
  1034         fold_map (add_def prfx types pfuns consts)
  1035           (sort_defs prfx funs pfuns (Symtab.defined tab) defs []) ||>>
  1036         fold_map (add_var prfx) (items vars) ||>>
  1037         add_init_vars prfx vcs');
  1038 
  1039     val ctxt =
  1040       [Element.Fixes (map (fn (s, T) =>
  1041          (Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)),
  1042        Element.Assumes (map (fn (id, rl) =>
  1043          ((mk_rulename id, []),
  1044           [(term_of_rule thy' prfx types pfuns ids rl, [])]))
  1045            other_rules),
  1046        Element.Notes ("", [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])]
  1047           
  1048   in
  1049     set_env ctxt defs' types funs ids vcs' path prfx thy'
  1050   end;
  1051 
  1052 end;