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