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